diff --git a/src/HOL/Analysis/Analysis.thy b/src/HOL/Analysis/Analysis.thy --- a/src/HOL/Analysis/Analysis.thy +++ b/src/HOL/Analysis/Analysis.thy @@ -1,50 +1,48 @@ theory Analysis imports (* Linear Algebra *) Convex Determinants (* Topology *) Connected Abstract_Limits Abstract_Euclidean_Space (* Functional Analysis *) Elementary_Normed_Spaces Norm_Arith (* Vector Analysis *) Convex_Euclidean_Space (* Measure and Integration Theory *) Ball_Volume Integral_Test Improper_Integral Equivalence_Measurable_On_Borel (* Unsorted *) Lebesgue_Integral_Substitution Embed_Measure Complete_Measure Radon_Nikodym Fashoda_Theorem Determinants Cross3 Homeomorphism Bounded_Continuous_Function Abstract_Topology Product_Topology Lindelof_Spaces Infinite_Products Infinite_Set_Sum Weierstrass_Theorems Polytope Jordan_Curve - Winding_Numbers - Riemann_Mapping Poly_Roots - Conformal_Mappings - FPS_Convergence Generalised_Binomial_Theorem Gamma_Function Change_Of_Vars Multivariate_Analysis Simplex_Content + FPS_Convergence + Smooth_Paths begin end diff --git a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy deleted file mode 100644 --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy +++ /dev/null @@ -1,7847 +0,0 @@ -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 -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) - -lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" - by (auto simp: C1_differentiable_on_eq) - -lemma C1_differentiable_on_add [simp, derivative_intros]: - "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x + g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_minus [simp, derivative_intros]: - "f C1_differentiable_on S \ (\x. - f x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_diff [simp, derivative_intros]: - "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x - g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) - -lemma C1_differentiable_on_mult [simp, derivative_intros]: - fixes f g :: "real \ 'a :: real_normed_algebra" - shows "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x * g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq - by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) - -lemma C1_differentiable_on_scaleR [simp, derivative_intros]: - "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x *\<^sub>R g x) C1_differentiable_on S" - unfolding C1_differentiable_on_eq - by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ - - -definition\<^marker>\tag important\ piecewise_C1_differentiable_on - (infixr "piecewise'_C1'_differentiable'_on" 50) - where "f piecewise_C1_differentiable_on i \ - continuous_on i f \ - (\S. finite S \ (f C1_differentiable_on (i - S)))" - -lemma C1_differentiable_imp_piecewise: - "f C1_differentiable_on S \ f piecewise_C1_differentiable_on S" - by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) - -lemma piecewise_C1_imp_differentiable: - "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" - by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def - C1_differentiable_on_def differentiable_def has_vector_derivative_def - intro: has_derivative_at_withinI) - -lemma piecewise_C1_differentiable_compose: - assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" - shows "(g \ f) piecewise_C1_differentiable_on S" -proof - - have "continuous_on S (\x. g (f x))" - by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) - moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" - proof - - obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" - using fg by (auto simp: piecewise_C1_differentiable_on_def) - obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" - using fg by (auto simp: piecewise_C1_differentiable_on_def) - show ?thesis - proof (intro exI conjI) - show "finite (F \ (\x\G. S \ f-`{x}))" - using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) - show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" - apply (rule C1_differentiable_compose) - apply (blast intro: C1_differentiable_on_subset [OF F]) - apply (blast intro: C1_differentiable_on_subset [OF G]) - by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) - qed - qed - ultimately show ?thesis - by (simp add: piecewise_C1_differentiable_on_def) -qed - -lemma piecewise_C1_differentiable_on_subset: - "f piecewise_C1_differentiable_on S \ T \ S \ f piecewise_C1_differentiable_on T" - by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) - -lemma C1_differentiable_imp_continuous_on: - "f C1_differentiable_on S \ continuous_on S f" - unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within - using differentiable_at_withinI differentiable_imp_continuous_within by blast - -lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" - unfolding C1_differentiable_on_def - by auto - -lemma piecewise_C1_differentiable_affine: - fixes m::real - assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` S)" - shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on S" -proof (cases "m = 0") - case True - then show ?thesis - unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def) -next - case False - have *: "\x. finite (S \ {y. m * y + c = x})" - using False not_finite_existsD by fastforce - show ?thesis - apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) - apply (rule * assms derivative_intros | simp add: False vimage_def)+ - done -qed - -lemma piecewise_C1_differentiable_cases: - fixes c::real - assumes "f piecewise_C1_differentiable_on {a..c}" - "g piecewise_C1_differentiable_on {c..b}" - "a \ c" "c \ b" "f c = g c" - shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" -proof - - obtain S T where st: "f C1_differentiable_on ({a..c} - S)" - "g C1_differentiable_on ({c..b} - T)" - "finite S" "finite T" - using assms - by (force simp: piecewise_C1_differentiable_on_def) - then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" - using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], - OF closed_real_atLeastAtMost [of c b], - of f g "\x. x\c"] assms - by (force simp: ivl_disj_un_two_touch) - { fix x - assume x: "x \ {a..b} - insert c (S \ T)" - have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") - proof (cases x c rule: le_cases) - case le show ?diff_fg - apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) - using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) - next - case ge show ?diff_fg - apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) - using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) - qed - } - then have "(\x \ {a..b} - insert c (S \ T). (\x. if x \ c then f x else g x) differentiable at x)" - by auto - moreover - { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" - and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" - have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - proof - - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" - if "a < x" "x < c" "x \ S" for x - proof - - have f: "f differentiable at x" - by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) - show ?thesis - using that - apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) - apply (auto simp: dist_norm vector_derivative_works [symmetric] f) - done - qed - then show ?thesis - by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) - qed - moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - proof - - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" - if "c < x" "x < b" "x \ T" for x - proof - - have g: "g differentiable at x" - by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) - show ?thesis - using that - apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) - apply (auto simp: dist_norm vector_derivative_works [symmetric] g) - done - qed - then show ?thesis - by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) - qed - ultimately have "continuous_on ({a<.. T)) - (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - by (rule continuous_on_subset [OF continuous_on_open_Un], auto) - } note * = this - have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" - using st - by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) - ultimately have "\S. finite S \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - S)" - apply (rule_tac x="{a,b,c} \ S \ T" in exI) - using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) - with cab show ?thesis - by (simp add: piecewise_C1_differentiable_on_def) -qed - -lemma piecewise_C1_differentiable_neg: - "f piecewise_C1_differentiable_on S \ (\x. -(f x)) piecewise_C1_differentiable_on S" - unfolding piecewise_C1_differentiable_on_def - by (auto intro!: continuous_on_minus C1_differentiable_on_minus) - -lemma piecewise_C1_differentiable_add: - assumes "f piecewise_C1_differentiable_on i" - "g piecewise_C1_differentiable_on i" - shows "(\x. f x + g x) piecewise_C1_differentiable_on i" -proof - - obtain S t where st: "finite S" "finite t" - "f C1_differentiable_on (i-S)" - "g C1_differentiable_on (i-t)" - using assms by (auto simp: piecewise_C1_differentiable_on_def) - then have "finite (S \ t) \ (\x. f x + g x) C1_differentiable_on i - (S \ t)" - by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) - moreover have "continuous_on i f" "continuous_on i g" - using assms piecewise_C1_differentiable_on_def by auto - ultimately show ?thesis - by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) -qed - -lemma piecewise_C1_differentiable_diff: - "\f piecewise_C1_differentiable_on S; g piecewise_C1_differentiable_on S\ - \ (\x. f x - g x) piecewise_C1_differentiable_on S" - unfolding diff_conv_add_uminus - by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) - -lemma piecewise_C1_differentiable_D1: - fixes g1 :: "real \ 'a::real_normed_field" - assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" - shows "g1 piecewise_C1_differentiable_on {0..1}" -proof - - obtain S where "finite S" - and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" - and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" - using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x - proof (rule differentiable_transform_within) - show "g1 +++ g2 \ (*) (inverse 2) differentiable at x" - using that g12D - apply (simp only: joinpaths_def) - by (rule differentiable_chain_at derivative_intros | force)+ - show "\x'. \dist x' x < dist (x/2) (1/2)\ - \ (g1 +++ g2 \ (*) (inverse 2)) x' = g1 x'" - using that by (auto simp: dist_real_def joinpaths_def) - qed (use that in \auto simp: dist_real_def\) - have [simp]: "vector_derivative (g1 \ (*) 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" - if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x - apply (subst vector_derivative_chain_at) - using that - apply (rule derivative_eq_intros g1D | simp)+ - done - have "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" - using co12 by (rule continuous_on_subset) force - then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 \ (*)2) (at x))" - proof (rule continuous_on_eq [OF _ vector_derivative_at]) - show "(g1 +++ g2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" - if "x \ {0..1/2} - insert (1/2) S" for x - proof (rule has_vector_derivative_transform_within) - show "(g1 \ (*) 2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" - using that - by (force intro: g1D differentiable_chain_at simp: vector_derivative_works [symmetric]) - show "\x'. \dist x' x < dist x (1/2)\ \ (g1 \ (*) 2) x' = (g1 +++ g2) x'" - using that by (auto simp: dist_norm joinpaths_def) - qed (use that in \auto simp: dist_norm\) - qed - have "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) - ((\x. 1/2 * vector_derivative (g1 \ (*)2) (at x)) \ (*)(1/2))" - apply (rule continuous_intros)+ - using coDhalf - apply (simp add: scaleR_conv_of_real image_set_diff image_image) - done - then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\x. vector_derivative g1 (at x))" - by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) - have "continuous_on {0..1} g1" - using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast - with \finite S\ show ?thesis - apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - apply (rule_tac x="insert 1 (((*)2)`S)" in exI) - apply (simp add: g1D con_g1) - done -qed - -lemma piecewise_C1_differentiable_D2: - fixes g2 :: "real \ 'a::real_normed_field" - assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" - shows "g2 piecewise_C1_differentiable_on {0..1}" -proof - - obtain S where "finite S" - and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" - and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" - using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x - proof (rule differentiable_transform_within) - show "g1 +++ g2 \ (\x. (x + 1) / 2) differentiable at x" - using g12D that - apply (simp only: joinpaths_def) - apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps) - apply (rule differentiable_chain_at derivative_intros | force)+ - done - show "\x'. dist x' x < dist ((x + 1) / 2) (1/2) \ (g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" - using that by (auto simp: dist_real_def joinpaths_def field_simps) - qed (use that in \auto simp: dist_norm\) - have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" - if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x - using that by (auto simp: vector_derivative_chain_at field_split_simps g2D) - have "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" - using co12 by (rule continuous_on_subset) force - then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g2 \ (\x. 2*x-1)) (at x))" - proof (rule continuous_on_eq [OF _ vector_derivative_at]) - show "(g1 +++ g2 has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) - (at x)" - if "x \ {1 / 2..1} - insert (1 / 2) S" for x - proof (rule_tac f="g2 \ (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) - show "(g2 \ (\x. 2 * x - 1) has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) - (at x)" - using that by (force intro: g2D differentiable_chain_at simp: vector_derivative_works [symmetric]) - show "\x'. \dist x' x < dist (3 / 4) ((x + 1) / 2)\ \ (g2 \ (\x. 2 * x - 1)) x' = (g1 +++ g2) x'" - using that by (auto simp: dist_norm joinpaths_def add_divide_distrib) - qed (use that in \auto simp: dist_norm\) - qed - have [simp]: "((\x. (x+1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` S))) = ({1/2..1} - insert (1/2) S)" - apply (simp add: image_set_diff inj_on_def image_image) - apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) - done - have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) - ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) \ (\x. (x+1)/2))" - by (rule continuous_intros | simp add: coDhalf)+ - then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) (\x. vector_derivative g2 (at x))" - by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) - have "continuous_on {0..1} g2" - using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast - with \finite S\ show ?thesis - apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` S)" in exI) - apply (simp add: g2D con_g2) - done -qed - -subsection \Valid paths, and their start and finish\ - -definition\<^marker>\tag important\ valid_path :: "(real \ 'a :: real_normed_vector) \ bool" - where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" - -definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" - where "closed_path g \ g 0 = g 1" - -text\In particular, all results for paths apply\ - -lemma valid_path_imp_path: "valid_path g \ path g" - by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) - -lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" - by (metis connected_path_image valid_path_imp_path) - -lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" - by (metis compact_path_image valid_path_imp_path) - -lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" - by (metis bounded_path_image valid_path_imp_path) - -lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" - by (metis closed_path_image valid_path_imp_path) - -lemma valid_path_compose: - assumes "valid_path g" - and der: "\x. x \ path_image g \ f field_differentiable (at x)" - and con: "continuous_on (path_image g) (deriv f)" - shows "valid_path (f \ g)" -proof - - obtain S where "finite S" and g_diff: "g C1_differentiable_on {0..1} - S" - using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto - have "f \ g differentiable at t" when "t\{0..1} - S" for t - proof (rule differentiable_chain_at) - show "g differentiable at t" using \valid_path g\ - by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - S\ that) - next - have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis - then show "f differentiable at (g t)" - using der[THEN field_differentiable_imp_differentiable] by auto - qed - moreover have "continuous_on ({0..1} - S) (\x. vector_derivative (f \ g) (at x))" - proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], - rule continuous_intros) - show "continuous_on ({0..1} - S) (\x. vector_derivative g (at x))" - using g_diff C1_differentiable_on_eq by auto - next - have "continuous_on {0..1} (\x. deriv f (g x))" - using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] - \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def - by blast - then show "continuous_on ({0..1} - S) (\x. deriv f (g x))" - using continuous_on_subset by blast - next - show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" - when "t \ {0..1} - S" for t - proof (rule vector_derivative_chain_at_general[symmetric]) - show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) - next - have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis - then show "f field_differentiable at (g t)" using der by auto - qed - qed - ultimately have "f \ g C1_differentiable_on {0..1} - S" - using C1_differentiable_on_eq by blast - moreover have "path (f \ g)" - apply (rule path_continuous_image[OF valid_path_imp_path[OF \valid_path g\]]) - using der - by (simp add: continuous_at_imp_continuous_on field_differentiable_imp_continuous_at) - ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def - using \finite S\ by auto -qed - -lemma valid_path_uminus_comp[simp]: - fixes g::"real \ 'a ::real_normed_field" - shows "valid_path (uminus \ g) \ valid_path g" -proof - show "valid_path g \ valid_path (uminus \ g)" for g::"real \ 'a" - by (auto intro!: valid_path_compose derivative_intros simp add: deriv_linear[of "-1",simplified]) - then show "valid_path g" when "valid_path (uminus \ g)" - by (metis fun.map_comp group_add_class.minus_comp_minus id_comp that) -qed - -lemma valid_path_offset[simp]: - shows "valid_path (\t. g t - z) \ valid_path g" -proof - show *: "valid_path (g::real\'a) \ valid_path (\t. g t - z)" for g z - unfolding valid_path_def - by (fastforce intro:derivative_intros C1_differentiable_imp_piecewise piecewise_C1_differentiable_diff) - show "valid_path (\t. g t - z) \ valid_path g" - using *[of "\t. g t - z" "-z",simplified] . -qed - - -subsection\Contour Integrals along a path\ - -text\This definition is for complex numbers only, and does not generalise to line integrals in a vector field\ - -text\piecewise differentiable function on [0,1]\ - -definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" - (infixr "has'_contour'_integral" 50) - where "(f has_contour_integral i) g \ - ((\x. f(g x) * vector_derivative g (at x within {0..1})) - has_integral i) {0..1}" - -definition\<^marker>\tag important\ contour_integrable_on - (infixr "contour'_integrable'_on" 50) - where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" - -definition\<^marker>\tag important\ contour_integral - where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" - -lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" - unfolding contour_integrable_on_def contour_integral_def by blast - -lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" - apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) - using has_integral_unique by blast - -lemma has_contour_integral_eqpath: - "\(f has_contour_integral y) p; f contour_integrable_on \; - contour_integral p f = contour_integral \ f\ - \ (f has_contour_integral y) \" -using contour_integrable_on_def contour_integral_unique by auto - -lemma has_contour_integral_integral: - "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" - by (metis contour_integral_unique contour_integrable_on_def) - -lemma has_contour_integral_unique: - "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" - using has_integral_unique - by (auto simp: has_contour_integral_def) - -lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" - using contour_integrable_on_def by blast - -text\Show that we can forget about the localized derivative.\ - -lemma has_integral_localized_vector_derivative: - "((\x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \ - ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" -proof - - have *: "{a..b} - {a,b} = interior {a..b}" - by (simp add: atLeastAtMost_diff_ends) - show ?thesis - apply (rule has_integral_spike_eq [of "{a,b}"]) - apply (auto simp: at_within_interior [of _ "{a..b}"]) - done -qed - -lemma integrable_on_localized_vector_derivative: - "(\x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \ - (\x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" - by (simp add: integrable_on_def has_integral_localized_vector_derivative) - -lemma has_contour_integral: - "(f has_contour_integral i) g \ - ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" - by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) - -lemma contour_integrable_on: - "f contour_integrable_on g \ - (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" - by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) - -subsection\<^marker>\tag unimportant\ \Reversing a path\ - -lemma valid_path_imp_reverse: - assumes "valid_path g" - shows "valid_path(reversepath g)" -proof - - obtain S where "finite S" and S: "g C1_differentiable_on ({0..1} - S)" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - then have "finite ((-) 1 ` S)" - by auto - moreover have "(reversepath g C1_differentiable_on ({0..1} - (-) 1 ` S))" - unfolding reversepath_def - apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) - using S - by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq elim!: continuous_on_subset)+ - ultimately show ?thesis using assms - by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) -qed - -lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" - using valid_path_imp_reverse by force - -lemma has_contour_integral_reversepath: - assumes "valid_path g" and f: "(f has_contour_integral i) g" - shows "(f has_contour_integral (-i)) (reversepath g)" -proof - - { fix S x - assume xs: "g C1_differentiable_on ({0..1} - S)" "x \ (-) 1 ` S" "0 \ x" "x \ 1" - have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - - vector_derivative g (at (1 - x) within {0..1})" - proof - - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" - using xs - by (force simp: has_vector_derivative_def C1_differentiable_on_def) - have "(g \ (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" - by (intro vector_diff_chain_within has_vector_derivative_at_within [OF f'] derivative_eq_intros | simp)+ - then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" - by (simp add: o_def) - show ?thesis - using xs - by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) - qed - } note * = this - obtain S where S: "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - have "((\x. - (f (g (1 - x)) * vector_derivative g (at (1 - x) within {0..1}))) has_integral -i) - {0..1}" - using has_integral_affinity01 [where m= "-1" and c=1, OF f [unfolded has_contour_integral_def]] - by (simp add: has_integral_neg) - then show ?thesis - using S - apply (clarsimp simp: reversepath_def has_contour_integral_def) - apply (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) - apply (auto simp: *) - done -qed - -lemma contour_integrable_reversepath: - "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" - using has_contour_integral_reversepath contour_integrable_on_def by blast - -lemma contour_integrable_reversepath_eq: - "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" - using contour_integrable_reversepath valid_path_reversepath by fastforce - -lemma contour_integral_reversepath: - assumes "valid_path g" - shows "contour_integral (reversepath g) f = - (contour_integral g f)" -proof (cases "f contour_integrable_on g") - case True then show ?thesis - by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) -next - case False then have "\ f contour_integrable_on (reversepath g)" - by (simp add: assms contour_integrable_reversepath_eq) - with False show ?thesis by (simp add: not_integrable_contour_integral) -qed - - -subsection\<^marker>\tag unimportant\ \Joining two paths together\ - -lemma valid_path_join: - assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" - shows "valid_path(g1 +++ g2)" -proof - - have "g1 1 = g2 0" - using assms by (auto simp: pathfinish_def pathstart_def) - moreover have "(g1 \ (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" - apply (rule piecewise_C1_differentiable_compose) - using assms - apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) - apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) - done - moreover have "(g2 \ (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" - apply (rule piecewise_C1_differentiable_compose) - using assms unfolding valid_path_def piecewise_C1_differentiable_on_def - by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI - simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) - ultimately show ?thesis - apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: o_def) - done -qed - -lemma valid_path_join_D1: - fixes g1 :: "real \ 'a::real_normed_field" - shows "valid_path (g1 +++ g2) \ valid_path g1" - unfolding valid_path_def - by (rule piecewise_C1_differentiable_D1) - -lemma valid_path_join_D2: - fixes g2 :: "real \ 'a::real_normed_field" - shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" - unfolding valid_path_def - by (rule piecewise_C1_differentiable_D2) - -lemma valid_path_join_eq [simp]: - fixes g2 :: "real \ 'a::real_normed_field" - shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" - using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast - -lemma has_contour_integral_join: - assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" - "valid_path g1" "valid_path g2" - shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" -proof - - obtain s1 s2 - where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" - and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" - using assms - by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" - and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" - using assms - by (auto simp: has_contour_integral) - have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" - and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" - using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] - has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] - by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) - have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = - 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\z - 1/2\"]]) - apply (simp_all add: dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) - apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - using s1 - apply (auto simp: algebra_simps vector_derivative_works) - done - have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = - 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2 (2*x - 1))" and d = "\z - 1/2\"]]) - apply (simp_all add: dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) - apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - using s2 - apply (auto simp: algebra_simps vector_derivative_works) - done - have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" - apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) - using s1 - apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) - apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) - done - moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" - apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) - using s2 - apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) - apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) - done - ultimately - show ?thesis - apply (simp add: has_contour_integral) - apply (rule has_integral_combine [where c = "1/2"], auto) - done -qed - -lemma contour_integrable_joinI: - assumes "f contour_integrable_on g1" "f contour_integrable_on g2" - "valid_path g1" "valid_path g2" - shows "f contour_integrable_on (g1 +++ g2)" - using assms - by (meson has_contour_integral_join contour_integrable_on_def) - -lemma contour_integrable_joinD1: - assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" - shows "f contour_integrable_on g1" -proof - - obtain s1 - where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" - using assms - apply (auto simp: contour_integrable_on) - apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) - apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) - done - then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" - by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) - have g1: "\0 < z; z < 1; z \ s1\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = - 2 *\<^sub>R vector_derivative g1 (at z)" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\(z-1)/2\"]]) - apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x*2" 2 _ g1, simplified o_def]) - using s1 - apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) - done - show ?thesis - using s1 - apply (auto simp: contour_integrable_on) - apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) - apply (auto simp: joinpaths_def scaleR_conv_of_real g1) - done -qed - -lemma contour_integrable_joinD2: - assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" - shows "f contour_integrable_on g2" -proof - - obtain s2 - where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" - using assms - apply (auto simp: contour_integrable_on) - apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) - apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) - apply (simp add: image_affinity_atLeastAtMost_diff) - done - then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) - integrable_on {0..1}" - by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) - have g2: "\0 < z; z < 1; z \ s2\ \ - vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = - 2 *\<^sub>R vector_derivative g2 (at z)" for z - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2(2*x-1))" and d = "\z/2\"]]) - apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x*2-1" 2 _ g2, simplified o_def]) - using s2 - apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left - vector_derivative_works add_divide_distrib) - done - show ?thesis - using s2 - apply (auto simp: contour_integrable_on) - apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) - apply (auto simp: joinpaths_def scaleR_conv_of_real g2) - done -qed - -lemma contour_integrable_join [simp]: - shows - "\valid_path g1; valid_path g2\ - \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" -using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast - -lemma contour_integral_join [simp]: - shows - "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ - \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" - by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) - - -subsection\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ - -lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" - by (auto simp: shiftpath_def) - -lemma valid_path_shiftpath [intro]: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "valid_path(shiftpath a g)" - using assms - apply (auto simp: valid_path_def shiftpath_alt_def) - apply (rule piecewise_C1_differentiable_cases) - apply (auto simp: algebra_simps) - apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) - apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) - apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) - done - -lemma has_contour_integral_shiftpath: - assumes f: "(f has_contour_integral i) g" "valid_path g" - and a: "a \ {0..1}" - shows "(f has_contour_integral i) (shiftpath a g)" -proof - - obtain s - where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" - using assms by (auto simp: has_contour_integral) - then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + - integral {0..a} (\x. f (g x) * vector_derivative g (at x))" - apply (rule has_integral_unique) - apply (subst add.commute) - apply (subst integral_combine) - using assms * integral_unique by auto - { fix x - have "0 \ x \ x + a < 1 \ x \ (\x. x - a) ` s \ - vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" - unfolding shiftpath_def - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x))" and d = "dist(1-a) x"]]) - apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x+a" 1 _ g, simplified o_def scaleR_one]) - apply (intro derivative_eq_intros | simp)+ - using g - apply (drule_tac x="x+a" in bspec) - using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) - done - } note vd1 = this - { fix x - have "1 < x + a \ x \ 1 \ x \ (\x. x - a + 1) ` s \ - vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" - unfolding shiftpath_def - apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x-1))" and d = "dist (1-a) x"]]) - apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) - apply (rule vector_diff_chain_at [of "\x. x+a-1" 1 _ g, simplified o_def scaleR_one]) - apply (intro derivative_eq_intros | simp)+ - using g - apply (drule_tac x="x+a-1" in bspec) - using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) - done - } note vd2 = this - have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" - using * a by (fastforce intro: integrable_subinterval_real) - have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" - apply (rule integrable_subinterval_real) - using * a by auto - have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) - has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" - apply (rule has_integral_spike_finite - [where S = "{1-a} \ (\x. x-a) ` s" and f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) - using s apply blast - using a apply (auto simp: algebra_simps vd1) - apply (force simp: shiftpath_def add.commute) - using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] - apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) - done - moreover - have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) - has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" - apply (rule has_integral_spike_finite - [where S = "{1-a} \ (\x. x-a+1) ` s" and f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) - using s apply blast - using a apply (auto simp: algebra_simps vd2) - apply (force simp: shiftpath_def add.commute) - using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] - apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) - apply (simp add: algebra_simps) - done - ultimately show ?thesis - using a - by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) -qed - -lemma has_contour_integral_shiftpath_D: - assumes "(f has_contour_integral i) (shiftpath a g)" - "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "(f has_contour_integral i) g" -proof - - obtain s - where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - { fix x - assume x: "0 < x" "x < 1" "x \ s" - then have gx: "g differentiable at x" - using g by auto - have "vector_derivative g (at x within {0..1}) = - vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" - apply (rule vector_derivative_at_within_ivl - [OF has_vector_derivative_transform_within_open - [where f = "(shiftpath (1 - a) (shiftpath a g))" and S = "{0<..<1}-s"]]) - using s g assms x - apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath - at_within_interior [of _ "{0..1}"] vector_derivative_works [symmetric]) - apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) - apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) - done - } note vd = this - have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" - using assms by (auto intro!: has_contour_integral_shiftpath) - show ?thesis - apply (simp add: has_contour_integral_def) - apply (rule has_integral_spike_finite [of "{0,1} \ s", OF _ _ fi [unfolded has_contour_integral_def]]) - using s assms vd - apply (auto simp: Path_Connected.shiftpath_shiftpath) - done -qed - -lemma has_contour_integral_shiftpath_eq: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" - using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast - -lemma contour_integrable_on_shiftpath_eq: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" -using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto - -lemma contour_integral_shiftpath: - assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" - shows "contour_integral (shiftpath a g) f = contour_integral g f" - using assms - by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) - - -subsection\<^marker>\tag unimportant\ \More about straight-line paths\ - -lemma has_vector_derivative_linepath_within: - "(linepath a b has_vector_derivative (b - a)) (at x within s)" -apply (simp add: linepath_def has_vector_derivative_def algebra_simps) -apply (rule derivative_eq_intros | simp)+ -done - -lemma vector_derivative_linepath_within: - "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" - apply (rule vector_derivative_within_cbox [of 0 "1::real", simplified]) - apply (auto simp: has_vector_derivative_linepath_within) - done - -lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" - by (simp add: has_vector_derivative_linepath_within vector_derivative_at) - -lemma valid_path_linepath [iff]: "valid_path (linepath a b)" - apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) - apply (rule_tac x="{}" in exI) - apply (simp add: differentiable_on_def differentiable_def) - using has_vector_derivative_def has_vector_derivative_linepath_within - apply (fastforce simp add: continuous_on_eq_continuous_within) - done - -lemma has_contour_integral_linepath: - shows "(f has_contour_integral i) (linepath a b) \ - ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" - by (simp add: has_contour_integral) - -lemma linepath_in_path: - shows "x \ {0..1} \ linepath a b x \ closed_segment a b" - by (auto simp: segment linepath_def) - -lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" - by (auto simp: segment linepath_def) - -lemma linepath_in_convex_hull: - fixes x::real - assumes a: "a \ convex hull s" - and b: "b \ convex hull s" - and x: "0\x" "x\1" - shows "linepath a b x \ convex hull s" - apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) - using x - apply (auto simp: linepath_image_01 [symmetric]) - done - -lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" - by (simp add: linepath_def) - -lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" - by (simp add: linepath_def) - -lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" - by (simp add: has_contour_integral_linepath) - -lemma has_contour_integral_trivial_iff [simp]: "(f has_contour_integral i) (linepath a a) \ i=0" - using has_contour_integral_unique by blast - -lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" - using has_contour_integral_trivial contour_integral_unique by blast - -lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" - by (auto simp: linepath_def) - -lemma bounded_linear_linepath: - assumes "bounded_linear f" - shows "f (linepath a b x) = linepath (f a) (f b) x" -proof - - interpret f: bounded_linear f by fact - show ?thesis by (simp add: linepath_def f.add f.scale) -qed - -lemma bounded_linear_linepath': - assumes "bounded_linear f" - shows "f \ linepath a b = linepath (f a) (f b)" - using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) - -lemma cnj_linepath: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" - by (simp add: linepath_def) - -lemma cnj_linepath': "cnj \ linepath a b = linepath (cnj a) (cnj b)" - by (simp add: linepath_def fun_eq_iff) - -subsection\Relation to subpath construction\ - -lemma valid_path_subpath: - fixes g :: "real \ 'a :: real_normed_vector" - assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" - shows "valid_path(subpath u v g)" -proof (cases "v=u") - case True - then show ?thesis - unfolding valid_path_def subpath_def - by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) -next - case False - have "(g \ (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" - apply (rule piecewise_C1_differentiable_compose) - apply (simp add: C1_differentiable_imp_piecewise) - apply (simp add: image_affinity_atLeastAtMost) - using assms False - apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) - apply (subst Int_commute) - apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) - done - then show ?thesis - by (auto simp: o_def valid_path_def subpath_def) -qed - -lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" - by (simp add: has_contour_integral subpath_def) - -lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" - using has_contour_integral_subpath_refl contour_integrable_on_def by blast - -lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" - by (simp add: contour_integral_unique) - -lemma has_contour_integral_subpath: - assumes f: "f contour_integrable_on g" and g: "valid_path g" - and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) - (subpath u v g)" -proof (cases "v=u") - case True - then show ?thesis - using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) -next - case False - obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" - using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast - have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) - has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) - {0..1}" - using f uv - apply (simp add: contour_integrable_on subpath_def has_contour_integral) - apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) - apply (simp_all add: has_integral_integral) - apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) - apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) - apply (simp add: divide_simps False) - done - { fix x - have "x \ {0..1} \ - x \ (\t. (v-u) *\<^sub>R t + u) -` s \ - vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" - apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) - apply (intro derivative_eq_intros | simp)+ - apply (cut_tac s [of "(v - u) * x + u"]) - using uv mult_left_le [of x "v-u"] - apply (auto simp: vector_derivative_works) - done - } note vd = this - show ?thesis - apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) - using fs assms - apply (simp add: False subpath_def has_contour_integral) - apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) - apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) - done -qed - -lemma contour_integrable_subpath: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" - shows "f contour_integrable_on (subpath u v g)" - apply (cases u v rule: linorder_class.le_cases) - apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) - apply (subst reversepath_subpath [symmetric]) - apply (rule contour_integrable_reversepath) - using assms apply (blast intro: valid_path_subpath) - apply (simp add: contour_integrable_on_def) - using assms apply (blast intro: has_contour_integral_subpath) - done - -lemma has_integral_contour_integral_subpath: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "(((\x. f(g x) * vector_derivative g (at x))) - has_integral contour_integral (subpath u v g) f) {u..v}" - using assms - apply (auto simp: has_integral_integrable_integral) - apply (rule integrable_on_subcbox [where a=u and b=v and S = "{0..1}", simplified]) - apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) - done - -lemma contour_integral_subcontour_integral: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" - shows "contour_integral (subpath u v g) f = - integral {u..v} (\x. f(g x) * vector_derivative g (at x))" - using assms has_contour_integral_subpath contour_integral_unique by blast - -lemma contour_integral_subpath_combine_less: - assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" - "u {0..1}" "v \ {0..1}" "w \ {0..1}" - shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = - contour_integral (subpath u w g) f" -proof (cases "u\v \ v\w \ u\w") - case True - have *: "subpath v u g = reversepath(subpath u v g) \ - subpath w u g = reversepath(subpath u w g) \ - subpath w v g = reversepath(subpath v w g)" - by (auto simp: reversepath_subpath) - have "u < v \ v < w \ - u < w \ w < v \ - v < u \ u < w \ - v < w \ w < u \ - w < u \ u < v \ - w < v \ v < u" - using True assms by linarith - with assms show ?thesis - using contour_integral_subpath_combine_less [of f g u v w] - contour_integral_subpath_combine_less [of f g u w v] - contour_integral_subpath_combine_less [of f g v u w] - contour_integral_subpath_combine_less [of f g v w u] - contour_integral_subpath_combine_less [of f g w u v] - contour_integral_subpath_combine_less [of f g w v u] - apply simp - apply (elim disjE) - apply (auto simp: * contour_integral_reversepath contour_integrable_subpath - valid_path_subpath algebra_simps) - done -next - case False - then show ?thesis - apply (auto) - using assms - by (metis eq_neg_iff_add_eq_0 contour_integral_reversepath reversepath_subpath valid_path_subpath) -qed - -lemma contour_integral_integral: - "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" - by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) - -lemma contour_integral_cong: - assumes "g = g'" "\x. x \ path_image g \ f x = f' x" - shows "contour_integral g f = contour_integral g' f'" - unfolding contour_integral_integral using assms - by (intro integral_cong) (auto simp: path_image_def) - - -text \Contour integral along a segment on the real axis\ - -lemma has_contour_integral_linepath_Reals_iff: - fixes a b :: complex and f :: "complex \ complex" - assumes "a \ Reals" "b \ Reals" "Re a < Re b" - shows "(f has_contour_integral I) (linepath a b) \ - ((\x. f (of_real x)) has_integral I) {Re a..Re b}" -proof - - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" - by (simp_all add: complex_eq_iff) - from assms have "a \ b" by auto - have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ - ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" - by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) - (insert assms, simp_all add: field_simps scaleR_conv_of_real) - also have "(\x. f (a + b * of_real x - a * of_real x)) = - (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" - using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) - also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ - ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms - by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) - also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def - by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) - finally show ?thesis by simp -qed - -lemma contour_integrable_linepath_Reals_iff: - fixes a b :: complex and f :: "complex \ complex" - assumes "a \ Reals" "b \ Reals" "Re a < Re b" - shows "(f contour_integrable_on linepath a b) \ - (\x. f (of_real x)) integrable_on {Re a..Re b}" - using has_contour_integral_linepath_Reals_iff[OF assms, of f] - by (auto simp: contour_integrable_on_def integrable_on_def) - -lemma contour_integral_linepath_Reals_eq: - fixes a b :: complex and f :: "complex \ complex" - assumes "a \ Reals" "b \ Reals" "Re a < Re b" - shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" -proof (cases "f contour_integrable_on linepath a b") - case True - thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] - using has_contour_integral_integral has_contour_integral_unique by blast -next - case False - thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] - by (simp add: not_integrable_contour_integral not_integrable_integral) -qed - - - -text\Cauchy's theorem where there's a primitive\ - -lemma contour_integral_primitive_lemma: - fixes f :: "complex \ complex" and g :: "real \ complex" - assumes "a \ b" - and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" - shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) - has_integral (f(g b) - f(g a))) {a..b}" -proof - - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" - using assms by (auto simp: piecewise_differentiable_on_def) - have cfg: "continuous_on {a..b} (\x. f (g x))" - apply (rule continuous_on_compose [OF cg, unfolded o_def]) - using assms - apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) - done - { fix x::real - assume a: "a < x" and b: "x < b" and xk: "x \ k" - then have "g differentiable at x within {a..b}" - using k by (simp add: differentiable_at_withinI) - then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" - by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) - then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" - by (simp add: has_vector_derivative_def scaleR_conv_of_real) - have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" - using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) - then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" - by (simp add: has_field_derivative_def) - have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" - using diff_chain_within [OF gdiff fdiff] - by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) - } note * = this - show ?thesis - apply (rule fundamental_theorem_of_calculus_interior_strong) - using k assms cfg * - apply (auto simp: at_within_Icc_at) - done -qed - -lemma contour_integral_primitive: - assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "valid_path g" "path_image g \ s" - shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" - using assms - apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) - apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) - done - -corollary Cauchy_theorem_primitive: - assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" - and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" - shows "(f' has_contour_integral 0) g" - using assms - by (metis diff_self contour_integral_primitive) - -text\Existence of path integral for continuous function\ -lemma contour_integrable_continuous_linepath: - assumes "continuous_on (closed_segment a b) f" - shows "f contour_integrable_on (linepath a b)" -proof - - have "continuous_on {0..1} ((\x. f x * (b - a)) \ linepath a b)" - apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) - apply (rule continuous_intros | simp add: assms)+ - done - then show ?thesis - apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) - apply (rule integrable_continuous [of 0 "1::real", simplified]) - apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) - apply (auto simp: vector_derivative_linepath_within) - done -qed - -lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" - by (rule has_derivative_imp_has_field_derivative) - (rule derivative_intros | simp)+ - -lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" - apply (rule contour_integral_unique) - using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] - apply (auto simp: field_simps has_field_der_id) - done - -lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" - by (simp add: contour_integrable_continuous_linepath) - -lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" - by (simp add: contour_integrable_continuous_linepath) - -subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ - -lemma has_contour_integral_neg: - "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" - by (simp add: has_integral_neg has_contour_integral_def) - -lemma has_contour_integral_add: - "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ - \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" - by (simp add: has_integral_add has_contour_integral_def algebra_simps) - -lemma has_contour_integral_diff: - "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ - \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" - by (simp add: has_integral_diff has_contour_integral_def algebra_simps) - -lemma has_contour_integral_lmul: - "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" -apply (simp add: has_contour_integral_def) -apply (drule has_integral_mult_right) -apply (simp add: algebra_simps) -done - -lemma has_contour_integral_rmul: - "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" -apply (drule has_contour_integral_lmul) -apply (simp add: mult.commute) -done - -lemma has_contour_integral_div: - "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" - by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) - -lemma has_contour_integral_eq: - "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" -apply (simp add: path_image_def has_contour_integral_def) -by (metis (no_types, lifting) image_eqI has_integral_eq) - -lemma has_contour_integral_bound_linepath: - assumes "(f has_contour_integral i) (linepath a b)" - "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" - shows "norm i \ B * norm(b - a)" -proof - - { fix x::real - assume x: "0 \ x" "x \ 1" - have "norm (f (linepath a b x)) * - norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" - by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) - } note * = this - have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" - apply (rule has_integral_bound - [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) - using assms * unfolding has_contour_integral_def - apply (auto simp: norm_mult) - done - then show ?thesis - by (auto simp: content_real) -qed - -(*UNUSED -lemma has_contour_integral_bound_linepath_strong: - fixes a :: real and f :: "complex \ real" - assumes "(f has_contour_integral i) (linepath a b)" - "finite k" - "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" - shows "norm i \ B*norm(b - a)" -*) - -lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" - unfolding has_contour_integral_linepath - by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) - -lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" - by (simp add: has_contour_integral_def) - -lemma has_contour_integral_is_0: - "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" - by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto - -lemma has_contour_integral_sum: - "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ - \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" - by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) - -subsection\<^marker>\tag unimportant\ \Operations on path integrals\ - -lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" - by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) - -lemma contour_integral_neg: - "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) - -lemma contour_integral_add: - "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = - contour_integral g f1 + contour_integral g f2" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) - -lemma contour_integral_diff: - "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = - contour_integral g f1 - contour_integral g f2" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) - -lemma contour_integral_lmul: - shows "f contour_integrable_on g - \ contour_integral g (\x. c * f x) = c*contour_integral g f" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) - -lemma contour_integral_rmul: - shows "f contour_integrable_on g - \ contour_integral g (\x. f x * c) = contour_integral g f * c" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) - -lemma contour_integral_div: - shows "f contour_integrable_on g - \ contour_integral g (\x. f x / c) = contour_integral g f / c" - by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) - -lemma contour_integral_eq: - "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" - apply (simp add: contour_integral_def) - using has_contour_integral_eq - by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) - -lemma contour_integral_eq_0: - "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" - by (simp add: has_contour_integral_is_0 contour_integral_unique) - -lemma contour_integral_bound_linepath: - shows - "\f contour_integrable_on (linepath a b); - 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ - \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" - apply (rule has_contour_integral_bound_linepath [of f]) - apply (auto simp: has_contour_integral_integral) - done - -lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" - by (simp add: contour_integral_unique has_contour_integral_0) - -lemma contour_integral_sum: - "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ - \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" - by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral) - -lemma contour_integrable_eq: - "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" - unfolding contour_integrable_on_def - by (metis has_contour_integral_eq) - - -subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ - -lemma contour_integrable_neg: - "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" - using has_contour_integral_neg contour_integrable_on_def by blast - -lemma contour_integrable_add: - "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" - using has_contour_integral_add contour_integrable_on_def - by fastforce - -lemma contour_integrable_diff: - "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" - using has_contour_integral_diff contour_integrable_on_def - by fastforce - -lemma contour_integrable_lmul: - "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" - using has_contour_integral_lmul contour_integrable_on_def - by fastforce - -lemma contour_integrable_rmul: - "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" - using has_contour_integral_rmul contour_integrable_on_def - by fastforce - -lemma contour_integrable_div: - "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" - using has_contour_integral_div contour_integrable_on_def - by fastforce - -lemma contour_integrable_sum: - "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ - \ (\x. sum (\a. f a x) s) contour_integrable_on p" - unfolding contour_integrable_on_def - by (metis has_contour_integral_sum) - - -subsection\<^marker>\tag unimportant\ \Reversing a path integral\ - -lemma has_contour_integral_reverse_linepath: - "(f has_contour_integral i) (linepath a b) - \ (f has_contour_integral (-i)) (linepath b a)" - using has_contour_integral_reversepath valid_path_linepath by fastforce - -lemma contour_integral_reverse_linepath: - "continuous_on (closed_segment a b) f - \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" -apply (rule contour_integral_unique) -apply (rule has_contour_integral_reverse_linepath) -by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) - - -(* Splitting a path integral in a flat way.*) - -lemma has_contour_integral_split: - assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "(f has_contour_integral (i + j)) (linepath a b)" -proof (cases "k = 0 \ k = 1") - case True - then show ?thesis - using assms by auto -next - case False - then have k: "0 < k" "k < 1" "complex_of_real k \ 1" - using assms by auto - have c': "c = k *\<^sub>R (b - a) + a" - by (metis diff_add_cancel c) - have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" - by (simp add: algebra_simps c') - { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" - have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" - using False apply (simp add: c' algebra_simps) - apply (simp add: real_vector.scale_left_distrib [symmetric] field_split_simps) - done - have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" - using k has_integral_affinity01 [OF *, of "inverse k" "0"] - apply (simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) - apply (auto dest: has_integral_cmul [where c = "inverse k"]) - done - } note fi = this - { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" - have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" - using k - apply (simp add: c' field_simps) - apply (simp add: scaleR_conv_of_real divide_simps) - apply (simp add: field_simps) - done - have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" - using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] - apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) - apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) - done - } note fj = this - show ?thesis - using f k - apply (simp add: has_contour_integral_linepath) - apply (simp add: linepath_def) - apply (rule has_integral_combine [OF _ _ fi fj], simp_all) - done -qed - -lemma continuous_on_closed_segment_transform: - assumes f: "continuous_on (closed_segment a b) f" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "continuous_on (closed_segment a c) f" -proof - - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" - using c by (simp add: algebra_simps) - have "closed_segment a c \ closed_segment a b" - by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) - then show "continuous_on (closed_segment a c) f" - by (rule continuous_on_subset [OF f]) -qed - -lemma contour_integral_split: - assumes f: "continuous_on (closed_segment a b) f" - and k: "0 \ k" "k \ 1" - and c: "c - a = k *\<^sub>R (b - a)" - shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" -proof - - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" - using c by (simp add: algebra_simps) - have "closed_segment a c \ closed_segment a b" - by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) - moreover have "closed_segment c b \ closed_segment a b" - by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) - ultimately - have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" - by (auto intro: continuous_on_subset [OF f]) - show ?thesis - by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) -qed - -lemma contour_integral_split_linepath: - assumes f: "continuous_on (closed_segment a b) f" - and c: "c \ closed_segment a b" - shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" - using c by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) - -text\The special case of midpoints used in the main quadrisection\ - -lemma has_contour_integral_midpoint: - assumes "(f has_contour_integral i) (linepath a (midpoint a b))" - "(f has_contour_integral j) (linepath (midpoint a b) b)" - shows "(f has_contour_integral (i + j)) (linepath a b)" - apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) - using assms - apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) - done - -lemma contour_integral_midpoint: - "continuous_on (closed_segment a b) f - \ contour_integral (linepath a b) f = - contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" - apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) - apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) - done - - -text\A couple of special case lemmas that are useful below\ - -lemma triangle_linear_has_chain_integral: - "((\x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" - apply (rule Cauchy_theorem_primitive [of UNIV "\x. m/2 * x^2 + d*x"]) - apply (auto intro!: derivative_eq_intros) - done - -lemma has_chain_integral_chain_integral3: - "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" - apply (subst contour_integral_unique [symmetric], assumption) - apply (drule has_contour_integral_integrable) - apply (simp add: valid_path_join) - done - -lemma has_chain_integral_chain_integral4: - "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" - apply (subst contour_integral_unique [symmetric], assumption) - apply (drule has_contour_integral_integrable) - apply (simp add: valid_path_join) - done - -subsection\Reversing the order in a double path integral\ - -text\The condition is stronger than needed but it's often true in typical situations\ - -lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" - by (auto simp: cbox_Pair_eq) - -lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" - by (auto simp: cbox_Pair_eq) - -proposition contour_integral_swap: - assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" - and vp: "valid_path g" "valid_path h" - and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" - and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" - shows "contour_integral g (\w. contour_integral h (f w)) = - contour_integral h (\z. contour_integral g (\w. f w z))" -proof - - have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" - using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) \ (\t. (g x, h t))" - by (rule ext) simp - have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) \ (\t. (g t, h x))" - by (rule ext) simp - have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" - by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) - have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" - by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) - have "\y. y \ {0..1} \ continuous_on {0..1} (\x. f (g x) (h y))" - by (subst fgh2) (rule fcon_im2 gcon continuous_intros | simp)+ - then have vdg: "\y. y \ {0..1} \ (\x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" - using continuous_on_mult gvcon integrable_continuous_real by blast - have "(\z. vector_derivative g (at (fst z))) = (\x. vector_derivative g (at x)) \ fst" - by auto - then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\x. vector_derivative g (at (fst x)))" - apply (rule ssubst) - apply (rule continuous_intros | simp add: gvcon)+ - done - have "(\z. vector_derivative h (at (snd z))) = (\x. vector_derivative h (at x)) \ snd" - by auto - then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" - apply (rule ssubst) - apply (rule continuous_intros | simp add: hvcon)+ - done - have "(\x. f (g (fst x)) (h (snd x))) = (\(y1,y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w))" - by auto - then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" - apply (rule ssubst) - apply (rule gcon hcon continuous_intros | simp)+ - apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) - done - have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = - integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" - proof (rule integral_cong [OF contour_integral_rmul [symmetric]]) - show "\x. x \ {0..1} \ f (g x) contour_integrable_on h" - unfolding contour_integrable_on - apply (rule integrable_continuous_real) - apply (rule continuous_on_mult [OF _ hvcon]) - apply (subst fgh1) - apply (rule fcon_im1 hcon continuous_intros | simp)+ - done - qed - also have "\ = integral {0..1} - (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" - unfolding contour_integral_integral - apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) - apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ - unfolding integral_mult_left [symmetric] - apply (simp only: mult_ac) - done - also have "\ = contour_integral h (\z. contour_integral g (\w. f w z))" - unfolding contour_integral_integral - apply (rule integral_cong) - unfolding integral_mult_left [symmetric] - apply (simp add: algebra_simps) - done - finally show ?thesis - by (simp add: contour_integral_integral) -qed - - -subsection\<^marker>\tag unimportant\ \The key quadrisection step\ - -lemma norm_sum_half: - assumes "norm(a + b) \ e" - shows "norm a \ e/2 \ norm b \ e/2" -proof - - have "e \ norm (- a - b)" - by (simp add: add.commute assms norm_minus_commute) - thus ?thesis - using norm_triangle_ineq4 order_trans by fastforce -qed - -lemma norm_sum_lemma: - assumes "e \ norm (a + b + c + d)" - shows "e / 4 \ norm a \ e / 4 \ norm b \ e / 4 \ norm c \ e / 4 \ norm d" -proof - - have "e \ norm ((a + b) + (c + d))" using assms - by (simp add: algebra_simps) - then show ?thesis - by (auto dest!: norm_sum_half) -qed - -lemma Cauchy_theorem_quadrisection: - assumes f: "continuous_on (convex hull {a,b,c}) f" - and dist: "dist a b \ K" "dist b c \ K" "dist c a \ K" - and e: "e * K^2 \ - norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" - shows "\a' b' c'. - a' \ convex hull {a,b,c} \ b' \ convex hull {a,b,c} \ c' \ convex hull {a,b,c} \ - dist a' b' \ K/2 \ dist b' c' \ K/2 \ dist c' a' \ K/2 \ - e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" - (is "\x y z. ?\ x y z") -proof - - note divide_le_eq_numeral1 [simp del] - define a' where "a' = midpoint b c" - define b' where "b' = midpoint c a" - define c' where "c' = midpoint a b" - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using f continuous_on_subset segments_subset_convex_hull by metis+ - have fcont': "continuous_on (closed_segment c' b') f" - "continuous_on (closed_segment a' c') f" - "continuous_on (closed_segment b' a') f" - unfolding a'_def b'_def c'_def - by (rule continuous_on_subset [OF f], - metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ - let ?pathint = "\x y. contour_integral(linepath x y) f" - have *: "?pathint a b + ?pathint b c + ?pathint c a = - (?pathint a c' + ?pathint c' b' + ?pathint b' a) + - (?pathint a' c' + ?pathint c' b + ?pathint b a') + - (?pathint a' c + ?pathint c b' + ?pathint b' a') + - (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" - by (simp add: fcont' contour_integral_reverse_linepath) (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) - have [simp]: "\x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" - by (metis left_diff_distrib mult.commute norm_mult_numeral1) - have [simp]: "\x y. cmod (x - y) = cmod (y - x)" - by (simp add: norm_minus_commute) - consider "e * K\<^sup>2 / 4 \ cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | - "e * K\<^sup>2 / 4 \ cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | - "e * K\<^sup>2 / 4 \ cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | - "e * K\<^sup>2 / 4 \ cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" - using assms unfolding * by (blast intro: that dest!: norm_sum_lemma) - then show ?thesis - proof cases - case 1 then have "?\ a c' b'" - using assms - apply (clarsimp simp: c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) - done - then show ?thesis by blast - next - case 2 then have "?\ a' c' b" - using assms - apply (clarsimp simp: a'_def c'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) - done - then show ?thesis by blast - next - case 3 then have "?\ a' c b'" - using assms - apply (clarsimp simp: a'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) - done - then show ?thesis by blast - next - case 4 then have "?\ a' b' c'" - using assms - apply (clarsimp simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) - apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) - done - then show ?thesis by blast - qed -qed - -subsection\<^marker>\tag unimportant\ \Cauchy's theorem for triangles\ - -lemma triangle_points_closer: - fixes a::complex - shows "\x \ convex hull {a,b,c}; y \ convex hull {a,b,c}\ - \ norm(x - y) \ norm(a - b) \ - norm(x - y) \ norm(b - c) \ - norm(x - y) \ norm(c - a)" - using simplex_extremal_le [of "{a,b,c}"] - by (auto simp: norm_minus_commute) - -lemma holomorphic_point_small_triangle: - assumes x: "x \ S" - and f: "continuous_on S f" - and cd: "f field_differentiable (at x within S)" - and e: "0 < e" - shows "\k>0. \a b c. dist a b \ k \ dist b c \ k \ dist c a \ k \ - x \ convex hull {a,b,c} \ convex hull {a,b,c} \ S - \ norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + - contour_integral(linepath c a) f) - \ e*(dist a b + dist b c + dist c a)^2" - (is "\k>0. \a b c. _ \ ?normle a b c") -proof - - have le_of_3: "\a x y z. \0 \ x*y; 0 \ x*z; 0 \ y*z; a \ (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\ - \ a \ e*(x + y + z)^2" - by (simp add: algebra_simps power2_eq_square) - have disj_le: "\x \ a \ x \ b \ x \ c; 0 \ a; 0 \ b; 0 \ c\ \ x \ a + b + c" - for x::real and a b c - by linarith - have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" - if "convex hull {a, b, c} \ S" for a b c - using segments_subset_convex_hull that - by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ - note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] - { fix f' a b c d - assume d: "0 < d" - and f': "\y. \cmod (y - x) \ d; y \ S\ \ cmod (f y - f x - f' * (y - x)) \ e * cmod (y - x)" - and le: "cmod (a - b) \ d" "cmod (b - c) \ d" "cmod (c - a) \ d" - and xc: "x \ convex hull {a, b, c}" - and S: "convex hull {a, b, c} \ S" - have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = - contour_integral (linepath a b) (\y. f y - f x - f'*(y - x)) + - contour_integral (linepath b c) (\y. f y - f x - f'*(y - x)) + - contour_integral (linepath c a) (\y. f y - f x - f'*(y - x))" - apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF S]) - apply (simp add: field_simps) - done - { fix y - assume yc: "y \ convex hull {a,b,c}" - have "cmod (f y - f x - f' * (y - x)) \ e*norm(y - x)" - proof (rule f') - show "cmod (y - x) \ d" - by (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) - qed (use S yc in blast) - also have "\ \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" - by (simp add: yc e xc disj_le [OF triangle_points_closer]) - finally have "cmod (f y - f x - f' * (y - x)) \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . - } note cm_le = this - have "?normle a b c" - unfolding dist_norm pa - apply (rule le_of_3) - using f' xc S e - apply simp_all - apply (intro norm_triangle_le add_mono path_bound) - apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) - apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ - done - } note * = this - show ?thesis - using cd e - apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) - apply (clarify dest!: spec mp) - using * unfolding dist_norm - apply blast - done -qed - - -text\Hence the most basic theorem for a triangle.\ - -locale Chain = - fixes x0 At Follows - assumes At0: "At x0 0" - and AtSuc: "\x n. At x n \ \x'. At x' (Suc n) \ Follows x' x" -begin - primrec f where - "f 0 = x0" - | "f (Suc n) = (SOME x. At x (Suc n) \ Follows x (f n))" - - lemma At: "At (f n) n" - proof (induct n) - case 0 show ?case - by (simp add: At0) - next - case (Suc n) show ?case - by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) - qed - - lemma Follows: "Follows (f(Suc n)) (f n)" - by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) - - declare f.simps(2) [simp del] -end - -lemma Chain3: - assumes At0: "At x0 y0 z0 0" - and AtSuc: "\x y z n. At x y z n \ \x' y' z'. At x' y' z' (Suc n) \ Follows x' y' z' x y z" - obtains f g h where - "f 0 = x0" "g 0 = y0" "h 0 = z0" - "\n. At (f n) (g n) (h n) n" - "\n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" -proof - - interpret three: Chain "(x0,y0,z0)" "\(x,y,z). At x y z" "\(x',y',z'). \(x,y,z). Follows x' y' z' x y z" - apply unfold_locales - using At0 AtSuc by auto - show ?thesis - apply (rule that [of "\n. fst (three.f n)" "\n. fst (snd (three.f n))" "\n. snd (snd (three.f n))"]) - using three.At three.Follows - apply simp_all - apply (simp_all add: split_beta') - done -qed - -proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle: - assumes "f holomorphic_on (convex hull {a,b,c})" - shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" -proof - - have contf: "continuous_on (convex hull {a,b,c}) f" - by (metis assms holomorphic_on_imp_continuous_on) - let ?pathint = "\x y. contour_integral(linepath x y) f" - { fix y::complex - assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" - and ynz: "y \ 0" - define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" - define e where "e = norm y / K^2" - have K1: "K \ 1" by (simp add: K_def max.coboundedI1) - then have K: "K > 0" by linarith - have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" - by (simp_all add: K_def) - have e: "e > 0" - unfolding e_def using ynz K1 by simp - define At where "At x y z n \ - convex hull {x,y,z} \ convex hull {a,b,c} \ - dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ - norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" - for x y z n - have At0: "At a b c 0" - using fy - by (simp add: At_def e_def has_chain_integral_chain_integral3) - { fix x y z n - assume At: "At x y z n" - then have contf': "continuous_on (convex hull {x,y,z}) f" - using contf At_def continuous_on_subset by metis - have "\x' y' z'. At x' y' z' (Suc n) \ convex hull {x',y',z'} \ convex hull {x,y,z}" - using At Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] - apply (simp add: At_def algebra_simps) - apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) - done - } note AtSuc = this - obtain fa fb fc - where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" - and cosb: "\n. convex hull {fa n, fb n, fc n} \ convex hull {a,b,c}" - and dist: "\n. dist (fa n) (fb n) \ K/2^n" - "\n. dist (fb n) (fc n) \ K/2^n" - "\n. dist (fc n) (fa n) \ K/2^n" - and no: "\n. norm(?pathint (fa n) (fb n) + - ?pathint (fb n) (fc n) + - ?pathint (fc n) (fa n)) \ e * (K/2^n)^2" - and conv_le: "\n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \ convex hull {fa n, fb n, fc n}" - apply (rule Chain3 [of At, OF At0 AtSuc]) - apply (auto simp: At_def) - done - obtain x where x: "\n. x \ convex hull {fa n, fb n, fc n}" - proof (rule bounded_closed_nest) - show "\n. closed (convex hull {fa n, fb n, fc n})" - by (simp add: compact_imp_closed finite_imp_compact_convex_hull) - show "\m n. m \ n \ convex hull {fa n, fb n, fc n} \ convex hull {fa m, fb m, fc m}" - by (erule transitive_stepwise_le) (auto simp: conv_le) - qed (fastforce intro: finite_imp_bounded_convex_hull)+ - then have xin: "x \ convex hull {a,b,c}" - using assms f0 by blast - then have fx: "f field_differentiable at x within (convex hull {a,b,c})" - using assms holomorphic_on_def by blast - { fix k n - assume k: "0 < k" - and le: - "\x' y' z'. - \dist x' y' \ k; dist y' z' \ k; dist z' x' \ k; - x \ convex hull {x',y',z'}; - convex hull {x',y',z'} \ convex hull {a,b,c}\ - \ - cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 - \ e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" - and Kk: "K / k < 2 ^ n" - have "K / 2 ^ n < k" using Kk k - by (auto simp: field_simps) - then have DD: "dist (fa n) (fb n) \ k" "dist (fb n) (fc n) \ k" "dist (fc n) (fa n) \ k" - using dist [of n] k - by linarith+ - have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 - \ (3 * K / 2 ^ n)\<^sup>2" - using dist [of n] e K - by (simp add: abs_le_square_iff [symmetric]) - have less10: "\x y::real. 0 < x \ y \ 9*x \ y < x*10" - by linarith - have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ e * (3 * K / 2 ^ n)\<^sup>2" - using ynz dle e mult_le_cancel_left_pos by blast - also have "\ < - cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" - using no [of n] e K - apply (simp add: e_def field_simps) - apply (simp only: zero_less_norm_iff [symmetric]) - done - finally have False - using le [OF DD x cosb] by auto - } then - have ?thesis - using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e - apply clarsimp - apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]], force+) - done - } - moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" - by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) - segments_subset_convex_hull(3) segments_subset_convex_hull(5)) - ultimately show ?thesis - using has_contour_integral_integral by fastforce -qed - -subsection\<^marker>\tag unimportant\ \Version needing function holomorphic in interior only\ - -lemma Cauchy_theorem_flat_lemma: - assumes f: "continuous_on (convex hull {a,b,c}) f" - and c: "c - a = k *\<^sub>R (b - a)" - and k: "0 \ k" - shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof - - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using f continuous_on_subset segments_subset_convex_hull by metis+ - show ?thesis - proof (cases "k \ 1") - case True show ?thesis - by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) - next - case False then show ?thesis - using fabc c - apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) - apply (metis closed_segment_commute fabc(3)) - apply (auto simp: k contour_integral_reverse_linepath) - done - qed -qed - -lemma Cauchy_theorem_flat: - assumes f: "continuous_on (convex hull {a,b,c}) f" - and c: "c - a = k *\<^sub>R (b - a)" - shows "contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof (cases "0 \ k") - case True with assms show ?thesis - by (blast intro: Cauchy_theorem_flat_lemma) -next - case False - have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using f continuous_on_subset segments_subset_convex_hull by metis+ - moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + - contour_integral (linepath c b) f = 0" - apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) - using False c - apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) - done - ultimately show ?thesis - apply (auto simp: contour_integral_reverse_linepath) - using add_eq_0_iff by force -qed - -lemma Cauchy_theorem_triangle_interior: - assumes contf: "continuous_on (convex hull {a,b,c}) f" - and holf: "f holomorphic_on interior (convex hull {a,b,c})" - shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" -proof - - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" - using contf continuous_on_subset segments_subset_convex_hull by metis+ - have "bounded (f ` (convex hull {a,b,c}))" - by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) - then obtain B where "0 < B" and Bnf: "\x. x \ convex hull {a,b,c} \ norm (f x) \ B" - by (auto simp: dest!: bounded_pos [THEN iffD1]) - have "bounded (convex hull {a,b,c})" - by (simp add: bounded_convex_hull) - then obtain C where C: "0 < C" and Cno: "\y. y \ convex hull {a,b,c} \ norm y < C" - using bounded_pos_less by blast - then have diff_2C: "norm(x - y) \ 2*C" - if x: "x \ convex hull {a, b, c}" and y: "y \ convex hull {a, b, c}" for x y - proof - - have "cmod x \ C" - using x by (meson Cno not_le not_less_iff_gr_or_eq) - hence "cmod (x - y) \ C + C" - using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) - thus "cmod (x - y) \ 2 * C" - by (metis mult_2) - qed - have contf': "continuous_on (convex hull {b,a,c}) f" - using contf by (simp add: insert_commute) - { fix y::complex - assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" - and ynz: "y \ 0" - have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" - by (rule has_chain_integral_chain_integral3 [OF fy]) - have ?thesis - proof (cases "c=a \ a=b \ b=c") - case True then show ?thesis - using Cauchy_theorem_flat [OF contf, of 0] - using has_chain_integral_chain_integral3 [OF fy] ynz - by (force simp: fabc contour_integral_reverse_linepath) - next - case False - then have car3: "card {a, b, c} = Suc (DIM(complex))" - by auto - { assume "interior(convex hull {a,b,c}) = {}" - then have "collinear{a,b,c}" - using interior_convex_hull_eq_empty [OF car3] - by (simp add: collinear_3_eq_affine_dependent) - with False obtain d where "c \ a" "a \ b" "b \ c" "c - b = d *\<^sub>R (a - b)" - by (auto simp: collinear_3 collinear_lemma) - then have "False" - using False Cauchy_theorem_flat [OF contf'] pi_eq_y ynz - by (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) - } - then obtain d where d: "d \ interior (convex hull {a, b, c})" - by blast - { fix d1 - assume d1_pos: "0 < d1" - and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ - \ cmod (f x' - f x) < cmod y / (24 * C)" - define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" - define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x - let ?pathint = "\x y. contour_integral(linepath x y) f" - have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" - using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) - then have eCB: "24 * e * C * B \ cmod y" - using \C>0\ \B>0\ by (simp add: field_simps) - have e_le_d1: "e * (4 * C) \ d1" - using e \C>0\ by (simp add: field_simps) - have "shrink a \ interior(convex hull {a,b,c})" - "shrink b \ interior(convex hull {a,b,c})" - "shrink c \ interior(convex hull {a,b,c})" - using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) - then have fhp0: "(f has_contour_integral 0) - (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" - by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal) - then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" - by (simp add: has_chain_integral_chain_integral3) - have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" - "f contour_integrable_on linepath (shrink b) (shrink c)" - "f contour_integrable_on linepath (shrink c) (shrink a)" - using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) - have cmod_shr: "\x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" - using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) - have sh_eq: "\a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" - by (simp add: algebra_simps) - have "cmod y / (24 * C) \ cmod y / cmod (b - a) / 12" - using False \C>0\ diff_2C [of b a] ynz - by (auto simp: field_split_simps hull_inc) - have less_C: "\u \ convex hull {a, b, c}; 0 \ x; x \ 1\ \ x * cmod u < C" for x u - apply (cases "x=0", simp add: \0) - using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast - { fix u v - assume uv: "u \ convex hull {a, b, c}" "v \ convex hull {a, b, c}" "u\v" - and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" - have shr_uv: "shrink u \ interior(convex hull {a,b,c})" - "shrink v \ interior(convex hull {a,b,c})" - using d e uv - by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) - have cmod_fuv: "\x. 0\x \ x\1 \ cmod (f (linepath (shrink u) (shrink v) x)) \ B" - using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) - have By_uv: "B * (12 * (e * cmod (u - v))) \ cmod y" - apply (rule order_trans [OF _ eCB]) - using e \B>0\ diff_2C [of u v] uv - by (auto simp: field_simps) - { fix x::real assume x: "0\x" "x\1" - have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" - apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) - using uv x d interior_subset - apply (auto simp: hull_inc intro!: less_C) - done - have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" - by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) - have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" - apply (simp only: ll norm_mult scaleR_diff_right) - using \e>0\ cmod_less_4C apply (force intro: norm_triangle_lt less_le_trans [OF _ e_le_d1]) - done - have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" - using x uv shr_uv cmod_less_dt - by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) - also have "\ \ cmod y / cmod (v - u) / 12" - using False uv \C>0\ diff_2C [of v u] ynz - by (auto simp: field_split_simps hull_inc) - finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / cmod (v - u) / 12" - by simp - then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \ cmod y" - using uv False by (auto simp: field_simps) - have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + - cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) - \ B * (cmod y / 24 / C / B * 2 * C) + 2 * C * (cmod y / 24 / C)" - apply (rule add_mono [OF mult_mono]) - using By_uv e \0 < B\ \0 < C\ x apply (simp_all add: cmod_fuv cmod_shr cmod_12_le) - apply (simp add: field_simps) - done - also have "\ \ cmod y / 6" - by simp - finally have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + - cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) - \ cmod y / 6" . - } note cmod_diff_le = this - have f_uv: "continuous_on (closed_segment u v) f" - by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) - have **: "\f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" - by (simp add: algebra_simps) - have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) - \ (B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)) * content (cbox 0 (1::real))" - apply (rule has_integral_bound - [of _ "\x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" - _ 0 1]) - using ynz \0 < B\ \0 < C\ - apply (simp_all del: le_divide_eq_numeral1) - apply (simp add: has_integral_diff has_contour_integral_linepath [symmetric] has_contour_integral_integral - fpi_uv f_uv contour_integrable_continuous_linepath) - apply (auto simp: ** norm_triangle_le norm_mult cmod_diff_le simp del: le_divide_eq_numeral1) - done - also have "\ \ norm y / 6" - by simp - finally have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ norm y / 6" . - } note * = this - have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \ norm y / 6" - using False fpi_abc by (rule_tac *) (auto simp: hull_inc) - moreover - have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \ norm y / 6" - using False fpi_abc by (rule_tac *) (auto simp: hull_inc) - moreover - have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \ norm y / 6" - using False fpi_abc by (rule_tac *) (auto simp: hull_inc) - ultimately - have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + - (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) - \ norm y / 6 + norm y / 6 + norm y / 6" - by (metis norm_triangle_le add_mono) - also have "\ = norm y / 2" - by simp - finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - - (?pathint a b + ?pathint b c + ?pathint c a)) - \ norm y / 2" - by (simp add: algebra_simps) - then - have "norm(?pathint a b + ?pathint b c + ?pathint c a) \ norm y / 2" - by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) - then have "False" - using pi_eq_y ynz by auto - } - moreover have "uniformly_continuous_on (convex hull {a,b,c}) f" - by (simp add: contf compact_convex_hull compact_uniformly_continuous) - ultimately have "False" - unfolding uniformly_continuous_on_def - by (force simp: ynz \0 < C\ dist_norm) - then show ?thesis .. - qed - } - moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" - using fabc contour_integrable_continuous_linepath by auto - ultimately show ?thesis - using has_contour_integral_integral by fastforce -qed - -subsection\<^marker>\tag unimportant\ \Version allowing finite number of exceptional points\ - -proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle_cofinite: - assumes "continuous_on (convex hull {a,b,c}) f" - and "finite S" - and "(\x. x \ interior(convex hull {a,b,c}) - S \ f field_differentiable (at x))" - shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" -using assms -proof (induction "card S" arbitrary: a b c S rule: less_induct) - case (less S a b c) - show ?case - proof (cases "S={}") - case True with less show ?thesis - by (fastforce simp: holomorphic_on_def field_differentiable_at_within Cauchy_theorem_triangle_interior) - next - case False - then obtain d S' where d: "S = insert d S'" "d \ S'" - by (meson Set.set_insert all_not_in_conv) - then show ?thesis - proof (cases "d \ convex hull {a,b,c}") - case False - show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" - proof (rule less.hyps) - show "\x. x \ interior (convex hull {a, b, c}) - S' \ f field_differentiable at x" - using False d interior_subset by (auto intro!: less.prems) - qed (use d less.prems in auto) - next - case True - have *: "convex hull {a, b, d} \ convex hull {a, b, c}" - by (meson True hull_subset insert_subset convex_hull_subset) - have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)" - proof (rule less.hyps) - show "\x. x \ interior (convex hull {a, b, d}) - S' \ f field_differentiable at x" - using d not_in_interior_convex_hull_3 - by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) - qed (use d continuous_on_subset [OF _ *] less.prems in auto) - have *: "convex hull {b, c, d} \ convex hull {a, b, c}" - by (meson True hull_subset insert_subset convex_hull_subset) - have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)" - proof (rule less.hyps) - show "\x. x \ interior (convex hull {b, c, d}) - S' \ f field_differentiable at x" - using d not_in_interior_convex_hull_3 - by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) - qed (use d continuous_on_subset [OF _ *] less.prems in auto) - have *: "convex hull {c, a, d} \ convex hull {a, b, c}" - by (meson True hull_subset insert_subset convex_hull_subset) - have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)" - proof (rule less.hyps) - show "\x. x \ interior (convex hull {c, a, d}) - S' \ f field_differentiable at x" - using d not_in_interior_convex_hull_3 - by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) - qed (use d continuous_on_subset [OF _ *] less.prems in auto) - have "f contour_integrable_on linepath a b" - using less.prems abd contour_integrable_joinD1 contour_integrable_on_def by blast - moreover have "f contour_integrable_on linepath b c" - using less.prems bcd contour_integrable_joinD1 contour_integrable_on_def by blast - moreover have "f contour_integrable_on linepath c a" - using less.prems cad contour_integrable_joinD1 contour_integrable_on_def by blast - ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" - by auto - { fix y::complex - assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" - and ynz: "y \ 0" - have cont_ad: "continuous_on (closed_segment a d) f" - by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3)) - have cont_bd: "continuous_on (closed_segment b d) f" - by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1)) - have cont_cd: "continuous_on (closed_segment c d) f" - by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2)) - have "contour_integral (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))" - "contour_integral (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))" - "contour_integral (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)" - using has_chain_integral_chain_integral3 [OF abd] - has_chain_integral_chain_integral3 [OF bcd] - has_chain_integral_chain_integral3 [OF cad] - by (simp_all add: algebra_simps add_eq_0_iff) - then have ?thesis - using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce - } - then show ?thesis - using fpi contour_integrable_on_def by blast - qed - qed -qed - -subsection\<^marker>\tag unimportant\ \Cauchy's theorem for an open starlike set\ - -lemma starlike_convex_subset: - assumes S: "a \ S" "closed_segment b c \ S" and subs: "\x. x \ S \ closed_segment a x \ S" - shows "convex hull {a,b,c} \ S" - using S - apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull) - apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE) - done - -lemma triangle_contour_integrals_starlike_primitive: - assumes contf: "continuous_on S f" - and S: "a \ S" "open S" - and x: "x \ S" - and subs: "\y. y \ S \ closed_segment a y \ S" - and zer: "\b c. closed_segment b c \ S - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)" -proof - - let ?pathint = "\x y. contour_integral(linepath x y) f" - { fix e y - assume e: "0 < e" and bxe: "ball x e \ S" and close: "cmod (y - x) < e" - have y: "y \ S" - using bxe close by (force simp: dist_norm norm_minus_commute) - have cont_ayf: "continuous_on (closed_segment a y) f" - using contf continuous_on_subset subs y by blast - have xys: "closed_segment x y \ S" - apply (rule order_trans [OF _ bxe]) - using close - by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound) - have "?pathint a y - ?pathint a x = ?pathint x y" - using zer [OF xys] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force - } note [simp] = this - { fix e::real - assume e: "0 < e" - have cont_atx: "continuous (at x) f" - using x S contf continuous_on_eq_continuous_at by blast - then obtain d1 where d1: "d1>0" and d1_less: "\y. cmod (y - x) < d1 \ cmod (f y - f x) < e/2" - unfolding continuous_at Lim_at dist_norm using e - by (drule_tac x="e/2" in spec) force - obtain d2 where d2: "d2>0" "ball x d2 \ S" using \open S\ x - by (auto simp: open_contains_ball) - have dpos: "min d1 d2 > 0" using d1 d2 by simp - { fix y - assume yx: "y \ x" and close: "cmod (y - x) < min d1 d2" - have y: "y \ S" - using d2 close by (force simp: dist_norm norm_minus_commute) - have "closed_segment x y \ S" - using close d2 by (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1)) - then have fxy: "f contour_integrable_on linepath x y" - by (metis contour_integrable_continuous_linepath continuous_on_subset [OF contf]) - then obtain i where i: "(f has_contour_integral i) (linepath x y)" - by (auto simp: contour_integrable_on_def) - then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" - by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) - then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" - proof (rule has_contour_integral_bound_linepath) - show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" - by (meson close d1_less le_less_trans less_imp_le min.strict_boundedE segment_bound1) - qed (use e in simp) - also have "\ < e * cmod (y - x)" - by (simp add: e yx) - finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using i yx by (simp add: contour_integral_unique divide_less_eq) - } - then have "\d>0. \y. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using dpos by blast - } - then have *: "(\y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \x\ 0" - by (simp add: Lim_at dist_norm inverse_eq_divide) - show ?thesis - apply (simp add: has_field_derivative_def has_derivative_at2 bounded_linear_mult_right) - apply (rule Lim_transform [OF * tendsto_eventually]) - using \open S\ x apply (force simp: dist_norm open_contains_ball inverse_eq_divide [symmetric] eventually_at) - done -qed - -(** Existence of a primitive.*) -lemma holomorphic_starlike_primitive: - fixes f :: "complex \ complex" - assumes contf: "continuous_on S f" - and S: "starlike S" and os: "open S" - and k: "finite k" - and fcd: "\x. x \ S - k \ f field_differentiable at x" - shows "\g. \x \ S. (g has_field_derivative f x) (at x)" -proof - - obtain a where a: "a\S" and a_cs: "\x. x\S \ closed_segment a x \ S" - using S by (auto simp: starlike_def) - { fix x b c - assume "x \ S" "closed_segment b c \ S" - then have abcs: "convex hull {a, b, c} \ S" - by (simp add: a a_cs starlike_convex_subset) - then have "continuous_on (convex hull {a, b, c}) f" - by (simp add: continuous_on_subset [OF contf]) - then have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" - using abcs interior_subset by (force intro: fcd Cauchy_theorem_triangle_cofinite [OF _ k]) - } note 0 = this - show ?thesis - apply (intro exI ballI) - apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption) - apply (metis a_cs) - apply (metis has_chain_integral_chain_integral3 0) - done -qed - -lemma Cauchy_theorem_starlike: - "\open S; starlike S; finite k; continuous_on S f; - \x. x \ S - k \ f field_differentiable at x; - valid_path g; path_image g \ S; pathfinish g = pathstart g\ - \ (f has_contour_integral 0) g" - by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open) - -lemma Cauchy_theorem_starlike_simple: - "\open S; starlike S; f holomorphic_on S; valid_path g; path_image g \ S; pathfinish g = pathstart g\ - \ (f has_contour_integral 0) g" -apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI]) -apply (simp_all add: holomorphic_on_imp_continuous_on) -apply (metis at_within_open holomorphic_on_def) -done - -subsection\Cauchy's theorem for a convex set\ - -text\For a convex set we can avoid assuming openness and boundary analyticity\ - -lemma triangle_contour_integrals_convex_primitive: - assumes contf: "continuous_on S f" - and S: "a \ S" "convex S" - and x: "x \ S" - and zer: "\b c. \b \ S; c \ S\ - \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within S)" -proof - - let ?pathint = "\x y. contour_integral(linepath x y) f" - { fix y - assume y: "y \ S" - have cont_ayf: "continuous_on (closed_segment a y) f" - using S y by (meson contf continuous_on_subset convex_contains_segment) - have xys: "closed_segment x y \ S" (*?*) - using convex_contains_segment S x y by auto - have "?pathint a y - ?pathint a x = ?pathint x y" - using zer [OF x y] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force - } note [simp] = this - { fix e::real - assume e: "0 < e" - have cont_atx: "continuous (at x within S) f" - using x S contf by (simp add: continuous_on_eq_continuous_within) - then obtain d1 where d1: "d1>0" and d1_less: "\y. \y \ S; cmod (y - x) < d1\ \ cmod (f y - f x) < e/2" - unfolding continuous_within Lim_within dist_norm using e - by (drule_tac x="e/2" in spec) force - { fix y - assume yx: "y \ x" and close: "cmod (y - x) < d1" and y: "y \ S" - have fxy: "f contour_integrable_on linepath x y" - using convex_contains_segment S x y - by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) - then obtain i where i: "(f has_contour_integral i) (linepath x y)" - by (auto simp: contour_integrable_on_def) - then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" - by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) - then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" - proof (rule has_contour_integral_bound_linepath) - show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" - by (meson assms(3) close convex_contains_segment d1_less le_less_trans less_imp_le segment_bound1 subset_iff x y) - qed (use e in simp) - also have "\ < e * cmod (y - x)" - by (simp add: e yx) - finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using i yx by (simp add: contour_integral_unique divide_less_eq) - } - then have "\d>0. \y\S. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" - using d1 by blast - } - then have *: "((\y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \ 0) (at x within S)" - by (simp add: Lim_within dist_norm inverse_eq_divide) - show ?thesis - apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) - apply (rule Lim_transform [OF * tendsto_eventually]) - using linordered_field_no_ub - apply (force simp: inverse_eq_divide [symmetric] eventually_at) - done -qed - -lemma contour_integral_convex_primitive: - assumes "convex S" "continuous_on S f" - "\a b c. \a \ S; b \ S; c \ S\ \ (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" - obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" -proof (cases "S={}") - case False - with assms that show ?thesis - by (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3) -qed auto - -lemma holomorphic_convex_primitive: - fixes f :: "complex \ complex" - assumes "convex S" "finite K" and contf: "continuous_on S f" - and fd: "\x. x \ interior S - K \ f field_differentiable at x" - obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" -proof (rule contour_integral_convex_primitive [OF \convex S\ contf Cauchy_theorem_triangle_cofinite]) - have *: "convex hull {a, b, c} \ S" if "a \ S" "b \ S" "c \ S" for a b c - by (simp add: \convex S\ hull_minimal that) - show "continuous_on (convex hull {a, b, c}) f" if "a \ S" "b \ S" "c \ S" for a b c - by (meson "*" contf continuous_on_subset that) - show "f field_differentiable at x" if "a \ S" "b \ S" "c \ S" "x \ interior (convex hull {a, b, c}) - K" for a b c x - by (metis "*" DiffD1 DiffD2 DiffI fd interior_mono subsetCE that) -qed (use assms in \force+\) - -lemma holomorphic_convex_primitive': - fixes f :: "complex \ complex" - assumes "convex S" and "open S" and "f holomorphic_on S" - obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" -proof (rule holomorphic_convex_primitive) - fix x assume "x \ interior S - {}" - with assms show "f field_differentiable at x" - by (auto intro!: holomorphic_on_imp_differentiable_at simp: interior_open) -qed (use assms in \auto intro: holomorphic_on_imp_continuous_on\) - -corollary\<^marker>\tag unimportant\ Cauchy_theorem_convex: - "\continuous_on S f; convex S; finite K; - \x. x \ interior S - K \ f field_differentiable at x; - valid_path g; path_image g \ S; pathfinish g = pathstart g\ - \ (f has_contour_integral 0) g" - by (metis holomorphic_convex_primitive Cauchy_theorem_primitive) - -corollary Cauchy_theorem_convex_simple: - "\f holomorphic_on S; convex S; - valid_path g; path_image g \ S; - pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" - apply (rule Cauchy_theorem_convex [where K = "{}"]) - apply (simp_all add: holomorphic_on_imp_continuous_on) - using at_within_interior holomorphic_on_def interior_subset by fastforce - -text\In particular for a disc\ -corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc: - "\finite K; continuous_on (cball a e) f; - \x. x \ ball a e - K \ f field_differentiable at x; - valid_path g; path_image g \ cball a e; - pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" - by (auto intro: Cauchy_theorem_convex) - -corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc_simple: - "\f holomorphic_on (ball a e); valid_path g; path_image g \ ball a e; - pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" -by (simp add: Cauchy_theorem_convex_simple) - -subsection\<^marker>\tag unimportant\ \Generalize integrability to local primitives\ - -lemma contour_integral_local_primitive_lemma: - fixes f :: "complex\complex" - shows - "\g piecewise_differentiable_on {a..b}; - \x. x \ s \ (f has_field_derivative f' x) (at x within s); - \x. x \ {a..b} \ g x \ s\ - \ (\x. f' (g x) * vector_derivative g (at x within {a..b})) - integrable_on {a..b}" - apply (cases "cbox a b = {}", force) - apply (simp add: integrable_on_def) - apply (rule exI) - apply (rule contour_integral_primitive_lemma, assumption+) - using atLeastAtMost_iff by blast - -lemma contour_integral_local_primitive_any: - fixes f :: "complex \ complex" - assumes gpd: "g piecewise_differentiable_on {a..b}" - and dh: "\x. x \ s - \ \d h. 0 < d \ - (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" - and gs: "\x. x \ {a..b} \ g x \ s" - shows "(\x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}" -proof - - { fix x - assume x: "a \ x" "x \ b" - obtain d h where d: "0 < d" - and h: "(\y. norm(y - g x) < d \ (h has_field_derivative f y) (at y within s))" - using x gs dh by (metis atLeastAtMost_iff) - have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast - then obtain e where e: "e>0" and lessd: "\x'. x' \ {a..b} \ \x' - x\ < e \ cmod (g x' - g x) < d" - using x d - apply (auto simp: dist_norm continuous_on_iff) - apply (drule_tac x=x in bspec) - using x apply simp - apply (drule_tac x=d in spec, auto) - done - have "\d>0. \u v. u \ x \ x \ v \ {u..v} \ ball x d \ (u \ v \ a \ u \ v \ b) \ - (\x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}" - apply (rule_tac x=e in exI) - using e - apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify) - apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma) - apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset) - apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force) - done - } then - show ?thesis - by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified]) -qed - -lemma contour_integral_local_primitive: - fixes f :: "complex \ complex" - assumes g: "valid_path g" "path_image g \ s" - and dh: "\x. x \ s - \ \d h. 0 < d \ - (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" - shows "f contour_integrable_on g" - using g - apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def - has_integral_localized_vector_derivative integrable_on_def [symmetric]) - using contour_integral_local_primitive_any [OF _ dh] - by (meson image_subset_iff piecewise_C1_imp_differentiable) - - -text\In particular if a function is holomorphic\ - -lemma contour_integrable_holomorphic: - assumes contf: "continuous_on s f" - and os: "open s" - and k: "finite k" - and g: "valid_path g" "path_image g \ s" - and fcd: "\x. x \ s - k \ f field_differentiable at x" - shows "f contour_integrable_on g" -proof - - { fix z - assume z: "z \ s" - obtain d where "d>0" and d: "ball z d \ s" using \open s\ z - by (auto simp: open_contains_ball) - then have contfb: "continuous_on (ball z d) f" - using contf continuous_on_subset by blast - obtain h where "\y\ball z d. (h has_field_derivative f y) (at y within ball z d)" - by (metis holomorphic_convex_primitive [OF convex_ball k contfb fcd] d interior_subset Diff_iff subsetD) - then have "\y\ball z d. (h has_field_derivative f y) (at y within s)" - by (metis open_ball at_within_open d os subsetCE) - then have "\h. (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" - by (force simp: dist_norm norm_minus_commute) - then have "\d h. 0 < d \ (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" - using \0 < d\ by blast - } - then show ?thesis - by (rule contour_integral_local_primitive [OF g]) -qed - -lemma contour_integrable_holomorphic_simple: - assumes fh: "f holomorphic_on S" - and os: "open S" - and g: "valid_path g" "path_image g \ S" - shows "f contour_integrable_on g" - apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g]) - apply (simp add: fh holomorphic_on_imp_continuous_on) - using fh by (simp add: field_differentiable_def holomorphic_on_open os) - -lemma continuous_on_inversediff: - fixes z:: "'a::real_normed_field" shows "z \ S \ continuous_on S (\w. 1 / (w - z))" - by (rule continuous_intros | force)+ - -lemma contour_integrable_inversediff: - "\valid_path g; z \ path_image g\ \ (\w. 1 / (w-z)) contour_integrable_on g" -apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"]) -apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros) -done - -text\Key fact that path integral is the same for a "nearby" path. This is the - main lemma for the homotopy form of Cauchy's theorem and is also useful - if we want "without loss of generality" to assume some nice properties of a - path (e.g. smoothness). It can also be used to define the integrals of - analytic functions over arbitrary continuous paths. This is just done for - winding numbers now. -\ - -text\A technical definition to avoid duplication of similar proofs, - for paths joined at the ends versus looping paths\ -definition linked_paths :: "bool \ (real \ 'a) \ (real \ 'a::topological_space) \ bool" - where "linked_paths atends g h == - (if atends then pathstart h = pathstart g \ pathfinish h = pathfinish g - else pathfinish g = pathstart g \ pathfinish h = pathstart h)" - -text\This formulation covers two cases: \<^term>\g\ and \<^term>\h\ share their - start and end points; \<^term>\g\ and \<^term>\h\ both loop upon themselves.\ -lemma contour_integral_nearby: - assumes os: "open S" and p: "path p" "path_image p \ S" - shows "\d. 0 < d \ - (\g h. valid_path g \ valid_path h \ - (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ - linked_paths atends g h - \ path_image g \ S \ path_image h \ S \ - (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" -proof - - have "\z. \e. z \ path_image p \ 0 < e \ ball z e \ S" - using open_contains_ball os p(2) by blast - then obtain ee where ee: "\z. z \ path_image p \ 0 < ee z \ ball z (ee z) \ S" - by metis - define cover where "cover = (\z. ball z (ee z/3)) ` (path_image p)" - have "compact (path_image p)" - by (metis p(1) compact_path_image) - moreover have "path_image p \ (\c\path_image p. ball c (ee c / 3))" - using ee by auto - ultimately have "\D \ cover. finite D \ path_image p \ \D" - by (simp add: compact_eq_Heine_Borel cover_def) - then obtain D where D: "D \ cover" "finite D" "path_image p \ \D" - by blast - then obtain k where k: "k \ {0..1}" "finite k" and D_eq: "D = ((\z. ball z (ee z / 3)) \ p) ` k" - apply (simp add: cover_def path_image_def image_comp) - apply (blast dest!: finite_subset_image [OF \finite D\]) - done - then have kne: "k \ {}" - using D by auto - have pi: "\i. i \ k \ p i \ path_image p" - using k by (auto simp: path_image_def) - then have eepi: "\i. i \ k \ 0 < ee((p i))" - by (metis ee) - define e where "e = Min((ee \ p) ` k)" - have fin_eep: "finite ((ee \ p) ` k)" - using k by blast - have "0 < e" - using ee k by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi) - have "uniformly_continuous_on {0..1} p" - using p by (simp add: path_def compact_uniformly_continuous) - then obtain d::real where d: "d>0" - and de: "\x x'. \x' - x\ < d \ x\{0..1} \ x'\{0..1} \ cmod (p x' - p x) < e/3" - unfolding uniformly_continuous_on_def dist_norm real_norm_def - by (metis divide_pos_pos \0 < e\ zero_less_numeral) - then obtain N::nat where N: "N>0" "inverse N < d" - using real_arch_inverse [of d] by auto - show ?thesis - proof (intro exI conjI allI; clarify?) - show "e/3 > 0" - using \0 < e\ by simp - fix g h - assume g: "valid_path g" and ghp: "\t\{0..1}. cmod (g t - p t) < e / 3 \ cmod (h t - p t) < e / 3" - and h: "valid_path h" - and joins: "linked_paths atends g h" - { fix t::real - assume t: "0 \ t" "t \ 1" - then obtain u where u: "u \ k" and ptu: "p t \ ball(p u) (ee(p u) / 3)" - using \path_image p \ \D\ D_eq by (force simp: path_image_def) - then have ele: "e \ ee (p u)" using fin_eep - by (simp add: e_def) - have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3" - using ghp t by auto - with ele have "cmod (g t - p t) < ee (p u) / 3" - "cmod (h t - p t) < ee (p u) / 3" - by linarith+ - then have "g t \ ball(p u) (ee(p u))" "h t \ ball(p u) (ee(p u))" - using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"] - norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u - by (force simp: dist_norm ball_def norm_minus_commute)+ - then have "g t \ S" "h t \ S" using ee u k - by (auto simp: path_image_def ball_def) - } - then have ghs: "path_image g \ S" "path_image h \ S" - by (auto simp: path_image_def) - moreover - { fix f - assume fhols: "f holomorphic_on S" - then have fpa: "f contour_integrable_on g" "f contour_integrable_on h" - using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple - by blast+ - have contf: "continuous_on S f" - by (simp add: fhols holomorphic_on_imp_continuous_on) - { fix z - assume z: "z \ path_image p" - have "f holomorphic_on ball z (ee z)" - using fhols ee z holomorphic_on_subset by blast - then have "\ff. (\w \ ball z (ee z). (ff has_field_derivative f w) (at w))" - using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified] - by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball) - } - then obtain ff where ff: - "\z w. \z \ path_image p; w \ ball z (ee z)\ \ (ff z has_field_derivative f w) (at w)" - by metis - { fix n - assume n: "n \ N" - then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f = - contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f" - proof (induct n) - case 0 show ?case by simp - next - case (Suc n) - obtain t where t: "t \ k" and "p (n/N) \ ball(p t) (ee(p t) / 3)" - using \path_image p \ \D\ [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems - by (force simp: path_image_def) - then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3" - by (simp add: dist_norm) - have e3le: "e/3 \ ee (p t) / 3" using fin_eep t - by (simp add: e_def) - { fix x - assume x: "n/N \ x" "x \ (1 + n)/N" - then have nN01: "0 \ n/N" "(1 + n)/N \ 1" - using Suc.prems by auto - then have x01: "0 \ x" "x \ 1" - using x by linarith+ - have "cmod (p t - p x) < ee (p t) / 3 + e/3" - proof (rule norm_diff_triangle_less [OF ptu de]) - show "\real n / real N - x\ < d" - using x N by (auto simp: field_simps) - qed (use x01 Suc.prems in auto) - then have ptx: "cmod (p t - p x) < 2*ee (p t)/3" - using e3le eepi [OF t] by simp - have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 " - apply (rule norm_diff_triangle_less [OF ptx]) - using ghp x01 by (simp add: norm_minus_commute) - also have "\ \ ee (p t)" - using e3le eepi [OF t] by simp - finally have gg: "cmod (p t - g x) < ee (p t)" . - have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 " - apply (rule norm_diff_triangle_less [OF ptx]) - using ghp x01 by (simp add: norm_minus_commute) - also have "\ \ ee (p t)" - using e3le eepi [OF t] by simp - finally have "cmod (p t - g x) < ee (p t)" - "cmod (p t - h x) < ee (p t)" - using gg by auto - } note ptgh_ee = this - have "closed_segment (g (real n / real N)) (h (real n / real N)) = path_image (linepath (h (n/N)) (g (n/N)))" - by (simp add: closed_segment_commute) - also have pi_hgn: "\ \ ball (p t) (ee (p t))" - using ptgh_ee [of "n/N"] Suc.prems - by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) - finally have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \ S" - using ee pi t by blast - have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) \ ball (p t) (ee (p t))" - using ptgh_ee [of "(1+n)/N"] Suc.prems - by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) - then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \ S" - using \N>0\ Suc.prems ee pi t - by (auto simp: Path_Connected.path_image_join field_simps) - have pi_subset_ball: - "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++ - subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N))) - \ ball (p t) (ee (p t))" - apply (intro subset_path_image_join pi_hgn pi_ghn') - using \N>0\ Suc.prems - apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) - done - have pi0: "(f has_contour_integral 0) - (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ - subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))" - apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"]) - apply (metis ff open_ball at_within_open pi t) - using Suc.prems pi_subset_ball apply (simp_all add: valid_path_join valid_path_subpath g h) - done - have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g" - using Suc.prems by (simp add: contour_integrable_subpath g fpa) - have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))" - using gh_n's - by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) - have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))" - using gh_ns - by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) - have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f + - contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f + - contour_integral (subpath ((Suc n) / N) (n/N) h) f + - contour_integral (linepath (h (n/N)) (g (n/N))) f = 0" - using contour_integral_unique [OF pi0] Suc.prems - by (simp add: g h fpa valid_path_subpath contour_integrable_subpath - fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) - have *: "\hn he hn' gn gd gn' hgn ghn gh0 ghn'. - \hn - gn = ghn - gh0; - gd + ghn' + he + hgn = (0::complex); - hn - he = hn'; gn + gd = gn'; hgn = -ghn\ \ hn' - gn' = ghn' - gh0" - by (auto simp: algebra_simps) - have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = - contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f" - unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"] - using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath) - also have "\ = contour_integral (subpath 0 ((Suc n) / N) h) f" - using Suc.prems by (simp add: contour_integral_subpath_combine h fpa) - finally have pi0_eq: - "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = - contour_integral (subpath 0 ((Suc n) / N) h) f" . - show ?case - apply (rule * [OF Suc.hyps eq0 pi0_eq]) - using Suc.prems - apply (simp_all add: g h fpa contour_integral_subpath_combine - contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath - continuous_on_subset [OF contf gh_ns]) - done - qed - } note ind = this - have "contour_integral h f = contour_integral g f" - using ind [OF order_refl] N joins - by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm) - } - ultimately - show "path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f)" - by metis - qed -qed - - -lemma - assumes "open S" "path p" "path_image p \ S" - shows contour_integral_nearby_ends: - "\d. 0 < d \ - (\g h. valid_path g \ valid_path h \ - (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ - pathstart h = pathstart g \ pathfinish h = pathfinish g - \ path_image g \ S \ - path_image h \ S \ - (\f. f holomorphic_on S - \ contour_integral h f = contour_integral g f))" - and contour_integral_nearby_loops: - "\d. 0 < d \ - (\g h. valid_path g \ valid_path h \ - (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ - pathfinish g = pathstart g \ pathfinish h = pathstart h - \ path_image g \ S \ - path_image h \ S \ - (\f. f holomorphic_on S - \ contour_integral h f = contour_integral g f))" - using contour_integral_nearby [OF assms, where atends=True] - using contour_integral_nearby [OF assms, where atends=False] - unfolding linked_paths_def by simp_all - -lemma C1_differentiable_polynomial_function: - fixes p :: "real \ 'a::euclidean_space" - shows "polynomial_function p \ p C1_differentiable_on S" - by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) - -lemma valid_path_polynomial_function: - fixes p :: "real \ 'a::euclidean_space" - shows "polynomial_function p \ valid_path p" -by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) - -lemma valid_path_subpath_trivial [simp]: - fixes g :: "real \ 'a::euclidean_space" - shows "z \ g x \ valid_path (subpath x x g)" - by (simp add: subpath_def valid_path_polynomial_function) - -lemma contour_integral_bound_exists: -assumes S: "open S" - and g: "valid_path g" - and pag: "path_image g \ S" - shows "\L. 0 < L \ - (\f B. f holomorphic_on S \ (\z \ S. norm(f z) \ B) - \ norm(contour_integral g f) \ L*B)" -proof - - have "path g" using g - by (simp add: valid_path_imp_path) - then obtain d::real and p - where d: "0 < d" - and p: "polynomial_function p" "path_image p \ S" - and pi: "\f. f holomorphic_on S \ contour_integral g f = contour_integral p f" - using contour_integral_nearby_ends [OF S \path g\ pag] - apply clarify - apply (drule_tac x=g in spec) - apply (simp only: assms) - apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function) - done - then obtain p' where p': "polynomial_function p'" - "\x. (p has_vector_derivative (p' x)) (at x)" - by (blast intro: has_vector_derivative_polynomial_function that) - then have "bounded(p' ` {0..1})" - using continuous_on_polymonial_function - by (force simp: intro!: compact_imp_bounded compact_continuous_image) - then obtain L where L: "L>0" and nop': "\x. \0 \ x; x \ 1\ \ norm (p' x) \ L" - by (force simp: bounded_pos) - { fix f B - assume f: "f holomorphic_on S" and B: "\z. z\S \ cmod (f z) \ B" - then have "f contour_integrable_on p \ valid_path p" - using p S - by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on) - moreover have "cmod (vector_derivative p (at x)) * cmod (f (p x)) \ L * B" if "0 \ x" "x \ 1" for x - proof (rule mult_mono) - show "cmod (vector_derivative p (at x)) \ L" - by (metis nop' p'(2) that vector_derivative_at) - show "cmod (f (p x)) \ B" - by (metis B atLeastAtMost_iff imageI p(2) path_defs(4) subset_eq that) - qed (use \L>0\ in auto) - ultimately have "cmod (contour_integral g f) \ L * B" - apply (simp only: pi [OF f]) - apply (simp only: contour_integral_integral) - apply (rule order_trans [OF integral_norm_bound_integral]) - apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult) - done - } then - show ?thesis - by (force simp: L contour_integral_integral) -qed - -text\We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\ - -subsection \Winding Numbers\ - -definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where - "winding_number_prop \ z e p n \ - valid_path p \ z \ path_image p \ - pathstart p = pathstart \ \ - pathfinish p = pathfinish \ \ - (\t \ {0..1}. norm(\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - -definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where - "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" - - -lemma winding_number: - assumes "path \" "z \ path_image \" "0 < e" - shows "\p. winding_number_prop \ z e p (winding_number \ z)" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain d - where d: "d>0" - and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ - (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ - pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ - path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ - (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" - using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) - then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ - (\t \ {0..1}. norm(h t - \ t) < d/2)" - using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto - define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" - have "\n. \e > 0. \p. winding_number_prop \ z e p n" - proof (rule_tac x=nn in exI, clarify) - fix e::real - assume e: "e>0" - obtain p where p: "polynomial_function p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d/2))" - using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by auto - have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" - by (auto simp: intro!: holomorphic_intros) - then show "\p. winding_number_prop \ z e p nn" - apply (rule_tac x=p in exI) - using pi_eq [of h p] h p d - apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def winding_number_prop_def) - done - qed - then show ?thesis - unfolding winding_number_def by (rule someI2_ex) (blast intro: \0) -qed - -lemma winding_number_unique: - assumes \: "path \" "z \ path_image \" - and pi: "\e. e>0 \ \p. winding_number_prop \ z e p n" - shows "winding_number \ z = n" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain e - where e: "e>0" - and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); - pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ - contour_integral h2 f = contour_integral h1 f" - using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) - obtain p where p: "winding_number_prop \ z e p n" - using pi [OF e] by blast - obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" - using winding_number [OF \ e] by blast - have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" - using p by (auto simp: winding_number_prop_def) - also have "\ = contour_integral q (\w. 1 / (w - z))" - proof (rule pi_eq) - show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" - by (auto intro!: holomorphic_intros) - qed (use p q in \auto simp: winding_number_prop_def norm_minus_commute\) - also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" - using q by (auto simp: winding_number_prop_def) - finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . - then show ?thesis - by simp -qed - -(*NB not winding_number_prop here due to the loop in p*) -lemma winding_number_unique_loop: - assumes \: "path \" "z \ path_image \" - and loop: "pathfinish \ = pathstart \" - and pi: - "\e. e>0 \ \p. valid_path p \ z \ path_image p \ - pathfinish p = pathstart p \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - shows "winding_number \ z = n" -proof - - have "path_image \ \ UNIV - {z}" - using assms by blast - then obtain e - where e: "e>0" - and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); - pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ - contour_integral h2 f = contour_integral h1 f" - using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) - obtain p where p: - "valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ - (\t \ {0..1}. norm (\ t - p t) < e) \ - contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" - using pi [OF e] by blast - obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" - using winding_number [OF \ e] by blast - have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" - using p by auto - also have "\ = contour_integral q (\w. 1 / (w - z))" - proof (rule pi_eq) - show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" - by (auto intro!: holomorphic_intros) - qed (use p q loop in \auto simp: winding_number_prop_def norm_minus_commute\) - also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" - using q by (auto simp: winding_number_prop_def) - finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . - then show ?thesis - by simp -qed - -proposition winding_number_valid_path: - assumes "valid_path \" "z \ path_image \" - shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" - by (rule winding_number_unique) - (use assms in \auto simp: valid_path_imp_path winding_number_prop_def\) - -proposition has_contour_integral_winding_number: - assumes \: "valid_path \" "z \ path_image \" - shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" -by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) - -lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" - by (simp add: winding_number_valid_path) - -lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" - by (simp add: path_image_subpath winding_number_valid_path) - -lemma winding_number_join: - assumes \1: "path \1" "z \ path_image \1" - and \2: "path \2" "z \ path_image \2" - and "pathfinish \1 = pathstart \2" - shows "winding_number(\1 +++ \2) z = winding_number \1 z + winding_number \2 z" -proof (rule winding_number_unique) - show "\p. winding_number_prop (\1 +++ \2) z e p - (winding_number \1 z + winding_number \2 z)" if "e > 0" for e - proof - - obtain p1 where "winding_number_prop \1 z e p1 (winding_number \1 z)" - using \0 < e\ \1 winding_number by blast - moreover - obtain p2 where "winding_number_prop \2 z e p2 (winding_number \2 z)" - using \0 < e\ \2 winding_number by blast - ultimately - have "winding_number_prop (\1+++\2) z e (p1+++p2) (winding_number \1 z + winding_number \2 z)" - using assms - apply (simp add: winding_number_prop_def not_in_path_image_join contour_integrable_inversediff algebra_simps) - apply (auto simp: joinpaths_def) - done - then show ?thesis - by blast - qed -qed (use assms in \auto simp: not_in_path_image_join\) - -lemma winding_number_reversepath: - assumes "path \" "z \ path_image \" - shows "winding_number(reversepath \) z = - (winding_number \ z)" -proof (rule winding_number_unique) - show "\p. winding_number_prop (reversepath \) z e p (- winding_number \ z)" if "e > 0" for e - proof - - obtain p where "winding_number_prop \ z e p (winding_number \ z)" - using \0 < e\ assms winding_number by blast - then have "winding_number_prop (reversepath \) z e (reversepath p) (- winding_number \ z)" - using assms - apply (simp add: winding_number_prop_def contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) - apply (auto simp: reversepath_def) - done - then show ?thesis - by blast - qed -qed (use assms in auto) - -lemma winding_number_shiftpath: - assumes \: "path \" "z \ path_image \" - and "pathfinish \ = pathstart \" "a \ {0..1}" - shows "winding_number(shiftpath a \) z = winding_number \ z" -proof (rule winding_number_unique_loop) - show "\p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ - (\t\{0..1}. cmod (shiftpath a \ t - p t) < e) \ - contour_integral p (\w. 1 / (w - z)) = - complex_of_real (2 * pi) * \ * winding_number \ z" - if "e > 0" for e - proof - - obtain p where "winding_number_prop \ z e p (winding_number \ z)" - using \0 < e\ assms winding_number by blast - then show ?thesis - apply (rule_tac x="shiftpath a p" in exI) - using assms that - apply (auto simp: winding_number_prop_def path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath contour_integral_shiftpath) - apply (simp add: shiftpath_def) - done - qed -qed (use assms in \auto simp: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath\) - -lemma winding_number_split_linepath: - assumes "c \ closed_segment a b" "z \ closed_segment a b" - shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" -proof - - have "z \ closed_segment a c" "z \ closed_segment c b" - using assms by (meson convex_contains_segment convex_segment ends_in_segment subsetCE)+ - then show ?thesis - using assms - by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) -qed - -lemma winding_number_cong: - "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" - by (simp add: winding_number_def winding_number_prop_def pathstart_def pathfinish_def) - -lemma winding_number_constI: - assumes "c\z" "\t. \0\t; t\1\ \ g t = c" - shows "winding_number g z = 0" -proof - - have "winding_number g z = winding_number (linepath c c) z" - apply (rule winding_number_cong) - using assms unfolding linepath_def by auto - moreover have "winding_number (linepath c c) z =0" - apply (rule winding_number_trivial) - using assms by auto - ultimately show ?thesis by auto -qed - -lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" - unfolding winding_number_def -proof (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) - fix n e g - assume "0 < e" and g: "winding_number_prop p z e g n" - then show "\r. winding_number_prop (\w. p w - z) 0 e r n" - by (rule_tac x="\t. g t - z" in exI) - (force simp: winding_number_prop_def contour_integral_integral valid_path_def path_defs - vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) -next - fix n e g - assume "0 < e" and g: "winding_number_prop (\w. p w - z) 0 e g n" - then show "\r. winding_number_prop p z e r n" - apply (rule_tac x="\t. g t + z" in exI) - apply (simp add: winding_number_prop_def contour_integral_integral valid_path_def path_defs - piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) - apply (force simp: algebra_simps) - done -qed - -subsubsection\<^marker>\tag unimportant\ \Some lemmas about negating a path\ - -lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" - unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast - -lemma has_contour_integral_negatepath: - assumes \: "valid_path \" and cint: "((\z. f (- z)) has_contour_integral - i) \" - shows "(f has_contour_integral i) (uminus \ \)" -proof - - obtain S where cont: "continuous_on {0..1} \" and "finite S" and diff: "\ C1_differentiable_on {0..1} - S" - using \ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) - have "((\x. - (f (- \ x) * vector_derivative \ (at x within {0..1}))) has_integral i) {0..1}" - using cint by (auto simp: has_contour_integral_def dest: has_integral_neg) - then - have "((\x. f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1})) has_integral i) {0..1}" - proof (rule rev_iffD1 [OF _ has_integral_spike_eq]) - show "negligible S" - by (simp add: \finite S\ negligible_finite) - show "f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1}) = - - (f (- \ x) * vector_derivative \ (at x within {0..1}))" - if "x \ {0..1} - S" for x - proof - - have "vector_derivative (uminus \ \) (at x within cbox 0 1) = - vector_derivative \ (at x within cbox 0 1)" - proof (rule vector_derivative_within_cbox) - show "(uminus \ \ has_vector_derivative - vector_derivative \ (at x within cbox 0 1)) (at x within cbox 0 1)" - using that unfolding o_def - by (metis C1_differentiable_on_eq UNIV_I diff differentiable_subset has_vector_derivative_minus subsetI that vector_derivative_works) - qed (use that in auto) - then show ?thesis - by simp - qed - qed - then show ?thesis by (simp add: has_contour_integral_def) -qed - -lemma winding_number_negatepath: - assumes \: "valid_path \" and 0: "0 \ path_image \" - shows "winding_number(uminus \ \) 0 = winding_number \ 0" -proof - - have "(/) 1 contour_integrable_on \" - using "0" \ contour_integrable_inversediff by fastforce - then have "((\z. 1/z) has_contour_integral contour_integral \ ((/) 1)) \" - by (rule has_contour_integral_integral) - then have "((\z. 1 / - z) has_contour_integral - contour_integral \ ((/) 1)) \" - using has_contour_integral_neg by auto - then show ?thesis - using assms - apply (simp add: winding_number_valid_path valid_path_negatepath image_def path_defs) - apply (simp add: contour_integral_unique has_contour_integral_negatepath) - done -qed - -lemma contour_integrable_negatepath: - assumes \: "valid_path \" and pi: "(\z. f (- z)) contour_integrable_on \" - shows "f contour_integrable_on (uminus \ \)" - by (metis \ add.inverse_inverse contour_integrable_on_def has_contour_integral_negatepath pi) - -(* A combined theorem deducing several things piecewise.*) -lemma winding_number_join_pos_combined: - "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); - valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ - \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" - by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) - - -subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ - -lemma Re_winding_number: - "\valid_path \; z \ path_image \\ - \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" -by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) - -lemma winding_number_pos_le: - assumes \: "valid_path \" "z \ path_image \" - and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" - shows "0 \ Re(winding_number \ z)" -proof - - have ge0: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x - using ge by (simp add: Complex.Im_divide algebra_simps x) - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" - let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" - have hi: "(?vd has_integral ?int z) (cbox 0 1)" - unfolding box_real - apply (subst has_contour_integral [symmetric]) - using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) - have "0 \ Im (?int z)" - proof (rule has_integral_component_nonneg [of \, simplified]) - show "\x. x \ cbox 0 1 \ 0 \ Im (if 0 < x \ x < 1 then ?vd x else 0)" - by (force simp: ge0) - show "((\x. if 0 < x \ x < 1 then ?vd x else 0) has_integral ?int z) (cbox 0 1)" - by (rule has_integral_spike_interior [OF hi]) simp - qed - then show ?thesis - by (simp add: Re_winding_number [OF \] field_simps) -qed - -lemma winding_number_pos_lt_lemma: - assumes \: "valid_path \" "z \ path_image \" - and e: "0 < e" - and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" - shows "0 < Re(winding_number \ z)" -proof - - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" - let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" - have hi: "(?vd has_integral ?int z) (cbox 0 1)" - unfolding box_real - apply (subst has_contour_integral [symmetric]) - using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) - have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" - proof (rule has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified]) - show "((\x. if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e) has_integral ?int z) {0..1}" - by (rule has_integral_spike_interior [OF hi, simplified box_real]) (use e in simp) - show "\x. 0 \ x \ x \ 1 \ - e \ Im (if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e)" - by (simp add: ge) - qed (use has_integral_const_real [of _ 0 1] in auto) - with e show ?thesis - by (simp add: Re_winding_number [OF \] field_simps) -qed - -lemma winding_number_pos_lt: - assumes \: "valid_path \" "z \ path_image \" - and e: "0 < e" - and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" - shows "0 < Re (winding_number \ z)" -proof - - have bm: "bounded ((\w. w - z) ` (path_image \))" - using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) - then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" - using bounded_pos [THEN iffD1, OF bm] by blast - { fix x::real assume x: "0 < x" "x < 1" - then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] - by (simp add: path_image_def power2_eq_square mult_mono') - with x have "\ x \ z" using \ - using path_image_def by fastforce - then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" - using B ge [OF x] B2 e - apply (rule_tac y="e / (cmod (\ x - z))\<^sup>2" in order_trans) - apply (auto simp: divide_left_mono divide_right_mono) - done - then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" - by (simp add: complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) - } note * = this - show ?thesis - using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) -qed - -subsection\The winding number is an integer\ - -text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, - Also on page 134 of Serge Lang's book with the name title, etc.\ - -lemma exp_fg: - fixes z::complex - assumes g: "(g has_vector_derivative g') (at x within s)" - and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" - and z: "g x \ z" - shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" -proof - - have *: "(exp \ (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" - using assms unfolding has_vector_derivative_def scaleR_conv_of_real - by (auto intro!: derivative_eq_intros) - show ?thesis - apply (rule has_vector_derivative_eq_rhs) - using z - apply (auto intro!: derivative_eq_intros * [unfolded o_def] g) - done -qed - -lemma winding_number_exp_integral: - fixes z::complex - assumes \: "\ piecewise_C1_differentiable_on {a..b}" - and ab: "a \ b" - and z: "z \ \ ` {a..b}" - shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" - (is "?thesis1") - "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" - (is "?thesis2") -proof - - let ?D\ = "\x. vector_derivative \ (at x)" - have [simp]: "\x. a \ x \ x \ b \ \ x \ z" - using z by force - have cong: "continuous_on {a..b} \" - using \ by (simp add: piecewise_C1_differentiable_on_def) - obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" - using \ by (force simp: piecewise_C1_differentiable_on_def) - have \: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) - moreover have "{a<.. {a..b} - k" - by force - ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" - by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open) - { fix w - assume "w \ z" - have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" - by (auto simp: dist_norm intro!: continuous_intros) - moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" - by (auto simp: intro!: derivative_eq_intros) - ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" - using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] - by (force simp: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) - } - then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" - by meson - have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" - unfolding integrable_on_def [symmetric] - proof (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \]]) - show "\d h. 0 < d \ - (\y. cmod (y - w) < d \ (h has_field_derivative inverse (y - z))(at y within - {z}))" - if "w \ - {z}" for w - apply (rule_tac x="norm(w - z)" in exI) - using that inverse_eq_divide has_field_derivative_at_within h - by (metis Compl_insert DiffD2 insertCI right_minus_eq zero_less_norm_iff) - qed simp - have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" - unfolding box_real [symmetric] divide_inverse_commute - by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) - with ab show ?thesis1 - by (simp add: divide_inverse_commute integral_def integrable_on_def) - { fix t - assume t: "t \ {a..b}" - have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" - using z by (auto intro!: continuous_intros simp: dist_norm) - have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" - unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) - obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ - (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" - using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] - by simp (auto simp: ball_def dist_norm that) - { fix x D - assume x: "x \ k" "a < x" "x < b" - then have "x \ interior ({a..b} - k)" - using open_subset_interior [OF \] by fastforce - then have con: "isCont ?D\ x" - using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) - then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" - by (rule continuous_at_imp_continuous_within) - have gdx: "\ differentiable at x" - using x by (simp add: g_diff_at) - have "\d. \x \ k; a < x; x < b; - (\ has_vector_derivative d) (at x); a \ t; t \ b\ - \ ((\x. integral {a..x} - (\x. ?D\ x / - (\ x - z))) has_vector_derivative - d / (\ x - z)) - (at x within {a..b})" - apply (rule has_vector_derivative_eq_rhs) - apply (rule integral_has_vector_derivative_continuous_at [where S = "{}", simplified]) - apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ - done - then have "((\c. exp (- integral {a..c} (\x. ?D\ x / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) - (at x within {a..b})" - using x gdx t - apply (clarsimp simp add: differentiable_iff_scaleR) - apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_withinI) - apply (simp_all add: has_vector_derivative_def [symmetric]) - done - } note * = this - have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" - apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) - using t - apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous_1 [OF vg_int] simp add: ab)+ - done - } - with ab show ?thesis2 - by (simp add: divide_inverse_commute integral_def) -qed - -lemma winding_number_exp_2pi: - "\path p; z \ path_image p\ - \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" -using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def winding_number_prop_def - by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) - -lemma integer_winding_number_eq: - assumes \: "path \" and z: "z \ path_image \" - shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" -proof - - obtain p where p: "valid_path p" "z \ path_image p" - "pathstart p = pathstart \" "pathfinish p = pathfinish \" - and eq: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF assms, of 1] unfolding winding_number_prop_def by auto - then have wneq: "winding_number \ z = winding_number p z" - using eq winding_number_valid_path by force - have iff: "(winding_number \ z \ \) \ (exp (contour_integral p (\w. 1 / (w - z))) = 1)" - using eq by (simp add: exp_eq_1 complex_is_Int_iff) - have "exp (contour_integral p (\w. 1 / (w - z))) = (\ 1 - z) / (\ 0 - z)" - using p winding_number_exp_integral(2) [of p 0 1 z] - apply (simp add: valid_path_def path_defs contour_integral_integral exp_minus field_split_simps) - by (metis path_image_def pathstart_def pathstart_in_path_image) - then have "winding_number p z \ \ \ pathfinish p = pathstart p" - using p wneq iff by (auto simp: path_defs) - then show ?thesis using p eq - by (auto simp: winding_number_valid_path) -qed - -theorem integer_winding_number: - "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" -by (metis integer_winding_number_eq) - - -text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) - We can thus bound the winding number of a path that doesn't intersect a given ray. \ - -lemma winding_number_pos_meets: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" - and w: "w \ z" - shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" -proof - - have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" - using z by (auto simp: path_image_def) - have [simp]: "z \ \ ` {0..1}" - using path_image_def z by auto - have gpd: "\ piecewise_C1_differentiable_on {0..1}" - using \ valid_path_def by blast - define r where "r = (w - z) / (\ 0 - z)" - have [simp]: "r \ 0" - using w z by (auto simp: r_def) - have cont: "continuous_on {0..1} - (\x. Im (integral {0..x} (\x. vector_derivative \ (at x) / (\ x - z))))" - by (intro continuous_intros indefinite_integral_continuous_1 winding_number_exp_integral [OF gpd]; simp) - have "Arg2pi r \ 2*pi" - by (simp add: Arg2pi less_eq_real_def) - also have "\ \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" - using 1 - apply (simp add: winding_number_valid_path [OF \ z] contour_integral_integral) - apply (simp add: Complex.Re_divide field_simps power2_eq_square) - done - finally have "Arg2pi r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . - then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" - by (simp add: Arg2pi_ge_0 cont IVT') - then obtain t where t: "t \ {0..1}" - and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" - by blast - define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" - have iArg: "Arg2pi r = Im i" - using eqArg by (simp add: i_def) - have gpdt: "\ piecewise_C1_differentiable_on {0..t}" - by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) - have "exp (- i) * (\ t - z) = \ 0 - z" - unfolding i_def - apply (rule winding_number_exp_integral [OF gpdt]) - using t z unfolding path_image_def by force+ - then have *: "\ t - z = exp i * (\ 0 - z)" - by (simp add: exp_minus field_simps) - then have "(w - z) = r * (\ 0 - z)" - by (simp add: r_def) - then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" - apply simp - apply (subst Complex_Transcendental.Arg2pi_eq [of r]) - apply (simp add: iArg) - using * apply (simp add: exp_eq_polar field_simps) - done - with t show ?thesis - by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) -qed - -lemma winding_number_big_meets: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" - and w: "w \ z" - shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" -proof - - { assume "Re (winding_number \ z) \ - 1" - then have "Re (winding_number (reversepath \) z) \ 1" - by (simp add: \ valid_path_imp_path winding_number_reversepath z) - moreover have "valid_path (reversepath \)" - using \ valid_path_imp_reverse by auto - moreover have "z \ path_image (reversepath \)" - by (simp add: z) - ultimately have "\a::real. 0 < a \ z + a*(w - z) \ path_image (reversepath \)" - using winding_number_pos_meets w by blast - then have ?thesis - by simp - } - then show ?thesis - using assms - by (simp add: abs_if winding_number_pos_meets split: if_split_asm) -qed - -lemma winding_number_less_1: - fixes z::complex - shows - "\valid_path \; z \ path_image \; w \ z; - \a::real. 0 < a \ z + a*(w - z) \ path_image \\ - \ Re(winding_number \ z) < 1" - by (auto simp: not_less dest: winding_number_big_meets) - -text\One way of proving that WN=1 for a loop.\ -lemma winding_number_eq_1: - fixes z::complex - assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" - and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" - shows "winding_number \ z = 1" -proof - - have "winding_number \ z \ Ints" - by (simp add: \ integer_winding_number loop valid_path_imp_path z) - then show ?thesis - using 0 2 by (auto simp: Ints_def) -qed - -subsection\Continuity of winding number and invariance on connected sets\ - -lemma continuous_at_winding_number: - fixes z::complex - assumes \: "path \" and z: "z \ path_image \" - shows "continuous (at z) (winding_number \)" -proof - - obtain e where "e>0" and cbg: "cball z e \ - path_image \" - using open_contains_cball [of "- path_image \"] z - by (force simp: closed_def [symmetric] closed_path_image [OF \]) - then have ppag: "path_image \ \ - cball z (e/2)" - by (force simp: cball_def dist_norm) - have oc: "open (- cball z (e / 2))" - by (simp add: closed_def [symmetric]) - obtain d where "d>0" and pi_eq: - "\h1 h2. \valid_path h1; valid_path h2; - (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); - pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ - \ - path_image h1 \ - cball z (e / 2) \ - path_image h2 \ - cball z (e / 2) \ - (\f. f holomorphic_on - cball z (e / 2) \ contour_integral h2 f = contour_integral h1 f)" - using contour_integral_nearby_ends [OF oc \ ppag] by metis - obtain p where p: "valid_path p" "z \ path_image p" - "pathstart p = pathstart \ \ pathfinish p = pathfinish \" - and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" - and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" - using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by (auto simp: winding_number_prop_def) - { fix w - assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" - then have wnotp: "w \ path_image p" - using cbg \d>0\ \e>0\ - apply (simp add: path_image_def cball_def dist_norm, clarify) - apply (frule pg) - apply (drule_tac c="\ x" in subsetD) - apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) - done - have wnotg: "w \ path_image \" - using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) - { fix k::real - assume k: "k>0" - then obtain q where q: "valid_path q" "w \ path_image q" - "pathstart q = pathstart \ \ pathfinish q = pathfinish \" - and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" - and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" - using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k - by (force simp: min_divide_distrib_right winding_number_prop_def) - have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" - apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) - apply (frule pg) - apply (frule qg) - using p q \d>0\ e2 - apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) - done - then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" - by (simp add: pi qi) - } note pip = this - have "path p" - using p by (simp add: valid_path_imp_path) - then have "winding_number p w = winding_number \ w" - apply (rule winding_number_unique [OF _ wnotp]) - apply (rule_tac x=p in exI) - apply (simp add: p wnotp min_divide_distrib_right pip winding_number_prop_def) - done - } note wnwn = this - obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" - using p open_contains_cball [of "- path_image p"] - by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) - obtain L - where "L>0" - and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); - \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ - cmod (contour_integral p f) \ L * B" - using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force - { fix e::real and w::complex - assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" - then have [simp]: "w \ path_image p" - using cbp p(2) \0 < pe\ - by (force simp: dist_norm norm_minus_commute path_image_def cball_def) - have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = - contour_integral p (\x. 1/(x - w) - 1/(x - z))" - by (simp add: p contour_integrable_inversediff contour_integral_diff) - { fix x - assume pe: "3/4 * pe < cmod (z - x)" - have "cmod (w - x) < pe/4 + cmod (z - x)" - by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) - then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp - have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" - using norm_diff_triangle_le by blast - also have "\ < pe/4 + cmod (w - x)" - using w by (simp add: norm_minus_commute) - finally have "pe/2 < cmod (w - x)" - using pe by auto - then have "(pe/2)^2 < cmod (w - x) ^ 2" - apply (rule power_strict_mono) - using \pe>0\ by auto - then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" - by (simp add: power_divide) - have "8 * L * cmod (w - z) < e * pe\<^sup>2" - using w \L>0\ by (simp add: field_simps) - also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" - using pe2 \e>0\ by (simp add: power2_eq_square) - also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" - using wx - apply (rule mult_strict_left_mono) - using pe2 e not_less_iff_gr_or_eq by fastforce - finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" - by simp - also have "\ \ e * cmod (w - x) * cmod (z - x)" - using e by simp - finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . - have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" - apply (cases "x=z \ x=w") - using pe \pe>0\ w \L>0\ - apply (force simp: norm_minus_commute) - using wx w(2) \L>0\ pe pe2 Lwz - apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) - done - } note L_cmod_le = this - have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" - apply (rule L) - using \pe>0\ w - apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) - using \pe>0\ w \L>0\ - apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) - done - have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" - apply simp - apply (rule le_less_trans [OF *]) - using \L>0\ e - apply (force simp: field_simps) - done - then have "cmod (winding_number p w - winding_number p z) < e" - using pi_ge_two e - by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) - } note cmod_wn_diff = this - then have "isCont (winding_number p) z" - apply (simp add: continuous_at_eps_delta, clarify) - apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) - using \pe>0\ \L>0\ - apply (simp add: dist_norm cmod_wn_diff) - done - then show ?thesis - apply (rule continuous_transform_within [where d = "min d e / 2"]) - apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) - done -qed - -corollary continuous_on_winding_number: - "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" - by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) - -subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ - -lemma winding_number_constant: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" - shows "winding_number \ constant_on S" -proof - - have *: "1 \ cmod (winding_number \ y - winding_number \ z)" - if ne: "winding_number \ y \ winding_number \ z" and "y \ S" "z \ S" for y z - proof - - have "winding_number \ y \ \" "winding_number \ z \ \" - using that integer_winding_number [OF \ loop] sg \y \ S\ by auto - with ne show ?thesis - by (auto simp: Ints_def simp flip: of_int_diff) - qed - have cont: "continuous_on S (\w. winding_number \ w)" - using continuous_on_winding_number [OF \] sg - by (meson continuous_on_subset disjoint_eq_subset_Compl) - show ?thesis - using "*" zero_less_one - by (blast intro: continuous_discrete_range_constant [OF cs cont]) -qed - -lemma winding_number_eq: - "\path \; pathfinish \ = pathstart \; w \ S; z \ S; connected S; S \ path_image \ = {}\ - \ winding_number \ w = winding_number \ z" - using winding_number_constant by (metis constant_on_def) - -lemma open_winding_number_levelsets: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - shows "open {z. z \ path_image \ \ winding_number \ z = k}" -proof - - have opn: "open (- path_image \)" - by (simp add: closed_path_image \ open_Compl) - { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" - obtain e where e: "e>0" "ball z e \ - path_image \" - using open_contains_ball [of "- path_image \"] opn z - by blast - have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" - apply (rule_tac x=e in exI) - using e apply (simp add: dist_norm ball_def norm_minus_commute) - apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where S = "ball z e"]) - done - } then - show ?thesis - by (auto simp: open_dist) -qed - -subsection\Winding number is zero "outside" a curve\ - -proposition winding_number_zero_in_outside: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" - shows "winding_number \ z = 0" -proof - - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" - using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto - obtain w::complex where w: "w \ ball 0 (B + 1)" - by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) - have "- ball 0 (B + 1) \ outside (path_image \)" - apply (rule outside_subset_convex) - using B subset_ball by auto - then have wout: "w \ outside (path_image \)" - using w by blast - moreover have "winding_number \ constant_on outside (path_image \)" - using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside - by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) - ultimately have "winding_number \ z = winding_number \ w" - by (metis (no_types, hide_lams) constant_on_def z) - also have "\ = 0" - proof - - have wnot: "w \ path_image \" using wout by (simp add: outside_def) - { fix e::real assume "0" "pathfinish p = pathfinish \" - and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" - and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" - using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force - have pip: "path_image p \ ball 0 (B + 1)" - using B - apply (clarsimp simp add: path_image_def dist_norm ball_def) - apply (frule (1) pg1) - apply (fastforce dest: norm_add_less) - done - then have "w \ path_image p" using w by blast - then have "\p. valid_path p \ w \ path_image p \ - pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ - (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" - apply (rule_tac x=p in exI) - apply (simp add: p valid_path_polynomial_function) - apply (intro conjI) - using pge apply (simp add: norm_minus_commute) - apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) - apply (rule holomorphic_intros | simp add: dist_norm)+ - using mem_ball_0 w apply blast - using p apply (simp_all add: valid_path_polynomial_function loop pip) - done - } - then show ?thesis - by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) - qed - finally show ?thesis . -qed - -corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" - by (rule winding_number_zero_in_outside) - (auto simp: pathfinish_def pathstart_def path_polynomial_function) - -corollary\<^marker>\tag unimportant\ winding_number_zero_outside: - "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" - by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) - -lemma winding_number_zero_at_infinity: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - shows "\B. \z. B \ norm z \ winding_number \ z = 0" -proof - - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" - using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto - then show ?thesis - apply (rule_tac x="B+1" in exI, clarify) - apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) - apply (meson less_add_one mem_cball_0 not_le order_trans) - using ball_subset_cball by blast -qed - -lemma winding_number_zero_point: - "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ - \ \z. z \ s \ winding_number \ z = 0" - using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside - by (fastforce simp add: compact_path_image) - - -text\If a path winds round a set, it winds rounds its inside.\ -lemma winding_number_around_inside: - assumes \: "path \" and loop: "pathfinish \ = pathstart \" - and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" - and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" - shows "winding_number \ w = winding_number \ z" -proof - - have ssb: "s \ inside(path_image \)" - proof - fix x :: complex - assume "x \ s" - hence "x \ path_image \" - by (meson disjoint_iff_not_equal s_disj) - thus "x \ inside (path_image \)" - using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) -qed - show ?thesis - apply (rule winding_number_eq [OF \ loop w]) - using z apply blast - apply (simp add: cls connected_with_inside cos) - apply (simp add: Int_Un_distrib2 s_disj, safe) - by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) - qed - - -text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ -lemma winding_number_subpath_continuous: - assumes \: "valid_path \" and z: "z \ path_image \" - shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" -proof - - have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = - winding_number (subpath 0 x \) z" - if x: "0 \ x" "x \ 1" for x - proof - - have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = - 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" - using assms x - apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) - done - also have "\ = winding_number (subpath 0 x \) z" - apply (subst winding_number_valid_path) - using assms x - apply (simp_all add: path_image_subpath valid_path_subpath) - by (force simp: path_image_def) - finally show ?thesis . - qed - show ?thesis - apply (rule continuous_on_eq - [where f = "\x. 1 / (2*pi*\) * - integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) - apply (rule continuous_intros)+ - apply (rule indefinite_integral_continuous_1) - apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) - using assms - apply (simp add: *) - done -qed - -lemma winding_number_ivt_pos: - assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" - shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" - apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) - apply (rule winding_number_subpath_continuous [OF \ z]) - using assms - apply (auto simp: path_image_def image_def) - done - -lemma winding_number_ivt_neg: - assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" - shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" - apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) - apply (rule winding_number_subpath_continuous [OF \ z]) - using assms - apply (auto simp: path_image_def image_def) - done - -lemma winding_number_ivt_abs: - assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" - shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" - using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] - by force - -lemma winding_number_lt_half_lemma: - assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" - shows "Re(winding_number \ z) < 1/2" -proof - - { assume "Re(winding_number \ z) \ 1/2" - then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" - using winding_number_ivt_pos [OF \ z, of "1/2"] by auto - have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" - using winding_number_exp_2pi [of "subpath 0 t \" z] - apply (simp add: t \ valid_path_imp_path) - using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) - have "b < a \ \ 0" - proof - - have "\ 0 \ {c. b < a \ c}" - by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) - thus ?thesis - by blast - qed - moreover have "b < a \ \ t" - proof - - have "\ t \ {c. b < a \ c}" - by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) - thus ?thesis - by blast - qed - ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az - by (simp add: inner_diff_right)+ - then have False - by (simp add: gt inner_mult_right mult_less_0_iff) - } - then show ?thesis by force -qed - -lemma winding_number_lt_half: - assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" - shows "\Re (winding_number \ z)\ < 1/2" -proof - - have "z \ path_image \" using assms by auto - with assms show ?thesis - apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) - apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] - winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) - done -qed - -lemma winding_number_le_half: - assumes \: "valid_path \" and z: "z \ path_image \" - and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" - shows "\Re (winding_number \ z)\ \ 1/2" -proof - - { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" - have "isCont (winding_number \) z" - by (metis continuous_at_winding_number valid_path_imp_path \ z) - then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" - using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast - define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" - have *: "a \ z' \ b - d / 3 * cmod a" - unfolding z'_def inner_mult_right' divide_inverse - apply (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square anz) - apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) - done - have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" - using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) - then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" - by simp - then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" - using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp - then have wnz_12': "\Re (winding_number \ z')\ > 1/2" - by linarith - moreover have "\Re (winding_number \ z')\ < 1/2" - apply (rule winding_number_lt_half [OF \ *]) - using azb \d>0\ pag - apply (auto simp: add_strict_increasing anz field_split_simps dest!: subsetD) - done - ultimately have False - by simp - } - then show ?thesis by force -qed - -lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" - using separating_hyperplane_closed_point [of "closed_segment a b" z] - apply auto - apply (simp add: closed_segment_def) - apply (drule less_imp_le) - apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) - apply (auto simp: segment) - done - - -text\ Positivity of WN for a linepath.\ -lemma winding_number_linepath_pos_lt: - assumes "0 < Im ((b - a) * cnj (b - z))" - shows "0 < Re(winding_number(linepath a b) z)" -proof - - have z: "z \ path_image (linepath a b)" - using assms - by (simp add: closed_segment_def) (force simp: algebra_simps) - show ?thesis - apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) - apply (simp add: linepath_def algebra_simps) - done -qed - - -subsection\Cauchy's integral formula, again for a convex enclosing set\ - -lemma Cauchy_integral_formula_weak: - assumes s: "convex s" and "finite k" and conf: "continuous_on s f" - and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" - and z: "z \ interior s - k" and vpg: "valid_path \" - and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - obtain f' where f': "(f has_field_derivative f') (at z)" - using fcd [OF z] by (auto simp: field_differentiable_def) - have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ - have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x - proof (cases "x = z") - case True then show ?thesis - apply (simp add: continuous_within) - apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) - using has_field_derivative_at_within has_field_derivative_iff f' - apply (fastforce simp add:)+ - done - next - case False - then have dxz: "dist x z > 0" by auto - have cf: "continuous (at x within s) f" - using conf continuous_on_eq_continuous_within that by blast - have "continuous (at x within s) (\w. (f w - f z) / (w - z))" - by (rule cf continuous_intros | simp add: False)+ - then show ?thesis - apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) - apply (force simp: dist_commute) - done - qed - have fink': "finite (insert z k)" using \finite k\ by blast - have *: "((\w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \" - apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) - using c apply (force simp: continuous_on_eq_continuous_within) - apply (rename_tac w) - apply (rule_tac d="dist w z" and f = "\w. (f w - f z)/(w - z)" in field_differentiable_transform_within) - apply (simp_all add: dist_pos_lt dist_commute) - apply (metis less_irrefl) - apply (rule derivative_intros fcd | simp)+ - done - show ?thesis - apply (rule has_contour_integral_eq) - using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] - apply (auto simp: ac_simps divide_simps) - done -qed - -theorem Cauchy_integral_formula_convex_simple: - "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; - pathfinish \ = pathstart \\ - \ ((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" - apply (rule Cauchy_integral_formula_weak [where k = "{}"]) - using holomorphic_on_imp_continuous_on - by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) - -subsection\Homotopy forms of Cauchy's theorem\ - -lemma Cauchy_theorem_homotopic: - assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" - and "open s" and f: "f holomorphic_on s" - and vpg: "valid_path g" and vph: "valid_path h" - shows "contour_integral g f = contour_integral h f" -proof - - have pathsf: "linked_paths atends g h" - using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) - obtain k :: "real \ real \ complex" - where contk: "continuous_on ({0..1} \ {0..1}) k" - and ks: "k ` ({0..1} \ {0..1}) \ s" - and k [simp]: "\x. k (0, x) = g x" "\x. k (1, x) = h x" - and ksf: "\t\{0..1}. linked_paths atends g (\x. k (t, x))" - using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) - have ucontk: "uniformly_continuous_on ({0..1} \ {0..1}) k" - by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) - { fix t::real assume t: "t \ {0..1}" - have pak: "path (k \ (\u. (t, u)))" - unfolding path_def - apply (rule continuous_intros continuous_on_subset [OF contk])+ - using t by force - have pik: "path_image (k \ Pair t) \ s" - using ks t by (auto simp: path_image_def) - obtain e where "e>0" and e: - "\g h. \valid_path g; valid_path h; - \u\{0..1}. cmod (g u - (k \ Pair t) u) < e \ cmod (h u - (k \ Pair t) u) < e; - linked_paths atends g h\ - \ contour_integral h f = contour_integral g f" - using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis - obtain d where "d>0" and d: - "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" - by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) - { fix t1 t2 - assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" - have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" - using \e > 0\ - apply (rule_tac y = k1 in norm_triangle_half_l) - apply (auto simp: norm_minus_commute intro: order_less_trans) - done - have "\d>0. \g1 g2. valid_path g1 \ valid_path g2 \ - (\u\{0..1}. cmod (g1 u - k (t1, u)) < d \ cmod (g2 u - k (t2, u)) < d) \ - linked_paths atends g1 g2 \ - contour_integral g2 f = contour_integral g1 f" - apply (rule_tac x="e/4" in exI) - using t t1 t2 ltd \e > 0\ - apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) - done - } - then have "\e. 0 < e \ - (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < e \ \t2 - t\ < e - \ (\d. 0 < d \ - (\g1 g2. valid_path g1 \ valid_path g2 \ - (\u \ {0..1}. - norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ - linked_paths atends g1 g2 - \ contour_integral g2 f = contour_integral g1 f)))" - by (rule_tac x=d in exI) (simp add: \d > 0\) - } - then obtain ee where ee: - "\t. t \ {0..1} \ ee t > 0 \ - (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < ee t \ \t2 - t\ < ee t - \ (\d. 0 < d \ - (\g1 g2. valid_path g1 \ valid_path g2 \ - (\u \ {0..1}. - norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ - linked_paths atends g1 g2 - \ contour_integral g2 f = contour_integral g1 f)))" - by metis - note ee_rule = ee [THEN conjunct2, rule_format] - define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" - obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" - proof (rule compactE [OF compact_interval]) - show "{0..1} \ \C" - using ee [THEN conjunct1] by (auto simp: C_def dist_norm) - qed (use C_def in auto) - define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" - have kk01: "kk \ {0..1}" by (auto simp: kk_def) - define e where "e = Min (ee ` kk)" - have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" - using C' by (auto simp: kk_def C_def) - have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" - by (simp add: kk_def ee) - moreover have "finite kk" - using \finite C'\ kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) - moreover have "kk \ {}" using \{0..1} \ \C'\ C'_eq by force - ultimately have "e > 0" - using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) - then obtain N::nat where "N > 0" and N: "1/N < e/3" - by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) - have e_le_ee: "\i. i \ kk \ e \ ee i" - using \finite kk\ by (simp add: e_def Min_le_iff [of "ee ` kk"]) - have plus: "\t \ kk. x \ ball t (ee t / 3)" if "x \ {0..1}" for x - using C' subsetD [OF C'01 that] unfolding C'_eq by blast - have [OF order_refl]: - "\d. 0 < d \ (\j. valid_path j \ (\u \ {0..1}. norm(j u - k (n/N, u)) < d) \ linked_paths atends g j - \ contour_integral j f = contour_integral g f)" - if "n \ N" for n - using that - proof (induct n) - case 0 show ?case using ee_rule [of 0 0 0] - apply clarsimp - apply (rule_tac x=d in exI, safe) - by (metis diff_self vpg norm_zero) - next - case (Suc n) - then have N01: "n/N \ {0..1}" "(Suc n)/N \ {0..1}" by auto - then obtain t where t: "t \ kk" "n/N \ ball t (ee t / 3)" - using plus [of "n/N"] by blast - then have nN_less: "\n/N - t\ < ee t" - by (simp add: dist_norm del: less_divide_eq_numeral1) - have n'N_less: "\real (Suc n) / real N - t\ < ee t" - using t N \N > 0\ e_le_ee [of t] - by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) - have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast - obtain d1 where "d1 > 0" and d1: - "\g1 g2. \valid_path g1; valid_path g2; - \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; - linked_paths atends g1 g2\ - \ contour_integral g2 f = contour_integral g1 f" - using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce - have "n \ N" using Suc.prems by auto - with Suc.hyps - obtain d2 where "d2 > 0" - and d2: "\j. \valid_path j; \u\{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\ - \ contour_integral j f = contour_integral g f" - by auto - have "continuous_on {0..1} (k \ (\u. (n/N, u)))" - apply (rule continuous_intros continuous_on_subset [OF contk])+ - using N01 by auto - then have pkn: "path (\u. k (n/N, u))" - by (simp add: path_def) - have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) - obtain p where "polynomial_function p" - and psf: "pathstart p = pathstart (\u. k (n/N, u))" - "pathfinish p = pathfinish (\u. k (n/N, u))" - and pk_le: "\t. t\{0..1} \ cmod (p t - k (n/N, t)) < min d1 d2" - using path_approx_polynomial_function [OF pkn min12] by blast - then have vpp: "valid_path p" using valid_path_polynomial_function by blast - have lpa: "linked_paths atends g p" - by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) - show ?case - proof (intro exI; safe) - fix j - assume "valid_path j" "linked_paths atends g j" - and "\u\{0..1}. cmod (j u - k (real (Suc n) / real N, u)) < min d1 d2" - then have "contour_integral j f = contour_integral p f" - using pk_le N01(1) ksf by (force intro!: vpp d1 simp add: linked_paths_def psf) - also have "... = contour_integral g f" - using pk_le by (force intro!: vpp d2 lpa) - finally show "contour_integral j f = contour_integral g f" . - qed (simp add: \0 < d1\ \0 < d2\) - qed - then obtain d where "0 < d" - "\j. valid_path j \ (\u \ {0..1}. norm(j u - k (1,u)) < d) \ linked_paths atends g j - \ contour_integral j f = contour_integral g f" - using \N>0\ by auto - then have "linked_paths atends g h \ contour_integral h f = contour_integral g f" - using \N>0\ vph by fastforce - then show ?thesis - by (simp add: pathsf) -qed - -proposition Cauchy_theorem_homotopic_paths: - assumes hom: "homotopic_paths s g h" - and "open s" and f: "f holomorphic_on s" - and vpg: "valid_path g" and vph: "valid_path h" - shows "contour_integral g f = contour_integral h f" - using Cauchy_theorem_homotopic [of True s g h] assms by simp - -proposition Cauchy_theorem_homotopic_loops: - assumes hom: "homotopic_loops s g h" - and "open s" and f: "f holomorphic_on s" - and vpg: "valid_path g" and vph: "valid_path h" - shows "contour_integral g f = contour_integral h f" - using Cauchy_theorem_homotopic [of False s g h] assms by simp - -lemma has_contour_integral_newpath: - "\(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\ - \ (f has_contour_integral y) g" - using has_contour_integral_integral contour_integral_unique by auto - -lemma Cauchy_theorem_null_homotopic: - "\f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\ \ (f has_contour_integral 0) g" - apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) - using contour_integrable_holomorphic_simple - apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) - by (simp add: Cauchy_theorem_homotopic_loops) - -subsection\<^marker>\tag unimportant\ \More winding number properties\ - -text\including the fact that it's +-1 inside a simple closed curve.\ - -lemma winding_number_homotopic_paths: - assumes "homotopic_paths (-{z}) g h" - shows "winding_number g z = winding_number h z" -proof - - have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto - moreover have pag: "z \ path_image g" and pah: "z \ path_image h" - using homotopic_paths_imp_subset [OF assms] by auto - ultimately obtain d e where "d > 0" "e > 0" - and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ - \ homotopic_paths (-{z}) g p" - and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ - \ homotopic_paths (-{z}) h q" - using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force - obtain p where p: - "valid_path p" "z \ path_image p" - "pathstart p = pathstart g" "pathfinish p = pathfinish g" - and gp_less:"\t\{0..1}. cmod (g t - p t) < d" - and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" - using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast - obtain q where q: - "valid_path q" "z \ path_image q" - "pathstart q = pathstart h" "pathfinish q = pathfinish h" - and hq_less: "\t\{0..1}. cmod (h t - q t) < e" - and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" - using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast - have "homotopic_paths (- {z}) g p" - by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) - moreover have "homotopic_paths (- {z}) h q" - by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) - ultimately have "homotopic_paths (- {z}) p q" - by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) - then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" - by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) - then show ?thesis - by (simp add: pap paq) -qed - -lemma winding_number_homotopic_loops: - assumes "homotopic_loops (-{z}) g h" - shows "winding_number g z = winding_number h z" -proof - - have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto - moreover have pag: "z \ path_image g" and pah: "z \ path_image h" - using homotopic_loops_imp_subset [OF assms] by auto - moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" - using homotopic_loops_imp_loop [OF assms] by auto - ultimately obtain d e where "d > 0" "e > 0" - and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ - \ homotopic_loops (-{z}) g p" - and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ - \ homotopic_loops (-{z}) h q" - using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force - obtain p where p: - "valid_path p" "z \ path_image p" - "pathstart p = pathstart g" "pathfinish p = pathfinish g" - and gp_less:"\t\{0..1}. cmod (g t - p t) < d" - and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" - using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast - obtain q where q: - "valid_path q" "z \ path_image q" - "pathstart q = pathstart h" "pathfinish q = pathfinish h" - and hq_less: "\t\{0..1}. cmod (h t - q t) < e" - and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" - using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast - have gp: "homotopic_loops (- {z}) g p" - by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) - have hq: "homotopic_loops (- {z}) h q" - by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) - have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" - proof (rule Cauchy_theorem_homotopic_loops) - show "homotopic_loops (- {z}) p q" - by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) - qed (auto intro!: holomorphic_intros simp: p q) - then show ?thesis - by (simp add: pap paq) -qed - -lemma winding_number_paths_linear_eq: - "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; - \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ - \ winding_number h z = winding_number g z" - by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) - -lemma winding_number_loops_linear_eq: - "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; - \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ - \ winding_number h z = winding_number g z" - by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) - -lemma winding_number_nearby_paths_eq: - "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; - \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ - \ winding_number h z = winding_number g z" - by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) - -lemma winding_number_nearby_loops_eq: - "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; - \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ - \ winding_number h z = winding_number g z" - by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) - - -lemma winding_number_subpath_combine: - "\path g; z \ path_image g; - u \ {0..1}; v \ {0..1}; w \ {0..1}\ - \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = - winding_number (subpath u w g) z" -apply (rule trans [OF winding_number_join [THEN sym] - winding_number_homotopic_paths [OF homotopic_join_subpaths]]) - using path_image_subpath_subset by auto - -subsection\Partial circle path\ - -definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" - where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" - -lemma pathstart_part_circlepath [simp]: - "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" -by (metis part_circlepath_def pathstart_def pathstart_linepath) - -lemma pathfinish_part_circlepath [simp]: - "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" -by (metis part_circlepath_def pathfinish_def pathfinish_linepath) - -lemma reversepath_part_circlepath[simp]: - "reversepath (part_circlepath z r s t) = part_circlepath z r t s" - unfolding part_circlepath_def reversepath_def linepath_def - by (auto simp:algebra_simps) - -lemma has_vector_derivative_part_circlepath [derivative_intros]: - "((part_circlepath z r s t) has_vector_derivative - (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) - (at x within X)" - apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) - apply (rule has_vector_derivative_real_field) - apply (rule derivative_eq_intros | simp)+ - done - -lemma differentiable_part_circlepath: - "part_circlepath c r a b differentiable at x within A" - using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast - -lemma vector_derivative_part_circlepath: - "vector_derivative (part_circlepath z r s t) (at x) = - \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" - using has_vector_derivative_part_circlepath vector_derivative_at by blast - -lemma vector_derivative_part_circlepath01: - "\0 \ x; x \ 1\ - \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = - \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" - using has_vector_derivative_part_circlepath - by (auto simp: vector_derivative_at_within_ivl) - -lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" - apply (simp add: valid_path_def) - apply (rule C1_differentiable_imp_piecewise) - apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath - intro!: continuous_intros) - done - -lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" - by (simp add: valid_path_imp_path) - -proposition path_image_part_circlepath: - assumes "s \ t" - shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" -proof - - { fix z::real - assume "0 \ z" "z \ 1" - with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" - apply (rule_tac x="(1 - z) * s + z * t" in exI) - apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) - apply (rule conjI) - using mult_right_mono apply blast - using affine_ineq by (metis "mult.commute") - } - moreover - { fix z - assume "s \ z" "z \ t" - then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" - apply (rule_tac x="(z - s)/(t - s)" in image_eqI) - apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) - apply (auto simp: field_split_simps) - done - } - ultimately show ?thesis - by (fastforce simp add: path_image_def part_circlepath_def) -qed - -lemma path_image_part_circlepath': - "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" -proof - - have "path_image (part_circlepath z r s t) = - (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" - by (simp add: image_image path_image_def part_circlepath_def) - also have "linepath s t ` {0..1} = closed_segment s t" - by (rule linepath_image_01) - finally show ?thesis by (simp add: cis_conv_exp) -qed - -lemma path_image_part_circlepath_subset: - "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" -by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) - -lemma in_path_image_part_circlepath: - assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" - shows "norm(w - z) = r" -proof - - have "w \ {c. dist z c = r}" - by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) - thus ?thesis - by (simp add: dist_norm norm_minus_commute) -qed - -lemma path_image_part_circlepath_subset': - assumes "r \ 0" - shows "path_image (part_circlepath z r s t) \ sphere z r" -proof (cases "s \ t") - case True - thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp -next - case False - thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms - by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all -qed - -lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" - by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) - -lemma contour_integral_bound_part_circlepath: - assumes "f contour_integrable_on part_circlepath c r a b" - assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" - shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" -proof - - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * - exp (\ * linepath a b x))" - have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" - proof (rule integral_norm_bound_integral, goal_cases) - case 1 - with assms(1) show ?case - by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) - next - case (3 x) - with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult - by (intro mult_mono) (auto simp: path_image_def) - qed auto - also have "?I = contour_integral (part_circlepath c r a b) f" - by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) - finally show ?thesis by simp -qed - -lemma has_contour_integral_part_circlepath_iff: - assumes "a < b" - shows "(f has_contour_integral I) (part_circlepath c r a b) \ - ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" -proof - - have "(f has_contour_integral I) (part_circlepath c r a b) \ - ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) - (at x within {0..1})) has_integral I) {0..1}" - unfolding has_contour_integral_def .. - also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * - cis (linepath a b x)) has_integral I) {0..1}" - by (intro has_integral_cong, subst vector_derivative_part_circlepath01) - (simp_all add: cis_conv_exp) - also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * - r * \ * exp (\ * linepath (of_real a) (of_real b) x) * - vector_derivative (linepath (of_real a) (of_real b)) - (at x within {0..1})) has_integral I) {0..1}" - by (intro has_integral_cong, subst vector_derivative_linepath_within) - (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) - also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) - (linepath (of_real a) (of_real b))" - by (simp add: has_contour_integral_def) - also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms - by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) - finally show ?thesis . -qed - -lemma contour_integrable_part_circlepath_iff: - assumes "a < b" - shows "f contour_integrable_on (part_circlepath c r a b) \ - (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" - using assms by (auto simp: contour_integrable_on_def integrable_on_def - has_contour_integral_part_circlepath_iff) - -lemma contour_integral_part_circlepath_eq: - assumes "a < b" - shows "contour_integral (part_circlepath c r a b) f = - integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" -proof (cases "f contour_integrable_on part_circlepath c r a b") - case True - hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" - using assms by (simp add: contour_integrable_part_circlepath_iff) - with True show ?thesis - using has_contour_integral_part_circlepath_iff[OF assms] - contour_integral_unique has_integral_integrable_integral by blast -next - case False - hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" - using assms by (simp add: contour_integrable_part_circlepath_iff) - with False show ?thesis - by (simp add: not_integrable_contour_integral not_integrable_integral) -qed - -lemma contour_integral_part_circlepath_reverse: - "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" - by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all - -lemma contour_integral_part_circlepath_reverse': - "b < a \ contour_integral (part_circlepath c r a b) f = - -contour_integral (part_circlepath c r b a) f" - by (rule contour_integral_part_circlepath_reverse) - -lemma finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" -proof (cases "w = 0") - case True then show ?thesis by auto -next - case False - have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" - apply (simp add: norm_mult finite_int_iff_bounded_le) - apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) - apply (auto simp: field_split_simps le_floor_iff) - done - have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" - by blast - show ?thesis - apply (subst exp_Ln [OF False, symmetric]) - apply (simp add: exp_eq) - using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) - done -qed - -lemma finite_bounded_log2: - fixes a::complex - assumes "a \ 0" - shows "finite {z. norm z \ b \ exp(a*z) = w}" -proof - - have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" - by (rule finite_imageI [OF finite_bounded_log]) - show ?thesis - by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) -qed - -lemma has_contour_integral_bound_part_circlepath_strong: - assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" - and "finite k" and le: "0 \ B" "0 < r" "s \ t" - and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" - shows "cmod i \ B * r * (t - s)" -proof - - consider "s = t" | "s < t" using \s \ t\ by linarith - then show ?thesis - proof cases - case 1 with fi [unfolded has_contour_integral] - have "i = 0" by (simp add: vector_derivative_part_circlepath) - with assms show ?thesis by simp - next - case 2 - have [simp]: "\r\ = r" using \r > 0\ by linarith - have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" - by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) - have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y - proof - - define w where "w = (y - z)/of_real r / exp(\ * of_real s)" - have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" - apply (rule finite_vimageI [OF finite_bounded_log2]) - using \s < t\ apply (auto simp: inj_of_real) - done - show ?thesis - apply (simp add: part_circlepath_def linepath_def vimage_def) - apply (rule finite_subset [OF _ fin]) - using le - apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) - done - qed - then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" - by (rule finite_finite_vimage_IntI [OF \finite k\]) - have **: "((\x. if (part_circlepath z r s t x) \ k then 0 - else f(part_circlepath z r s t x) * - vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" - by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) - have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" - by (auto intro!: B [unfolded path_image_def image_def, simplified]) - show ?thesis - apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) - using assms apply force - apply (simp add: norm_mult vector_derivative_part_circlepath) - using le * "2" \r > 0\ by auto - qed -qed - -lemma has_contour_integral_bound_part_circlepath: - "\(f has_contour_integral i) (part_circlepath z r s t); - 0 \ B; 0 < r; s \ t; - \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ - \ norm i \ B*r*(t - s)" - by (auto intro: has_contour_integral_bound_part_circlepath_strong) - -lemma contour_integrable_continuous_part_circlepath: - "continuous_on (path_image (part_circlepath z r s t)) f - \ f contour_integrable_on (part_circlepath z r s t)" - apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) - apply (rule integrable_continuous_real) - apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) - done - -proposition winding_number_part_circlepath_pos_less: - assumes "s < t" and no: "norm(w - z) < r" - shows "0 < Re (winding_number(part_circlepath z r s t) w)" -proof - - have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) - note valid_path_part_circlepath - moreover have " w \ path_image (part_circlepath z r s t)" - using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) - moreover have "0 < r * (t - s) * (r - cmod (w - z))" - using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) - ultimately show ?thesis - apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) - apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) - apply (rule mult_left_mono)+ - using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] - apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) - using assms \0 < r\ by auto -qed - -lemma simple_path_part_circlepath: - "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" -proof (cases "r = 0 \ s = t") - case True - then show ?thesis - unfolding part_circlepath_def simple_path_def - by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ -next - case False then have "r \ 0" "s \ t" by auto - have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" - by (simp add: algebra_simps) - have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 - \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" - by auto - have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ - (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" - by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] - intro: exI [where x = "-n" for n]) - have 1: "\s - t\ \ 2 * pi" - if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" - proof (rule ccontr) - assume "\ \s - t\ \ 2 * pi" - then have *: "\n. t - s \ of_int n * \s - t\" - using False that [of "2*pi / \t - s\"] - by (simp add: abs_minus_commute divide_simps) - show False - using * [of 1] * [of "-1"] by auto - qed - have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n - proof - - have "t-s = 2 * (real_of_int n * pi)/x" - using that by (simp add: field_simps) - then show ?thesis by (metis abs_minus_commute) - qed - have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" - by force - show ?thesis using False - apply (simp add: simple_path_def) - apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) - apply (subst abs_away) - apply (auto simp: 1) - apply (rule ccontr) - apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) - done -qed - -lemma arc_part_circlepath: - assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" - shows "arc (part_circlepath z r s t)" -proof - - have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" - and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n - proof (rule ccontr) - assume "x \ y" - have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" - by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_div_cancel_left eq) - then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" - by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) - with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" - by (force simp: field_simps) - have "\real_of_int n\ < \y - x\" - using assms \x \ y\ by (simp add: st abs_mult field_simps) - then show False - using assms x y st by (auto dest: of_int_lessD) - qed - show ?thesis - using assms - apply (simp add: arc_def) - apply (simp add: part_circlepath_def inj_on_def exp_eq) - apply (blast intro: *) - done -qed - -subsection\Special case of one complete circle\ - -definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" - where "circlepath z r \ part_circlepath z r 0 (2*pi)" - -lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" - by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) - -lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" - by (simp add: circlepath_def) - -lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" - by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) - -lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" -proof - - have "z + of_real r * exp (2 * pi * \ * (x + 1/2)) = - z + of_real r * exp (2 * pi * \ * x + pi * \)" - by (simp add: divide_simps) (simp add: algebra_simps) - also have "\ = z - r * exp (2 * pi * \ * x)" - by (simp add: exp_add) - finally show ?thesis - by (simp add: circlepath path_image_def sphere_def dist_norm) -qed - -lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" - using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] - by (simp add: add.commute) - -lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" - using circlepath_add1 [of z r "x-1/2"] - by (simp add: add.commute) - -lemma path_image_circlepath_minus_subset: - "path_image (circlepath z (-r)) \ path_image (circlepath z r)" - apply (simp add: path_image_def image_def circlepath_minus, clarify) - apply (case_tac "xa \ 1/2", force) - apply (force simp: circlepath_add_half)+ - done - -lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" - using path_image_circlepath_minus_subset by fastforce - -lemma has_vector_derivative_circlepath [derivative_intros]: - "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) - (at x within X)" - apply (simp add: circlepath_def scaleR_conv_of_real) - apply (rule derivative_eq_intros) - apply (simp add: algebra_simps) - done - -lemma vector_derivative_circlepath: - "vector_derivative (circlepath z r) (at x) = - 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" -using has_vector_derivative_circlepath vector_derivative_at by blast - -lemma vector_derivative_circlepath01: - "\0 \ x; x \ 1\ - \ vector_derivative (circlepath z r) (at x within {0..1}) = - 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" - using has_vector_derivative_circlepath - by (auto simp: vector_derivative_at_within_ivl) - -lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" - by (simp add: circlepath_def) - -lemma path_circlepath [simp]: "path (circlepath z r)" - by (simp add: valid_path_imp_path) - -lemma path_image_circlepath_nonneg: - assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" -proof - - have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x - proof (cases "x = z") - case True then show ?thesis by force - next - case False - define w where "w = x - z" - then have "w \ 0" by (simp add: False) - have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" - using cis_conv_exp complex_eq_iff by auto - show ?thesis - apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) - apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) - apply (rule_tac x="t / (2*pi)" in image_eqI) - apply (simp add: field_simps \w \ 0\) - using False ** - apply (auto simp: w_def) - done - qed - show ?thesis - unfolding circlepath path_image_def sphere_def dist_norm - by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) -qed - -lemma path_image_circlepath [simp]: - "path_image (circlepath z r) = sphere z \r\" - using path_image_circlepath_minus - by (force simp: path_image_circlepath_nonneg abs_if) - -lemma has_contour_integral_bound_circlepath_strong: - "\(f has_contour_integral i) (circlepath z r); - finite k; 0 \ B; 0 < r; - \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ - \ norm i \ B*(2*pi*r)" - unfolding circlepath_def - by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) - -lemma has_contour_integral_bound_circlepath: - "\(f has_contour_integral i) (circlepath z r); - 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ - \ norm i \ B*(2*pi*r)" - by (auto intro: has_contour_integral_bound_circlepath_strong) - -lemma contour_integrable_continuous_circlepath: - "continuous_on (path_image (circlepath z r)) f - \ f contour_integrable_on (circlepath z r)" - by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) - -lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" - by (simp add: circlepath_def simple_path_part_circlepath) - -lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" - by (simp add: sphere_def dist_norm norm_minus_commute) - -lemma contour_integral_circlepath: - assumes "r > 0" - shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" -proof (rule contour_integral_unique) - show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" - unfolding has_contour_integral_def using assms - apply (subst has_integral_cong) - apply (simp add: vector_derivative_circlepath01) - using has_integral_const_real [of _ 0 1] apply (force simp: circlepath) - done -qed - -lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" - apply (rule winding_number_unique_loop) - apply (simp_all add: sphere_def valid_path_imp_path) - apply (rule_tac x="circlepath z r" in exI) - apply (simp add: sphere_def contour_integral_circlepath) - done - -proposition winding_number_circlepath: - assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" -proof (cases "w = z") - case True then show ?thesis - using assms winding_number_circlepath_centre by auto -next - case False - have [simp]: "r > 0" - using assms le_less_trans norm_ge_zero by blast - define r' where "r' = norm(w - z)" - have "r' < r" - by (simp add: assms r'_def) - have disjo: "cball z r' \ sphere z r = {}" - using \r' < r\ by (force simp: cball_def sphere_def) - have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" - proof (rule winding_number_around_inside [where s = "cball z r'"]) - show "winding_number (circlepath z r) z \ 0" - by (simp add: winding_number_circlepath_centre) - show "cball z r' \ path_image (circlepath z r) = {}" - by (simp add: disjo less_eq_real_def) - qed (auto simp: r'_def dist_norm norm_minus_commute) - also have "\ = 1" - by (simp add: winding_number_circlepath_centre) - finally show ?thesis . -qed - - -text\ Hence the Cauchy formula for points inside a circle.\ - -theorem Cauchy_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r" - shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) - (circlepath z r)" -proof - - have "r > 0" - using assms le_less_trans norm_ge_zero by blast - have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) - (circlepath z r)" - proof (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) - show "\x. x \ interior (cball z r) - {} \ - f field_differentiable at x" - using holf holomorphic_on_imp_differentiable_at by auto - have "w \ sphere z r" - by simp (metis dist_commute dist_norm not_le order_refl wz) - then show "path_image (circlepath z r) \ cball z r - {w}" - using \r > 0\ by (auto simp add: cball_def sphere_def) - qed (use wz in \simp_all add: dist_norm norm_minus_commute contf\) - then show ?thesis - by (simp add: winding_number_circlepath assms) -qed - -corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: - assumes "f holomorphic_on cball z r" "norm(w - z) < r" - shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) - (circlepath z r)" -using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) - - -lemma no_bounded_connected_component_imp_winding_number_zero: - assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" - and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" - shows "winding_number g z = 0" -apply (rule winding_number_zero_in_outside) -apply (simp_all add: assms) -by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) - -lemma no_bounded_path_component_imp_winding_number_zero: - assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" - and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" - shows "winding_number g z = 0" -apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) -by (simp add: bounded_subset nb path_component_subset_connected_component) - - -subsection\ Uniform convergence of path integral\ - -text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ - -proposition contour_integral_uniform_limit: - assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" - and ul_f: "uniform_limit (path_image \) f l F" - and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" - and \: "valid_path \" - and [simp]: "\ trivial_limit F" - shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" -proof - - have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) - { fix e::real - assume "0 < e" - then have "0 < e / (\B\ + 1)" by simp - then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" - using ul_f [unfolded uniform_limit_iff dist_norm] by auto - with ev_fint - obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" - and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" - using eventually_happens [OF eventually_conj] - by (fastforce simp: contour_integrable_on path_image_def) - have Ble: "B * e / (\B\ + 1) \ e" - using \0 \ B\ \0 < e\ by (simp add: field_split_simps) - have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" - proof (intro exI conjI ballI) - show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" - if "x \ {0..1}" for x - apply (rule order_trans [OF _ Ble]) - using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ - apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) - apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le]) - done - qed (rule inta) - } - then show lintg: "l contour_integrable_on \" - unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) - { fix e::real - define B' where "B' = B + 1" - have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) - assume "0 < e" - then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" - using ul_f [unfolded uniform_limit_iff dist_norm, rule_format, of "e / B' / 2"] B' - by (simp add: field_simps) - have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp - have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" - if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t - proof - - have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" - using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto - also have "\ < e" - by (simp add: B' \0 < e\ mult_imp_div_pos_less) - finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . - then show ?thesis - by (simp add: left_diff_distrib [symmetric] norm_mult) - qed - have le_e: "\x. \\xa\{0..1}. 2 * cmod (f x (\ xa) - l (\ xa)) < e / B'; f x contour_integrable_on \\ - \ cmod (integral {0..1} - (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" - apply (rule le_less_trans [OF integral_norm_bound_integral ie]) - apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) - apply (blast intro: *)+ - done - have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" - apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) - apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) - apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric] le_e) - done - } - then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" - by (rule tendstoI) -qed - -corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: - assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" - and "uniform_limit (sphere z r) f l F" - and "\ trivial_limit F" "0 < r" - shows "l contour_integrable_on (circlepath z r)" - "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" - using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) - - -subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ - -lemma Cauchy_next_derivative: - assumes "continuous_on (path_image \) f'" - and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" - and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" - and k: "k \ 0" - and "open s" - and \: "valid_path \" - and w: "w \ s - path_image \" - shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" - and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) - (at w)" (is "?thes2") -proof - - have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast - then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w - using open_contains_ball by blast - have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" - by (metis norm_of_nat of_nat_Suc) - have cint: "\x. \x \ w; cmod (x - w) < d\ - \ (\z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \" - apply (rule contour_integrable_div [OF contour_integrable_diff]) - using int w d - by (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ - have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) - contour_integrable_on \" - unfolding eventually_at - apply (rule_tac x=d in exI) - apply (simp add: \d > 0\ dist_norm field_simps cint) - done - have bim_g: "bounded (image f' (path_image \))" - by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) - then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" - by (force simp: bounded_pos path_image_def) - have twom: "\\<^sub>F n in at w. - \x\path_image \. - cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" - if "0 < e" for e - proof - - have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" - if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" - and uw_less: "cmod (u - w) < e * (d/2) ^ (k+2) / (1 + real k)" - for u x - proof - - define ff where [abs_def]: - "ff n w = - (if n = 0 then inverse(x - w)^k - else if n = 1 then k / (x - w)^(Suc k) - else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w - have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" - by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) - have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d/2))" - if "z \ ball w (d/2)" "i \ 1" for i z - proof - - have "z \ path_image \" - using \x \ path_image \\ d that ball_divide_subset_numeral by blast - then have xz[simp]: "x \ z" using \x \ path_image \\ by blast - then have neq: "x * x + z * z \ x * (z * 2)" - by (blast intro: dest!: sum_sqs_eq) - with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto - then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" - by (simp add: algebra_simps) - show ?thesis using \i \ 1\ - apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) - apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ - done - qed - { fix a::real and b::real assume ab: "a > 0" "b > 0" - then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" - by (subst mult_le_cancel_left_pos) - (use \k \ 0\ in \auto simp: divide_simps\) - with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" - by (simp add: field_simps) - } note canc = this - have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d/2) ^ (k + 2)" - if "v \ ball w (d/2)" for v - proof - - have lessd: "\z. cmod (\ z - v) < d/2 \ cmod (w - \ z) < d" - by (metis that norm_minus_commute norm_triangle_half_r dist_norm mem_ball) - have "d/2 \ cmod (x - v)" using d x that - using lessd d x - by (auto simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps) - then have "d \ cmod (x - v) * 2" - by (simp add: field_split_simps) - then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" - using \0 < d\ order_less_imp_le power_mono by blast - have "x \ v" using that - using \x \ path_image \\ ball_divide_subset_numeral d by fastforce - then show ?thesis - using \d > 0\ apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) - using dpow_le apply (simp add: field_split_simps) - done - qed - have ub: "u \ ball w (d/2)" - using uwd by (simp add: dist_commute dist_norm) - have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d/2) ^ k))" - using complex_Taylor [OF _ ff1 ff2 _ ub, of w, simplified] - by (simp add: ff_def \0 < d\) - then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" - by (simp add: field_simps) - then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) - / (cmod (u - w) * real k) - \ (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" - using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) - also have "\ < e" - using uw_less \0 < d\ by (simp add: mult_ac divide_simps) - finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) - / cmod ((u - w) * real k) < e" - by (simp add: norm_mult) - have "x \ u" - using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) - show ?thesis - apply (rule le_less_trans [OF _ e]) - using \k \ 0\ \x \ u\ \u \ w\ - apply (simp add: field_simps norm_divide [symmetric]) - done - qed - show ?thesis - unfolding eventually_at - apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) - apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) - done - qed - have 2: "uniform_limit (path_image \) (\n x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) (\x. f' x / (x - w) ^ Suc k) (at w)" - unfolding uniform_limit_iff dist_norm - proof clarify - fix e::real - assume "0 < e" - have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" - if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" - and x: "0 \ x" "x \ 1" - for u x - proof (cases "(f' (\ x)) = 0") - case True then show ?thesis by (simp add: \0 < e\) - next - case False - have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = - cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k))" - by (simp add: field_simps) - also have "\ = cmod (f' (\ x)) * - cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - - inverse (\ x - w) * inverse (\ x - w) ^ k)" - by (simp add: norm_mult) - also have "\ < cmod (f' (\ x)) * (e/C)" - using False mult_strict_left_mono [OF ec] by force - also have "\ \ e" using C - by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) - finally show ?thesis . - qed - show "\\<^sub>F n in at w. - \x\path_image \. - cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" - using twom [OF divide_pos_pos [OF \0 < e\ \C > 0\]] unfolding path_image_def - by (force intro: * elim: eventually_mono) - qed - show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" - by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto - have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) - \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" - by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto - have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = - (f u - f w) / (u - w) / k" - if "dist u w < d" for u - proof - - have u: "u \ s - path_image \" - by (metis subsetD d dist_commute mem_ball that) - show ?thesis - apply (rule contour_integral_unique) - apply (simp add: diff_divide_distrib algebra_simps) - apply (intro has_contour_integral_diff has_contour_integral_div) - using u w apply (simp_all add: field_simps int) - done - qed - show ?thes2 - apply (simp add: has_field_derivative_iff del: power_Suc) - apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) - apply (simp add: \k \ 0\ **) - done -qed - -lemma Cauchy_next_derivative_circlepath: - assumes contf: "continuous_on (path_image (circlepath z r)) f" - and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" - and k: "k \ 0" - and w: "w \ ball z r" - shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" - (is "?thes2") -proof - - have "r > 0" using w - using ball_eq_empty by fastforce - have wim: "w \ ball z r - path_image (circlepath z r)" - using w by (auto simp: dist_norm) - show ?thes1 ?thes2 - by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; - auto simp: vector_derivative_circlepath norm_mult)+ -qed - - -text\ In particular, the first derivative formula.\ - -lemma Cauchy_derivative_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" - (is "?thes2") -proof - - have [simp]: "r \ 0" using w - using ball_eq_empty by fastforce - have f: "continuous_on (path_image (circlepath z r)) f" - by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def) - have int: "\w. dist z w < r \ - ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" - by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) - show ?thes1 - apply (simp add: power2_eq_square) - apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) - apply (blast intro: int) - done - have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" - apply (simp add: power2_eq_square) - apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) - apply (blast intro: int) - done - then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" - by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) - show ?thes2 - by simp (rule fder) -qed - -subsection\Existence of all higher derivatives\ - -proposition derivative_is_holomorphic: - assumes "open S" - and fder: "\z. z \ S \ (f has_field_derivative f' z) (at z)" - shows "f' holomorphic_on S" -proof - - have *: "\h. (f' has_field_derivative h) (at z)" if "z \ S" for z - proof - - obtain r where "r > 0" and r: "cball z r \ S" - using open_contains_cball \z \ S\ \open S\ by blast - then have holf_cball: "f holomorphic_on cball z r" - apply (simp add: holomorphic_on_def) - using field_differentiable_at_within field_differentiable_def fder by blast - then have "continuous_on (path_image (circlepath z r)) f" - using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) - then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" - by (auto intro: continuous_intros)+ - have contf_cball: "continuous_on (cball z r) f" using holf_cball - by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) - have holf_ball: "f holomorphic_on ball z r" using holf_cball - using ball_subset_cball holomorphic_on_subset by blast - { fix w assume w: "w \ ball z r" - have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" - by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) - have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) - (at w)" - by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) - have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" - using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) - have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral - contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) - (circlepath z r)" - by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) - then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral - contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) - (circlepath z r)" - by (simp add: algebra_simps) - then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" - by (simp add: f'_eq) - } note * = this - show ?thesis - apply (rule exI) - apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) - apply (simp_all add: \0 < r\ * dist_norm) - done - qed - show ?thesis - by (simp add: holomorphic_on_open [OF \open S\] *) -qed - -lemma holomorphic_deriv [holomorphic_intros]: - "\f holomorphic_on S; open S\ \ (deriv f) holomorphic_on S" -by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) - -lemma analytic_deriv [analytic_intros]: "f analytic_on S \ (deriv f) analytic_on S" - using analytic_on_holomorphic holomorphic_deriv by auto - -lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv ^^ n) f holomorphic_on S" - by (induction n) (auto simp: holomorphic_deriv) - -lemma analytic_higher_deriv [analytic_intros]: "f analytic_on S \ (deriv ^^ n) f analytic_on S" - unfolding analytic_on_def using holomorphic_higher_deriv by blast - -lemma has_field_derivative_higher_deriv: - "\f holomorphic_on S; open S; x \ S\ - \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" -by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply - funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) - -lemma valid_path_compose_holomorphic: - assumes "valid_path g" and holo:"f holomorphic_on S" and "open S" "path_image g \ S" - shows "valid_path (f \ g)" -proof (rule valid_path_compose[OF \valid_path g\]) - fix x assume "x \ path_image g" - then show "f field_differentiable at x" - using analytic_on_imp_differentiable_at analytic_on_open assms holo by blast -next - have "deriv f holomorphic_on S" - using holomorphic_deriv holo \open S\ by auto - then show "continuous_on (path_image g) (deriv f)" - using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto -qed - - -subsection\Morera's theorem\ - -lemma Morera_local_triangle_ball: - assumes "\z. z \ S - \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ - (\b c. closed_segment b c \ ball a e - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0)" - shows "f analytic_on S" -proof - - { fix z assume "z \ S" - with assms obtain e a where - "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" - and 0: "\b c. closed_segment b c \ ball a e - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - by fastforce - have az: "dist a z < e" using mem_ball z by blast - have sb_ball: "ball z (e - dist a z) \ ball a e" - by (simp add: dist_commute ball_subset_ball_iff) - have "\e>0. f holomorphic_on ball z e" - proof (intro exI conjI) - have sub_ball: "\y. dist a y < e \ closed_segment a y \ ball a e" - by (meson \0 < e\ centre_in_ball convex_ball convex_contains_segment mem_ball) - show "f holomorphic_on ball z (e - dist a z)" - apply (rule holomorphic_on_subset [OF _ sb_ball]) - apply (rule derivative_is_holomorphic[OF open_ball]) - apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) - apply (simp_all add: 0 \0 < e\ sub_ball) - done - qed (simp add: az) - } - then show ?thesis - by (simp add: analytic_on_def) -qed - -lemma Morera_local_triangle: - assumes "\z. z \ S - \ \t. open t \ z \ t \ continuous_on t f \ - (\a b c. convex hull {a,b,c} \ t - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0)" - shows "f analytic_on S" -proof - - { fix z assume "z \ S" - with assms obtain t where - "open t" and z: "z \ t" and contf: "continuous_on t f" - and 0: "\a b c. convex hull {a,b,c} \ t - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" - by force - then obtain e where "e>0" and e: "ball z e \ t" - using open_contains_ball by blast - have [simp]: "continuous_on (ball z e) f" using contf - using continuous_on_subset e by blast - have eq0: "\b c. closed_segment b c \ ball z e \ - contour_integral (linepath z b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c z) f = 0" - by (meson 0 z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) - have "\e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ - (\b c. closed_segment b c \ ball a e \ - contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" - using \e > 0\ eq0 by force - } - then show ?thesis - by (simp add: Morera_local_triangle_ball) -qed - -proposition Morera_triangle: - "\continuous_on S f; open S; - \a b c. convex hull {a,b,c} \ S - \ contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0\ - \ f analytic_on S" - using Morera_local_triangle by blast - -subsection\Combining theorems for higher derivatives including Leibniz rule\ - -lemma higher_deriv_linear [simp]: - "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" - by (induction n) auto - -lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" - by (induction n) auto - -lemma higher_deriv_ident [simp]: - "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" - apply (induction n, simp) - apply (metis higher_deriv_linear lambda_one) - done - -lemma higher_deriv_id [simp]: - "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" - by (simp add: id_def) - -lemma has_complex_derivative_funpow_1: - "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" - apply (induction n, auto) - apply (simp add: id_def) - by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) - -lemma higher_deriv_uminus: - assumes "f holomorphic_on S" "open S" and z: "z \ S" - shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - have "((deriv ^^ n) (\w. - f w) has_field_derivative - deriv ((deriv ^^ n) f) z) (at z)" - apply (rule has_field_derivative_transform_within_open [of "\w. -((deriv ^^ n) f w)"]) - apply (rule derivative_eq_intros | rule * refl assms)+ - apply (auto simp add: Suc) - done - then show ?case - by (simp add: DERIV_imp_deriv) -qed - -lemma higher_deriv_add: - fixes z::complex - assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" - shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - have "((deriv ^^ n) (\w. f w + g w) has_field_derivative - deriv ((deriv ^^ n) f) z + deriv ((deriv ^^ n) g) z) (at z)" - apply (rule has_field_derivative_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) - apply (rule derivative_eq_intros | rule * refl assms)+ - apply (auto simp add: Suc) - done - then show ?case - by (simp add: DERIV_imp_deriv) -qed - -lemma higher_deriv_diff: - fixes z::complex - assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" - shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" - apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add) - apply (subst higher_deriv_add) - using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus) - done - -lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" - by (cases k) simp_all - -lemma higher_deriv_mult: - fixes z::complex - assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" - shows "(deriv ^^ n) (\w. f w * g w) z = - (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" - "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - have sumeq: "(\i = 0..n. - of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = - g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" - apply (simp add: bb algebra_simps sum.distrib) - apply (subst (4) sum_Suc_reindex) - apply (auto simp: algebra_simps Suc_diff_le intro: sum.cong) - done - have "((deriv ^^ n) (\w. f w * g w) has_field_derivative - (\i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z)) - (at z)" - apply (rule has_field_derivative_transform_within_open - [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) - apply (simp add: algebra_simps) - apply (rule DERIV_cong [OF DERIV_sum]) - apply (rule DERIV_cmult) - apply (auto intro: DERIV_mult * sumeq \open S\ Suc.prems Suc.IH [symmetric]) - done - then show ?case - unfolding funpow.simps o_apply - by (simp add: DERIV_imp_deriv) -qed - -lemma higher_deriv_transform_within_open: - fixes z::complex - assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" - and fg: "\w. w \ S \ f w = g w" - shows "(deriv ^^ i) f z = (deriv ^^ i) g z" -using z -by (induction i arbitrary: z) - (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) - -lemma higher_deriv_compose_linear: - fixes z::complex - assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \ S" - and fg: "\w. w \ S \ u * w \ T" - shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" -using z -proof (induction n arbitrary: z) - case 0 then show ?case by simp -next - case (Suc n z) - have holo0: "f holomorphic_on (*) u ` S" - by (meson fg f holomorphic_on_subset image_subset_iff) - have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S" - by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T) - have holo3: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S" - by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros) - have holo1: "(\w. f (u * w)) holomorphic_on S" - apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) - apply (rule holo0 holomorphic_intros)+ - done - have "deriv ((deriv ^^ n) (\w. f (u * w))) z = deriv (\z. u^n * (deriv ^^ n) f (u*z)) z" - apply (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems]) - apply (rule holomorphic_higher_deriv [OF holo1 S]) - apply (simp add: Suc.IH) - done - also have "\ = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" - apply (rule deriv_cmult) - apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) - apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and T=T, unfolded o_def]) - apply (simp) - apply (simp add: analytic_on_open f holomorphic_higher_deriv T) - apply (blast intro: fg) - done - also have "\ = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" - apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "(*) u", unfolded o_def]) - apply (rule derivative_intros) - using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv T apply blast - apply (simp) - done - finally show ?case - by simp -qed - -lemma higher_deriv_add_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_add show ?thesis - by (auto simp: analytic_at_two) -qed - -lemma higher_deriv_diff_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_diff show ?thesis - by (auto simp: analytic_at_two) -qed - -lemma higher_deriv_uminus_at: - "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" - using higher_deriv_uminus - by (auto simp: analytic_at) - -lemma higher_deriv_mult_at: - assumes "f analytic_on {z}" "g analytic_on {z}" - shows "(deriv ^^ n) (\w. f w * g w) z = - (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" -proof - - have "f analytic_on {z} \ g analytic_on {z}" - using assms by blast - with higher_deriv_mult show ?thesis - by (auto simp: analytic_at_two) -qed - - -text\ Nonexistence of isolated singularities and a stronger integral formula.\ - -proposition no_isolated_singularity: - fixes z::complex - assumes f: "continuous_on S f" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" - shows "f holomorphic_on S" -proof - - { fix z - assume "z \ S" and cdf: "\x. x \ S - K \ f field_differentiable at x" - have "f field_differentiable at z" - proof (cases "z \ K") - case False then show ?thesis by (blast intro: cdf \z \ S\) - next - case True - with finite_set_avoid [OF K, of z] - obtain d where "d>0" and d: "\x. \x\K; x \ z\ \ d \ dist z x" - by blast - obtain e where "e>0" and e: "ball z e \ S" - using S \z \ S\ by (force simp: open_contains_ball) - have fde: "continuous_on (ball z (min d e)) f" - by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) - have cont: "{a,b,c} \ ball z (min d e) \ continuous_on (convex hull {a, b, c}) f" for a b c - by (simp add: hull_minimal continuous_on_subset [OF fde]) - have fd: "\{a,b,c} \ ball z (min d e); x \ interior (convex hull {a, b, c}) - K\ - \ f field_differentiable at x" for a b c x - by (metis cdf Diff_iff Int_iff ball_min_Int subsetD convex_ball e interior_mono interior_subset subset_hull) - obtain g where "\w. w \ ball z (min d e) \ (g has_field_derivative f w) (at w within ball z (min d e))" - apply (rule contour_integral_convex_primitive - [OF convex_ball fde Cauchy_theorem_triangle_cofinite [OF _ K]]) - using cont fd by auto - then have "f holomorphic_on ball z (min d e)" - by (metis open_ball at_within_open derivative_is_holomorphic) - then show ?thesis - unfolding holomorphic_on_def - by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) - qed - } - with holf S K show ?thesis - by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) -qed - -lemma no_isolated_singularity': - fixes z::complex - assumes f: "\z. z \ K \ (f \ f z) (at z within S)" - and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" - shows "f holomorphic_on S" -proof (rule no_isolated_singularity[OF _ assms(2-)]) - show "continuous_on S f" unfolding continuous_on_def - proof - fix z assume z: "z \ S" - show "(f \ f z) (at z within S)" - proof (cases "z \ K") - case False - from holf have "continuous_on (S - K) f" - by (rule holomorphic_on_imp_continuous_on) - with z False have "(f \ f z) (at z within (S - K))" - by (simp add: continuous_on_def) - also from z K S False have "at z within (S - K) = at z within S" - by (subst (1 2) at_within_open) (auto intro: finite_imp_closed) - finally show "(f \ f z) (at z within S)" . - qed (insert assms z, simp_all) - qed -qed - -proposition Cauchy_integral_formula_convex: - assumes S: "convex S" and K: "finite K" and contf: "continuous_on S f" - and fcd: "(\x. x \ interior S - K \ f field_differentiable at x)" - and z: "z \ interior S" and vpg: "valid_path \" - and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - have *: "\x. x \ interior S \ f field_differentiable at x" - unfolding holomorphic_on_open [symmetric] field_differentiable_def - using no_isolated_singularity [where S = "interior S"] - by (meson K contf continuous_at_imp_continuous_on continuous_on_interior fcd - field_differentiable_at_within field_differentiable_def holomorphic_onI - holomorphic_on_imp_differentiable_at open_interior) - show ?thesis - by (rule Cauchy_integral_formula_weak [OF S finite.emptyI contf]) (use * assms in auto) -qed - -text\ Formula for higher derivatives.\ - -lemma Cauchy_has_contour_integral_higher_derivative_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) - (circlepath z r)" -using w -proof (induction k arbitrary: w) - case 0 then show ?case - using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) -next - case (Suc k) - have [simp]: "r > 0" using w - using ball_eq_empty by fastforce - have f: "continuous_on (path_image (circlepath z r)) f" - by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def less_imp_le) - obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" - using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] - by (auto simp: contour_integrable_on_def) - then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" - by (rule contour_integral_unique) - have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" - using Suc.prems assms has_field_derivative_higher_deriv by auto - then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" - by (force simp: field_differentiable_def) - have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = - of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" - by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) - also have "\ = of_nat (Suc k) * X" - by (simp only: con) - finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . - then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" - by (metis deriv_cmult dnf_diff) - then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" - by (simp add: field_simps) - then show ?case - using of_nat_eq_0_iff X by fastforce -qed - -lemma Cauchy_higher_derivative_integral_circlepath: - assumes contf: "continuous_on (cball z r) f" - and holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" - (is "?thes1") - and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" - (is "?thes2") -proof - - have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) - (circlepath z r)" - using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] - by simp - show ?thes1 using * - using contour_integrable_on_def by blast - show ?thes2 - unfolding contour_integral_unique [OF *] by (simp add: field_split_simps) -qed - -corollary Cauchy_contour_integral_circlepath: - assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" - shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" -by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) - -lemma Cauchy_contour_integral_circlepath_2: - assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" - shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" - using Cauchy_contour_integral_circlepath [OF assms, of 1] - by (simp add: power2_eq_square) - - -subsection\A holomorphic function is analytic, i.e. has local power series\ - -theorem holomorphic_power_series: - assumes holf: "f holomorphic_on ball z r" - and w: "w \ ball z r" - shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" -proof - - \ \Replacing \<^term>\r\ and the original (weak) premises with stronger ones\ - obtain r where "r > 0" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" - proof - have "cball z ((r + dist w z) / 2) \ ball z r" - using w by (simp add: dist_commute field_sum_of_halves subset_eq) - then show "f holomorphic_on cball z ((r + dist w z) / 2)" - by (rule holomorphic_on_subset [OF holf]) - have "r > 0" - using w by clarsimp (metis dist_norm le_less_trans norm_ge_zero) - then show "0 < (r + dist w z) / 2" - by simp (use zero_le_dist [of w z] in linarith) - qed (use w in \auto simp: dist_commute\) - then have holf: "f holomorphic_on ball z r" - using ball_subset_cball holomorphic_on_subset by blast - have contf: "continuous_on (cball z r) f" - by (simp add: holfc holomorphic_on_imp_continuous_on) - have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" - by (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) (simp add: \0 < r\) - obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" - by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) - obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" - and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" - proof - show "\u. cmod (u - z) = r \ r - dist z w \ cmod (u - w)" - by (metis add_diff_eq diff_add_cancel dist_norm norm_diff_ineq) - qed (use w in \auto simp: dist_norm norm_minus_commute\) - have ul: "uniform_limit (sphere z r) (\n x. (\kx. f x / (x - w)) sequentially" - unfolding uniform_limit_iff dist_norm - proof clarify - fix e::real - assume "0 < e" - have rr: "0 \ (r - k) / r" "(r - k) / r < 1" using k by auto - obtain n where n: "((r - k) / r) ^ n < e / B * k" - using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force - have "norm ((\k N" and r: "r = dist z u" for N u - proof - - have N: "((r - k) / r) ^ N < e / B * k" - apply (rule le_less_trans [OF power_decreasing n]) - using \n \ N\ k by auto - have u [simp]: "(u \ z) \ (u \ w)" - using \0 < r\ r w by auto - have wzu_not1: "(w - z) / (u - z) \ 1" - by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) - have "norm ((\kk = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" - using \0 < B\ - apply (auto simp: geometric_sum [OF wzu_not1]) - apply (simp add: field_simps norm_mult [symmetric]) - done - also have "\ = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" - using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) - also have "\ = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" - by (simp add: algebra_simps) - also have "\ = norm (w - z) ^ N * norm (f u) / r ^ N" - by (simp add: norm_mult norm_power norm_minus_commute) - also have "\ \ (((r - k)/r)^N) * B" - using \0 < r\ w k - apply (simp add: divide_simps) - apply (rule mult_mono [OF power_mono]) - apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) - done - also have "\ < e * k" - using \0 < B\ N by (simp add: divide_simps) - also have "\ \ e * norm (u - w)" - using r kle \0 < e\ by (simp add: dist_commute dist_norm) - finally show ?thesis - by (simp add: field_split_simps norm_divide del: power_Suc) - qed - with \0 < r\ show "\\<^sub>F n in sequentially. \x\sphere z r. - norm ((\k\<^sub>F x in sequentially. - contour_integral (circlepath z r) (\u. \kku. f u / (u - z) ^ Suc k) * (w - z) ^ k)" - apply (rule eventuallyI) - apply (subst contour_integral_sum, simp) - using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) - apply (simp only: contour_integral_lmul cint algebra_simps) - done - have cic: "\u. (\y. \k0 < r\ by (force intro!: Cauchy_higher_derivative_integral_circlepath [OF contf holf]) - have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) - sums contour_integral (circlepath z r) (\u. f u/(u - w))" - unfolding sums_def - apply (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul] cic) - using \0 < r\ apply auto - done - then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) - sums (2 * of_real pi * \ * f w)" - using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) - then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) - sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" - by (rule sums_divide) - then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) - sums f w" - by (simp add: field_simps) - then show ?thesis - by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) -qed - - -subsection\The Liouville theorem and the Fundamental Theorem of Algebra\ - -text\ These weak Liouville versions don't even need the derivative formula.\ - -lemma Liouville_weak_0: - assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" - shows "f z = 0" -proof (rule ccontr) - assume fz: "f z \ 0" - with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] - obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" - by (auto simp: dist_norm) - define R where "R = 1 + \B\ + norm z" - have "R > 0" unfolding R_def - proof - - have "0 \ cmod z + \B\" - by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) - then show "0 < 1 + \B\ + cmod z" - by linarith - qed - have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" - apply (rule Cauchy_integral_circlepath) - using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ - done - have "cmod (x - z) = R \ cmod (f x) * 2 < cmod (f z)" for x - unfolding R_def - by (rule B) (use norm_triangle_ineq4 [of x z] in auto) - with \R > 0\ fz show False - using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] - by (auto simp: less_imp_le norm_mult norm_divide field_split_simps) -qed - -proposition Liouville_weak: - assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" - shows "f z = l" - using Liouville_weak_0 [of "\z. f z - l"] - by (simp add: assms holomorphic_on_diff LIM_zero) - -proposition Liouville_weak_inverse: - assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" - obtains z where "f z = 0" -proof - - { assume f: "\z. f z \ 0" - have 1: "(\x. 1 / f x) holomorphic_on UNIV" - by (simp add: holomorphic_on_divide assms f) - have 2: "((\x. 1 / f x) \ 0) at_infinity" - apply (rule tendstoI [OF eventually_mono]) - apply (rule_tac B="2/e" in unbounded) - apply (simp add: dist_norm norm_divide field_split_simps) - done - have False - using Liouville_weak_0 [OF 1 2] f by simp - } - then show ?thesis - using that by blast -qed - -text\ In particular we get the Fundamental Theorem of Algebra.\ - -theorem fundamental_theorem_of_algebra: - fixes a :: "nat \ complex" - assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" - obtains z where "(\i\n. a i * z^i) = 0" -using assms -proof (elim disjE bexE) - assume "a 0 = 0" then show ?thesis - by (auto simp: that [of 0]) -next - fix i - assume i: "i \ {1..n}" and nz: "a i \ 0" - have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" - by (rule holomorphic_intros)+ - show thesis - proof (rule Liouville_weak_inverse [OF 1]) - show "\\<^sub>F x in at_infinity. B \ cmod (\i\n. a i * x ^ i)" for B - using i polyfun_extremal nz by force - qed (use that in auto) -qed - -subsection\Weierstrass convergence theorem\ - -lemma holomorphic_uniform_limit: - assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" - and ulim: "uniform_limit (cball z r) f g F" - and F: "\ trivial_limit F" - obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" -proof (cases r "0::real" rule: linorder_cases) - case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) -next - case equal then show ?thesis - by (force simp: holomorphic_on_def intro: that) -next - case greater - have contg: "continuous_on (cball z r) g" - using cont uniform_limit_theorem [OF eventually_mono ulim F] by blast - have "path_image (circlepath z r) \ cball z r" - using \0 < r\ by auto - then have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" - by (intro continuous_intros continuous_on_subset [OF contg]) - have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" - if w: "w \ ball z r" for w - proof - - define d where "d = (r - norm(w - z))" - have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) - have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" - unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) - have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" - apply (rule eventually_mono [OF cont]) - using w - apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) - done - have ul_less: "uniform_limit (sphere z r) (\n x. f n x / (x - w)) (\x. g x / (x - w)) F" - using greater \0 < d\ - apply (clarsimp simp add: uniform_limit_iff dist_norm norm_divide diff_divide_distrib [symmetric] divide_simps) - apply (rule_tac e1="e * d" in eventually_mono [OF uniform_limitD [OF ulim]]) - apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+ - done - have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" - by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) - have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" - by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) - have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" - proof (rule Lim_transform_eventually) - show "\\<^sub>F x in F. contour_integral (circlepath z r) (\u. f x u / (u - w)) - = 2 * of_real pi * \ * f x w" - apply (rule eventually_mono [OF cont contour_integral_unique [OF Cauchy_integral_circlepath]]) - using w\0 < d\ d_def by auto - qed (auto simp: cif_tends_cig) - have "\e. 0 < e \ \\<^sub>F n in F. dist (f n w) (g w) < e" - by (rule eventually_mono [OF uniform_limitD [OF ulim]]) (use w in auto) - then have "((\n. 2 * of_real pi * \ * f n w) \ 2 * of_real pi * \ * g w) F" - by (rule tendsto_mult_left [OF tendstoI]) - then have "((\u. g u / (u - w)) has_contour_integral 2 * of_real pi * \ * g w) (circlepath z r)" - using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w - by fastforce - then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" - using has_contour_integral_div [where c = "2 * of_real pi * \"] - by (force simp: field_simps) - then show ?thesis - by (simp add: dist_norm) - qed - show ?thesis - using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] - by (fastforce simp add: holomorphic_on_open contg intro: that) -qed - - -text\ Version showing that the limit is the limit of the derivatives.\ - -proposition has_complex_derivative_uniform_limit: - fixes z::complex - assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ - (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" - and ulim: "uniform_limit (cball z r) f g F" - and F: "\ trivial_limit F" and "0 < r" - obtains g' where - "continuous_on (cball z r) g" - "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" -proof - - let ?conint = "contour_integral (circlepath z r)" - have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" - by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] ulim F]; - auto simp: holomorphic_on_open field_differentiable_def)+ - then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" - using DERIV_deriv_iff_has_field_derivative - by (fastforce simp add: holomorphic_on_open) - then have derg: "\x. x \ ball z r \ deriv g x = g' x" - by (simp add: DERIV_imp_deriv) - have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w - proof - - have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" - if cont_fn: "continuous_on (cball z r) (f n)" - and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n - proof - - have hol_fn: "f n holomorphic_on ball z r" - using fnd by (force simp: holomorphic_on_open) - have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" - by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) - then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" - using DERIV_unique [OF fnd] w by blast - show ?thesis - by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] field_split_simps) - qed - define d where "d = (r - norm(w - z))^2" - have "d > 0" - using w by (simp add: dist_commute dist_norm d_def) - have dle: "d \ cmod ((y - w)\<^sup>2)" if "r = cmod (z - y)" for y - proof - - have "w \ ball z (cmod (z - y))" - using that w by fastforce - then have "cmod (w - z) \ cmod (z - y)" - by (simp add: dist_complex_def norm_minus_commute) - moreover have "cmod (z - y) - cmod (w - z) \ cmod (y - w)" - by (metis diff_add_cancel diff_add_eq_diff_diff_swap norm_minus_commute norm_triangle_ineq2) - ultimately show ?thesis - using that by (simp add: d_def norm_power power_mono) - qed - have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" - by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) - have 2: "uniform_limit (sphere z r) (\n x. f n x / (x - w)\<^sup>2) (\x. g x / (x - w)\<^sup>2) F" - unfolding uniform_limit_iff - proof clarify - fix e::real - assume "0 < e" - with \r > 0\ show "\\<^sub>F n in F. \x\sphere z r. dist (f n x / (x - w)\<^sup>2) (g x / (x - w)\<^sup>2) < e" - apply (simp add: norm_divide field_split_simps sphere_def dist_norm) - apply (rule eventually_mono [OF uniform_limitD [OF ulim], of "e*d"]) - apply (simp add: \0 < d\) - apply (force simp: dist_norm dle intro: less_le_trans) - done - qed - have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) - \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" - by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) - then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" - using Lim_null by (force intro!: tendsto_mult_right_zero) - have "((\n. f' n w - g' w) \ 0) F" - apply (rule Lim_transform_eventually [OF tendsto_0]) - apply (force simp: divide_simps intro: eq_f' eventually_mono [OF cont]) - done - then show ?thesis using Lim_null by blast - qed - obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" - by (blast intro: tends_f'n_g' g') - then show ?thesis using g - using that by blast -qed - - -subsection\<^marker>\tag unimportant\ \Some more simple/convenient versions for applications\ - -lemma holomorphic_uniform_sequence: - assumes S: "open S" - and hol_fn: "\n. (f n) holomorphic_on S" - and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" - shows "g holomorphic_on S" -proof - - have "\f'. (g has_field_derivative f') (at z)" if "z \ S" for z - proof - - obtain r where "0 < r" and r: "cball z r \ S" - and ul: "uniform_limit (cball z r) f g sequentially" - using ulim_g [OF \z \ S\] by blast - have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" - proof (intro eventuallyI conjI) - show "continuous_on (cball z r) (f x)" for x - using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast - show "f x holomorphic_on ball z r" for x - by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) - qed - show ?thesis - apply (rule holomorphic_uniform_limit [OF *]) - using \0 < r\ centre_in_ball ul - apply (auto simp: holomorphic_on_open) - done - qed - with S show ?thesis - by (simp add: holomorphic_on_open) -qed - -lemma has_complex_derivative_uniform_sequence: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ ((f n) has_field_derivative f' n x) (at x)" - and ulim_g: "\x. x \ S - \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" - shows "\g'. \x \ S. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" -proof - - have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ S" for z - proof - - obtain r where "0 < r" and r: "cball z r \ S" - and ul: "uniform_limit (cball z r) f g sequentially" - using ulim_g [OF \z \ S\] by blast - have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ - (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" - proof (intro eventuallyI conjI ballI) - show "continuous_on (cball z r) (f x)" for x - by (meson S continuous_on_subset hfd holomorphic_on_imp_continuous_on holomorphic_on_open r) - show "w \ ball z r \ (f x has_field_derivative f' x w) (at w)" for w x - using ball_subset_cball hfd r by blast - qed - show ?thesis - by (rule has_complex_derivative_uniform_limit [OF *, of g]) (use \0 < r\ ul in \force+\) - qed - show ?thesis - by (rule bchoice) (blast intro: y) -qed - -subsection\On analytic functions defined by a series\ - -lemma series_and_derivative_comparison: - fixes S :: "complex set" - assumes S: "open S" - and h: "summable h" - and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\\<^sub>F n in sequentially. \x\S. norm (f n x) \ h n" - obtains g g' where "\x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -proof - - obtain g where g: "uniform_limit S (\n x. \id>0. cball x d \ S \ uniform_limit (cball x d) (\n x. \i S" for x - proof - - obtain d where "d>0" and d: "cball x d \ S" - using open_contains_cball [of "S"] \x \ S\ S by blast - show ?thesis - proof (intro conjI exI) - show "uniform_limit (cball x d) (\n x. \id > 0\ d in auto) - qed - have "\x. x \ S \ (\n. \i g x" - by (metis tendsto_uniform_limitI [OF g]) - moreover have "\g'. \x\S. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" - by (rule has_complex_derivative_uniform_sequence [OF S]) (auto intro: * hfd DERIV_sum)+ - ultimately show ?thesis - by (metis sums_def that) -qed - -text\A version where we only have local uniform/comparative convergence.\ - -lemma series_and_derivative_comparison_local: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ (\\<^sub>F n in sequentially. \y\ball x d \ S. norm (f n y) \ h n)" - shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -proof - - have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" - if "z \ S" for z - proof - - obtain d h where "0 < d" "summable h" and le_h: "\\<^sub>F n in sequentially. \y\ball z d \ S. norm (f n y) \ h n" - using to_g \z \ S\ by meson - then obtain r where "r>0" and r: "ball z r \ ball z d \ S" using \z \ S\ S - by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) - have 1: "open (ball z d \ S)" - by (simp add: open_Int S) - have 2: "\n x. x \ ball z d \ S \ (f n has_field_derivative f' n x) (at x)" - by (auto simp: hfd) - obtain g g' where gg': "\x \ ball z d \ S. ((\n. f n x) sums g x) \ - ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" - by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) - then have "(\n. f' n z) sums g' z" - by (meson \0 < r\ centre_in_ball contra_subsetD r) - moreover have "(\n. f n z) sums (\n. f n z)" - using summable_sums centre_in_ball \0 < d\ \summable h\ le_h - by (metis (full_types) Int_iff gg' summable_def that) - moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" - proof (rule has_field_derivative_transform_within) - show "\x. dist x z < r \ g x = (\n. f n x)" - by (metis subsetD dist_commute gg' mem_ball r sums_unique) - qed (use \0 < r\ gg' \z \ S\ \0 < d\ in auto) - ultimately show ?thesis by auto - qed - then show ?thesis - by (rule_tac x="\x. suminf (\n. f n x)" in exI) meson -qed - - -text\Sometimes convenient to compare with a complex series of positive reals. (?)\ - -lemma series_and_derivative_comparison_complex: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" - and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" - shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" -apply (rule series_and_derivative_comparison_local [OF S hfd], assumption) -apply (rule ex_forward [OF to_g], assumption) -apply (erule exE) -apply (rule_tac x="Re \ h" in exI) -apply (force simp: summable_Re o_def nonneg_Reals_cmod_eq_Re image_subset_iff) -done - -text\Sometimes convenient to compare with a complex series of positive reals. (?)\ -lemma series_differentiable_comparison_complex: - fixes S :: "complex set" - assumes S: "open S" - and hfd: "\n x. x \ S \ f n field_differentiable (at x)" - and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ range h \ \\<^sub>\\<^sub>0 \ (\\<^sub>F n in sequentially. \y\ball x d \ S. cmod(f n y) \ cmod (h n))" - obtains g where "\x \ S. ((\n. f n x) sums g x) \ g field_differentiable (at x)" -proof - - have hfd': "\n x. x \ S \ (f n has_field_derivative deriv (f n) x) (at x)" - using hfd field_differentiable_derivI by blast - have "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. deriv (f n) x) sums g' x) \ (g has_field_derivative g' x) (at x)" - by (metis series_and_derivative_comparison_complex [OF S hfd' to_g]) - then show ?thesis - using field_differentiable_def that by blast -qed - -text\In particular, a power series is analytic inside circle of convergence.\ - -lemma power_series_and_derivative_0: - fixes a :: "nat \ complex" and r::real - assumes "summable (\n. a n * r^n)" - shows "\g g'. \z. cmod z < r \ - ((\n. a n * z^n) sums g z) \ ((\n. of_nat n * a n * z^(n - 1)) sums g' z) \ (g has_field_derivative g' z) (at z)" -proof (cases "0 < r") - case True - have der: "\n z. ((\x. a n * x ^ n) has_field_derivative of_nat n * a n * z ^ (n - 1)) (at z)" - by (rule derivative_eq_intros | simp)+ - have y_le: "\cmod (z - y) * 2 < r - cmod z\ \ cmod y \ cmod (of_real r + of_real (cmod z)) / 2" for z y - using \r > 0\ - apply (auto simp: algebra_simps norm_mult norm_divide norm_power simp flip: of_real_add) - using norm_triangle_ineq2 [of y z] - apply (simp only: diff_le_eq norm_minus_commute mult_2) - done - have "summable (\n. a n * complex_of_real r ^ n)" - using assms \r > 0\ by simp - moreover have "\z. cmod z < r \ cmod ((of_real r + of_real (cmod z)) / 2) < cmod (of_real r)" - using \r > 0\ - by (simp flip: of_real_add) - ultimately have sum: "\z. cmod z < r \ summable (\n. of_real (cmod (a n)) * ((of_real r + complex_of_real (cmod z)) / 2) ^ n)" - by (rule power_series_conv_imp_absconv_weak) - have "\g g'. \z \ ball 0 r. (\n. (a n) * z ^ n) sums g z \ - (\n. of_nat n * (a n) * z ^ (n - 1)) sums g' z \ (g has_field_derivative g' z) (at z)" - apply (rule series_and_derivative_comparison_complex [OF open_ball der]) - apply (rule_tac x="(r - norm z)/2" in exI) - apply (rule_tac x="\n. of_real(norm(a n)*((r + norm z)/2)^n)" in exI) - using \r > 0\ - apply (auto simp: sum eventually_sequentially norm_mult norm_power dist_norm intro!: mult_left_mono power_mono y_le) - done - then show ?thesis - by (simp add: ball_def) -next - case False then show ?thesis - apply (simp add: not_less) - using less_le_trans norm_not_less_zero by blast -qed - -proposition\<^marker>\tag unimportant\ power_series_and_derivative: - fixes a :: "nat \ complex" and r::real - assumes "summable (\n. a n * r^n)" - obtains g g' where "\z \ ball w r. - ((\n. a n * (z - w) ^ n) sums g z) \ ((\n. of_nat n * a n * (z - w) ^ (n - 1)) sums g' z) \ - (g has_field_derivative g' z) (at z)" - using power_series_and_derivative_0 [OF assms] - apply clarify - apply (rule_tac g="(\z. g(z - w))" in that) - using DERIV_shift [where z="-w"] - apply (auto simp: norm_minus_commute Ball_def dist_norm) - done - -proposition\<^marker>\tag unimportant\ power_series_holomorphic: - assumes "\w. w \ ball z r \ ((\n. a n*(w - z)^n) sums f w)" - shows "f holomorphic_on ball z r" -proof - - have "\f'. (f has_field_derivative f') (at w)" if w: "dist z w < r" for w - proof - - have inb: "z + complex_of_real ((dist z w + r) / 2) \ ball z r" - proof - - have wz: "cmod (w - z) < r" using w - by (auto simp: field_split_simps dist_norm norm_minus_commute) - then have "0 \ r" - by (meson less_eq_real_def norm_ge_zero order_trans) - show ?thesis - using w by (simp add: dist_norm \0\r\ flip: of_real_add) - qed - have sum: "summable (\n. a n * of_real (((cmod (z - w) + r) / 2) ^ n))" - using assms [OF inb] by (force simp: summable_def dist_norm) - obtain g g' where gg': "\u. u \ ball z ((cmod (z - w) + r) / 2) \ - (\n. a n * (u - z) ^ n) sums g u \ - (\n. of_nat n * a n * (u - z) ^ (n - 1)) sums g' u \ (g has_field_derivative g' u) (at u)" - by (rule power_series_and_derivative [OF sum, of z]) fastforce - have [simp]: "g u = f u" if "cmod (u - w) < (r - cmod (z - w)) / 2" for u - proof - - have less: "cmod (z - u) * 2 < cmod (z - w) + r" - using that dist_triangle2 [of z u w] - by (simp add: dist_norm [symmetric] algebra_simps) - show ?thesis - apply (rule sums_unique2 [of "\n. a n*(u - z)^n"]) - using gg' [of u] less w - apply (auto simp: assms dist_norm) - done - qed - have "(f has_field_derivative g' w) (at w)" - by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"]) - (use w gg' [of w] in \(force simp: dist_norm)+\) - then show ?thesis .. - qed - then show ?thesis by (simp add: holomorphic_on_open) -qed - -corollary holomorphic_iff_power_series: - "f holomorphic_on ball z r \ - (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" - apply (intro iffI ballI holomorphic_power_series, assumption+) - apply (force intro: power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) - done - -lemma power_series_analytic: - "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" - by (force simp: analytic_on_open intro!: power_series_holomorphic) - -lemma analytic_iff_power_series: - "f analytic_on ball z r \ - (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" - by (simp add: analytic_on_open holomorphic_iff_power_series) - -subsection\<^marker>\tag unimportant\ \Equality between holomorphic functions, on open ball then connected set\ - -lemma holomorphic_fun_eq_on_ball: - "\f holomorphic_on ball z r; g holomorphic_on ball z r; - w \ ball z r; - \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ - \ f w = g w" - apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) - apply (auto simp: holomorphic_iff_power_series) - done - -lemma holomorphic_fun_eq_0_on_ball: - "\f holomorphic_on ball z r; w \ ball z r; - \n. (deriv ^^ n) f z = 0\ - \ f w = 0" - apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) - apply (auto simp: holomorphic_iff_power_series) - done - -lemma holomorphic_fun_eq_0_on_connected: - assumes holf: "f holomorphic_on S" and "open S" - and cons: "connected S" - and der: "\n. (deriv ^^ n) f z = 0" - and "z \ S" "w \ S" - shows "f w = 0" -proof - - have *: "ball x e \ (\n. {w \ S. (deriv ^^ n) f w = 0})" - if "\u. (deriv ^^ u) f x = 0" "ball x e \ S" for x e - proof - - have "\x' n. dist x x' < e \ (deriv ^^ n) f x' = 0" - apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) - apply (rule holomorphic_on_subset [OF holf]) - using that apply simp_all - by (metis funpow_add o_apply) - with that show ?thesis by auto - qed - have 1: "openin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" - apply (rule open_subset, force) - using \open S\ - apply (simp add: open_contains_ball Ball_def) - apply (erule all_forward) - using "*" by auto blast+ - have 2: "closedin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" - using assms - by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) - obtain e where "e>0" and e: "ball w e \ S" using openE [OF \open S\ \w \ S\] . - then have holfb: "f holomorphic_on ball w e" - using holf holomorphic_on_subset by blast - have 3: "(\n. {w \ S. (deriv ^^ n) f w = 0}) = S \ f w = 0" - using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) - show ?thesis - using cons der \z \ S\ - apply (simp add: connected_clopen) - apply (drule_tac x="\n. {w \ S. (deriv ^^ n) f w = 0}" in spec) - apply (auto simp: 1 2 3) - done -qed - -lemma holomorphic_fun_eq_on_connected: - assumes "f holomorphic_on S" "g holomorphic_on S" and "open S" "connected S" - and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" - and "z \ S" "w \ S" - shows "f w = g w" -proof (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" S z, simplified]) - show "(\x. f x - g x) holomorphic_on S" - by (intro assms holomorphic_intros) - show "\n. (deriv ^^ n) (\x. f x - g x) z = 0" - using assms higher_deriv_diff by auto -qed (use assms in auto) - -lemma holomorphic_fun_eq_const_on_connected: - assumes holf: "f holomorphic_on S" and "open S" - and cons: "connected S" - and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" - and "z \ S" "w \ S" - shows "f w = f z" -proof (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" S z, simplified]) - show "(\w. f w - f z) holomorphic_on S" - by (intro assms holomorphic_intros) - show "\n. (deriv ^^ n) (\w. f w - f z) z = 0" - by (subst higher_deriv_diff) (use assms in \auto intro: holomorphic_intros\) -qed (use assms in auto) - -subsection\<^marker>\tag unimportant\ \Some basic lemmas about poles/singularities\ - -lemma pole_lemma: - assumes holf: "f holomorphic_on S" and a: "a \ interior S" - shows "(\z. if z = a then deriv f a - else (f z - f a) / (z - a)) holomorphic_on S" (is "?F holomorphic_on S") -proof - - have F1: "?F field_differentiable (at u within S)" if "u \ S" "u \ a" for u - proof - - have fcd: "f field_differentiable at u within S" - using holf holomorphic_on_def by (simp add: \u \ S\) - have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within S" - by (rule fcd derivative_intros | simp add: that)+ - have "0 < dist a u" using that dist_nz by blast - then show ?thesis - by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ S\) - qed - have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ S" for e - proof - - have holfb: "f holomorphic_on ball a e" - by (rule holomorphic_on_subset [OF holf \ball a e \ S\]) - have 2: "?F holomorphic_on ball a e - {a}" - apply (simp add: holomorphic_on_def flip: field_differentiable_def) - using mem_ball that - apply (auto intro: F1 field_differentiable_within_subset) - done - have "isCont (\z. if z = a then deriv f a else (f z - f a) / (z - a)) x" - if "dist a x < e" for x - proof (cases "x=a") - case True - then have "f field_differentiable at a" - using holfb \0 < e\ holomorphic_on_imp_differentiable_at by auto - with True show ?thesis - by (auto simp: continuous_at has_field_derivative_iff simp flip: DERIV_deriv_iff_field_differentiable - elim: rev_iffD1 [OF _ LIM_equal]) - next - case False with 2 that show ?thesis - by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) - qed - then have 1: "continuous_on (ball a e) ?F" - by (clarsimp simp: continuous_on_eq_continuous_at) - have "?F holomorphic_on ball a e" - by (auto intro: no_isolated_singularity [OF 1 2]) - with that show ?thesis - by (simp add: holomorphic_on_open field_differentiable_def [symmetric] - field_differentiable_at_within) - qed - show ?thesis - proof - fix x assume "x \ S" show "?F field_differentiable at x within S" - proof (cases "x=a") - case True then show ?thesis - using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) - next - case False with F1 \x \ S\ - show ?thesis by blast - qed - qed -qed - -lemma pole_theorem: - assumes holg: "g holomorphic_on S" and a: "a \ interior S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) holomorphic_on S" - using pole_lemma [OF holg a] - by (rule holomorphic_transform) (simp add: eq field_split_simps) - -lemma pole_lemma_open: - assumes "f holomorphic_on S" "open S" - shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on S" -proof (cases "a \ S") - case True with assms interior_eq pole_lemma - show ?thesis by fastforce -next - case False with assms show ?thesis - apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) - apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) - apply (rule derivative_intros | force)+ - done -qed - -lemma pole_theorem_open: - assumes holg: "g holomorphic_on S" and S: "open S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) holomorphic_on S" - using pole_lemma_open [OF holg S] - by (rule holomorphic_transform) (auto simp: eq divide_simps) - -lemma pole_theorem_0: - assumes holg: "g holomorphic_on S" and a: "a \ interior S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f holomorphic_on S" - using pole_theorem [OF holg a eq] - by (rule holomorphic_transform) (auto simp: eq field_split_simps) - -lemma pole_theorem_open_0: - assumes holg: "g holomorphic_on S" and S: "open S" - and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f holomorphic_on S" - using pole_theorem_open [OF holg S eq] - by (rule holomorphic_transform) (auto simp: eq field_split_simps) - -lemma pole_theorem_analytic: - assumes g: "g analytic_on S" - and eq: "\z. z \ S - \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" - shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" (is "?F analytic_on S") - unfolding analytic_on_def -proof - fix x - assume "x \ S" - with g obtain e where "0 < e" and e: "g holomorphic_on ball x e" - by (auto simp add: analytic_on_def) - obtain d where "0 < d" and d: "\w. w \ ball x d - {a} \ g w = (w - a) * f w" - using \x \ S\ eq by blast - have "?F holomorphic_on ball x (min d e)" - using d e \x \ S\ by (fastforce simp: holomorphic_on_subset subset_ball intro!: pole_theorem_open) - then show "\e>0. ?F holomorphic_on ball x e" - using \0 < d\ \0 < e\ not_le by fastforce -qed - -lemma pole_theorem_analytic_0: - assumes g: "g analytic_on S" - and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f analytic_on S" -proof - - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" - by auto - show ?thesis - using pole_theorem_analytic [OF g eq] by simp -qed - -lemma pole_theorem_analytic_open_superset: - assumes g: "g analytic_on S" and "S \ T" "open T" - and eq: "\z. z \ T - {a} \ g z = (z - a) * f z" - shows "(\z. if z = a then deriv g a - else f z - g a/(z - a)) analytic_on S" -proof (rule pole_theorem_analytic [OF g]) - fix z - assume "z \ S" - then obtain e where "0 < e" and e: "ball z e \ T" - using assms openE by blast - then show "\d>0. \w\ball z d - {a}. g w = (w - a) * f w" - using eq by auto -qed - -lemma pole_theorem_analytic_open_superset_0: - assumes g: "g analytic_on S" "S \ T" "open T" "\z. z \ T - {a} \ g z = (z - a) * f z" - and [simp]: "f a = deriv g a" "g a = 0" - shows "f analytic_on S" -proof - - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" - by auto - have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" - by (rule pole_theorem_analytic_open_superset [OF g]) - then show ?thesis by simp -qed - - -subsection\General, homology form of Cauchy's theorem\ - -text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ - -lemma contour_integral_continuous_on_linepath_2D: - assumes "open U" and cont_dw: "\w. w \ U \ F w contour_integrable_on (linepath a b)" - and cond_uu: "continuous_on (U \ U) (\(x,y). F x y)" - and abu: "closed_segment a b \ U" - shows "continuous_on U (\w. contour_integral (linepath a b) (F w))" -proof - - have *: "\d>0. \x'\U. dist x' w < d \ - dist (contour_integral (linepath a b) (F x')) - (contour_integral (linepath a b) (F w)) \ \" - if "w \ U" "0 < \" "a \ b" for w \ - proof - - obtain \ where "\>0" and \: "cball w \ \ U" using open_contains_cball \open U\ \w \ U\ by force - let ?TZ = "cball w \ \ closed_segment a b" - have "uniformly_continuous_on ?TZ (\(x,y). F x y)" - proof (rule compact_uniformly_continuous) - show "continuous_on ?TZ (\(x,y). F x y)" - by (rule continuous_on_subset[OF cond_uu]) (use SigmaE \ abu in blast) - show "compact ?TZ" - by (simp add: compact_Times) - qed - then obtain \ where "\>0" - and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ - dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" - apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) - using \0 < \\ \a \ b\ by auto - have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; - norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ - \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" - for x1 x2 x1' x2' - using \ [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm) - have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" - if "x' \ U" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' - proof - - have "(\x. F x' x - F w x) contour_integrable_on linepath a b" - by (simp add: \w \ U\ cont_dw contour_integrable_diff that) - then have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" - apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) - using \0 < \\ \0 < \\ that apply (auto simp: norm_minus_commute) - done - also have "\ = \" using \a \ b\ by simp - finally show ?thesis . - qed - show ?thesis - apply (rule_tac x="min \ \" in exI) - using \0 < \\ \0 < \\ - apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ U\ intro: le_ee) - done - qed - show ?thesis - proof (cases "a=b") - case True - then show ?thesis by simp - next - case False - show ?thesis - by (rule continuous_onI) (use False in \auto intro: *\) - qed -qed - -text\This version has \<^term>\polynomial_function \\ as an additional assumption.\ -lemma Cauchy_integral_formula_global_weak: - assumes "open U" and holf: "f holomorphic_on U" - and z: "z \ U" and \: "polynomial_function \" - and pasz: "path_image \ \ U - {z}" and loop: "pathfinish \ = pathstart \" - and zero: "\w. w \ U \ winding_number \ w = 0" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" - using has_vector_derivative_polynomial_function [OF \] by blast - then have "bounded(path_image \')" - by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) - then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" - using bounded_pos by force - define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w - define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" - have "path \" "valid_path \" using \ - by (auto simp: path_polynomial_function valid_path_polynomial_function) - then have ov: "open v" - by (simp add: v_def open_winding_number_levelsets loop) - have uv_Un: "U \ v = UNIV" - using pasz zero by (auto simp: v_def) - have conf: "continuous_on U f" - by (metis holf holomorphic_on_imp_continuous_on) - have hol_d: "(d y) holomorphic_on U" if "y \ U" for y - proof - - have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U" - by (simp add: holf pole_lemma_open \open U\) - then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" - using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \open U\ by fastforce - then have "continuous_on U (d y)" - apply (simp add: d_def continuous_on_eq_continuous_at \open U\, clarify) - using * holomorphic_on_def - by (meson field_differentiable_within_open field_differentiable_imp_continuous_at \open U\) - moreover have "d y holomorphic_on U - {y}" - proof - - have "\w. w \ U - {y} \ - (\w. if w = y then deriv f y else (f w - f y) / (w - y)) field_differentiable at w" - apply (rule_tac d="dist w y" and f = "\w. (f w - f y)/(w - y)" in field_differentiable_transform_within) - apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros) - using \open U\ holf holomorphic_on_imp_differentiable_at by blast - then show ?thesis - unfolding field_differentiable_def by (simp add: d_def holomorphic_on_open \open U\ open_delete) - qed - ultimately show ?thesis - by (rule no_isolated_singularity) (auto simp: \open U\) - qed - have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y - proof (rule contour_integrable_holomorphic_simple [where S = "U-{y}"]) - show "(\x. (f x - f y) / (x - y)) holomorphic_on U - {y}" - by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) - show "path_image \ \ U - {y}" - using pasz that by blast - qed (auto simp: \open U\ open_delete \valid_path \\) - define h where - "h z = (if z \ U then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z - have U: "((d z) has_contour_integral h z) \" if "z \ U" for z - proof - - have "d z holomorphic_on U" - by (simp add: hol_d that) - with that show ?thesis - apply (simp add: h_def) - by (meson Diff_subset \open U\ \valid_path \\ contour_integrable_holomorphic_simple has_contour_integral_integral pasz subset_trans) - qed - have V: "((\w. f w / (w - z)) has_contour_integral h z) \" if z: "z \ v" for z - proof - - have 0: "0 = (f z) * 2 * of_real (2 * pi) * \ * winding_number \ z" - using v_def z by auto - then have "((\x. 1 / (x - z)) has_contour_integral 0) \" - using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce - then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" - using has_contour_integral_lmul by fastforce - then have "((\x. f z / (x - z)) has_contour_integral 0) \" - by (simp add: field_split_simps) - moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" - using z - apply (auto simp: v_def) - apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) - done - ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" - by (rule has_contour_integral_add) - have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" - if "z \ U" - using * by (auto simp: divide_simps has_contour_integral_eq) - moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" - if "z \ U" - apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=U]]) - using U pasz \valid_path \\ that - apply (auto intro: holomorphic_on_imp_continuous_on hol_d) - apply (rule continuous_intros conf holomorphic_intros holf assms | force)+ - done - ultimately show ?thesis - using z by (simp add: h_def) - qed - have znot: "z \ path_image \" - using pasz by blast - obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - U \ d0 \ dist x y" - using separate_compact_closed [of "path_image \" "-U"] pasz \open U\ - by (fastforce simp add: \path \\ compact_path_image) - obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ U" - apply (rule that [of "d0/2"]) - using \0 < d0\ - apply (auto simp: dist_norm dest: d0) - done - have "\x x'. \x \ path_image \; dist x x' * 2 < dd\ \ \y k. x' = y + k \ y \ path_image \ \ dist 0 k * 2 \ dd" - apply (rule_tac x=x in exI) - apply (rule_tac x="x'-x" in exI) - apply (force simp: dist_norm) - done - then have 1: "path_image \ \ interior {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" - apply (clarsimp simp add: mem_interior) - using \0 < dd\ - apply (rule_tac x="dd/2" in exI, auto) - done - obtain T where "compact T" and subt: "path_image \ \ interior T" and T: "T \ U" - apply (rule that [OF _ 1]) - apply (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) - apply (rule order_trans [OF _ dd]) - using \0 < dd\ by fastforce - obtain L where "L>0" - and L: "\f B. \f holomorphic_on interior T; \z. z\interior T \ cmod (f z) \ B\ \ - cmod (contour_integral \ f) \ L * B" - using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] - by blast - have "bounded(f ` T)" - by (meson \compact T\ compact_continuous_image compact_imp_bounded conf continuous_on_subset T) - then obtain D where "D>0" and D: "\x. x \ T \ norm (f x) \ D" - by (auto simp: bounded_pos) - obtain C where "C>0" and C: "\x. x \ T \ norm x \ C" - using \compact T\ bounded_pos compact_imp_bounded by force - have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y - proof - - have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp - with le have ybig: "norm y > C" by force - with C have "y \ T" by force - then have ynot: "y \ path_image \" - using subt interior_subset by blast - have [simp]: "winding_number \ y = 0" - apply (rule winding_number_zero_outside [of _ "cball 0 C"]) - using ybig interior_subset subt - apply (force simp: loop \path \\ dist_norm intro!: C)+ - done - have [simp]: "h y = contour_integral \ (\w. f w/(w - y))" - by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) - have holint: "(\w. f w / (w - y)) holomorphic_on interior T" - apply (rule holomorphic_on_divide) - using holf holomorphic_on_subset interior_subset T apply blast - apply (rule holomorphic_intros)+ - using \y \ T\ interior_subset by auto - have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior T" for z - proof - - have "D * L / e + cmod z \ cmod y" - using le C [of z] z using interior_subset by force - then have DL2: "D * L / e \ cmod (z - y)" - using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) - have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" - by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) - also have "\ \ D * (e / L / D)" - apply (rule mult_mono) - using that D interior_subset apply blast - using \L>0\ \e>0\ \D>0\ DL2 - apply (auto simp: norm_divide field_split_simps) - done - finally show ?thesis . - qed - have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" - by (simp add: dist_norm) - also have "\ \ L * (D * (e / L / D))" - by (rule L [OF holint leD]) - also have "\ = e" - using \L>0\ \0 < D\ by auto - finally show ?thesis . - qed - then have "(h \ 0) at_infinity" - by (meson Lim_at_infinityI) - moreover have "h holomorphic_on UNIV" - proof - - have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" - if "x \ U" "z \ U" "x \ z" for x z - using that conf - apply (simp add: split_def continuous_on_eq_continuous_at \open U\) - apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ - done - have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" - by (rule continuous_intros)+ - have open_uu_Id: "open (U \ U - Id)" - apply (rule open_Diff) - apply (simp add: open_Times \open U\) - using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] - apply (auto simp: Id_fstsnd_eq algebra_simps) - done - have con_derf: "continuous (at z) (deriv f)" if "z \ U" for z - apply (rule continuous_on_interior [of U]) - apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open U\) - by (simp add: interior_open that \open U\) - have tendsto_f': "((\(x,y). if y = x then deriv f (x) - else (f (y) - f (x)) / (y - x)) \ deriv f x) - (at (x, x) within U \ U)" if "x \ U" for x - proof (rule Lim_withinI) - fix e::real assume "0 < e" - obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" - using \0 < e\ continuous_within_E [OF con_derf [OF \x \ U\]] - by (metis UNIV_I dist_norm) - obtain k2 where "k2>0" and k2: "ball x k2 \ U" - by (blast intro: openE [OF \open U\] \x \ U\) - have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" - if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" - for x' z' - proof - - have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w - apply (drule segment_furthest_le [where y=x]) - by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) - have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w - by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) - have f_has_der: "\x. x \ U \ (f has_field_derivative deriv f x) (at x within U)" - by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \open U\) - have "closed_segment x' z' \ U" - by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) - then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" - using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp - then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" - by (rule has_contour_integral_div) - have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" - apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) - using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] - \e > 0\ \z' \ x'\ - apply (auto simp: norm_divide divide_simps derf_le) - done - also have "\ \ e" using \0 < e\ by simp - finally show ?thesis . - qed - show "\d>0. \xa\U \ U. - 0 < dist xa (x, x) \ dist xa (x, x) < d \ - dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" - apply (rule_tac x="min k1 k2" in exI) - using \k1>0\ \k2>0\ \e>0\ - apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) - done - qed - have con_pa_f: "continuous_on (path_image \) f" - by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt T) - have le_B: "\T. T \ {0..1} \ cmod (vector_derivative \ (at T)) \ B" - apply (rule B) - using \' using path_image_def vector_derivative_at by fastforce - have f_has_cint: "\w. w \ v - path_image \ \ ((\u. f u / (u - w) ^ 1) has_contour_integral h w) \" - by (simp add: V) - have cond_uu: "continuous_on (U \ U) (\(x,y). d x y)" - apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') - apply (simp add: tendsto_within_open_NO_MATCH open_Times \open U\, clarify) - apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) - using con_ff - apply (auto simp: continuous_within) - done - have hol_dw: "(\z. d z w) holomorphic_on U" if "w \ U" for w - proof - - have "continuous_on U ((\(x,y). d x y) \ (\z. (w,z)))" - by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ - then have *: "continuous_on U (\z. if w = z then deriv f z else (f w - f z) / (w - z))" - by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) - have **: "\x. \x \ U; x \ w\ \ (\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" - apply (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) - apply (rule \open U\ derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+ - done - show ?thesis - unfolding d_def - apply (rule no_isolated_singularity [OF * _ \open U\, where K = "{w}"]) - apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \open U\ **) - done - qed - { fix a b - assume abu: "closed_segment a b \ U" - then have "\w. w \ U \ (\z. d z w) contour_integrable_on (linepath a b)" - by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) - then have cont_cint_d: "continuous_on U (\w. contour_integral (linepath a b) (\z. d z w))" - apply (rule contour_integral_continuous_on_linepath_2D [OF \open U\ _ _ abu]) - apply (auto intro: continuous_on_swap_args cond_uu) - done - have cont_cint_d\: "continuous_on {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" - proof (rule continuous_on_compose) - show "continuous_on {0..1} \" - using \path \\ path_def by blast - show "continuous_on (\ ` {0..1}) (\w. contour_integral (linepath a b) (\z. d z w))" - using pasz unfolding path_image_def - by (auto intro!: continuous_on_subset [OF cont_cint_d]) - qed - have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" - apply (simp add: contour_integrable_on) - apply (rule integrable_continuous_real) - apply (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) - using pf\' - by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) - have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\z. contour_integral \ (d z))" - using abu by (force simp: h_def intro: contour_integral_eq) - also have "\ = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" - apply (rule contour_integral_swap) - apply (rule continuous_on_subset [OF cond_uu]) - using abu pasz \valid_path \\ - apply (auto intro!: continuous_intros) - by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) - finally have cint_h_eq: - "contour_integral (linepath a b) h = - contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . - note cint_cint cint_h_eq - } note cint_h = this - have conthu: "continuous_on U h" - proof (simp add: continuous_on_sequentially, clarify) - fix a x - assume x: "x \ U" and au: "\n. a n \ U" and ax: "a \ x" - then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" - by (meson U contour_integrable_on_def eventuallyI) - obtain dd where "dd>0" and dd: "cball x dd \ U" using open_contains_cball \open U\ x by force - have A2: "uniform_limit (path_image \) (\n. d (a n)) (d x) sequentially" - unfolding uniform_limit_iff dist_norm - proof clarify - fix ee::real - assume "0 < ee" - show "\\<^sub>F n in sequentially. \\\path_image \. cmod (d (a n) \ - d x \) < ee" - proof - - let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" - have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" - apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) - using dd pasz \valid_path \\ - apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) - done - then obtain kk where "kk>0" - and kk: "\x x'. \x \ ?ddpa; x' \ ?ddpa; dist x' x < kk\ \ - dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" - by (rule uniformly_continuous_onE [where e = ee]) (use \0 < ee\ in auto) - have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" - for w z - using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm) - show ?thesis - using ax unfolding lim_sequentially eventually_sequentially - apply (drule_tac x="min dd kk" in spec) - using \dd > 0\ \kk > 0\ - apply (fastforce simp: kk dist_norm) - done - qed - qed - have "(\n. contour_integral \ (d (a n))) \ contour_integral \ (d x)" - by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \valid_path \\) - then have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" - by (simp add: h_def x) - then show "(h \ a) \ h x" - by (simp add: h_def x au o_def) - qed - show ?thesis - proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) - fix z0 - consider "z0 \ v" | "z0 \ U" using uv_Un by blast - then show "h field_differentiable at z0" - proof cases - assume "z0 \ v" then show ?thesis - using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ - by (auto simp: field_differentiable_def v_def) - next - assume "z0 \ U" then - obtain e where "e>0" and e: "ball z0 e \ U" by (blast intro: openE [OF \open U\]) - have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" - if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c - proof - - have *: "\x1 x2 z. z \ U \ closed_segment x1 x2 \ U \ (\w. d w z) contour_integrable_on linepath x1 x2" - using hol_dw holomorphic_on_imp_continuous_on \open U\ - by (auto intro!: contour_integrable_holomorphic_simple) - have abc: "closed_segment a b \ U" "closed_segment b c \ U" "closed_segment c a \ U" - using that e segments_subset_convex_hull by fastforce+ - have eq0: "\w. w \ U \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" - apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) - apply (rule holomorphic_on_subset [OF hol_dw]) - using e abc_subset by auto - have "contour_integral \ - (\x. contour_integral (linepath a b) (\z. d z x) + - (contour_integral (linepath b c) (\z. d z x) + - contour_integral (linepath c a) (\z. d z x))) = 0" - apply (rule contour_integral_eq_0) - using abc pasz U - apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ - done - then show ?thesis - by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) - qed - show ?thesis - using e \e > 0\ - by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic - Morera_triangle continuous_on_subset [OF conthu] *) - qed - qed - qed - ultimately have [simp]: "h z = 0" for z - by (meson Liouville_weak) - have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" - by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) - then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" - by (metis mult.commute has_contour_integral_lmul) - then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" - by (simp add: field_split_simps) - moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" - using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) - show ?thesis - using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) -qed - -theorem Cauchy_integral_formula_global: - assumes S: "open S" and holf: "f holomorphic_on S" - and z: "z \ S" and vpg: "valid_path \" - and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" - and zero: "\w. w \ S \ winding_number \ w = 0" - shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" -proof - - have "path \" using vpg by (blast intro: valid_path_imp_path) - have hols: "(\w. f w / (w - z)) holomorphic_on S - {z}" "(\w. 1 / (w - z)) holomorphic_on S - {z}" - by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ - then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" - by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete S vpg pasz) - obtain d where "d>0" - and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; - pathstart h = pathstart g \ pathfinish h = pathfinish g\ - \ path_image h \ S - {z} \ (\f. f holomorphic_on S - {z} \ contour_integral h f = contour_integral g f)" - using contour_integral_nearby_ends [OF _ \path \\ pasz] S by (simp add: open_Diff) metis - obtain p where polyp: "polynomial_function p" - and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" - using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast - then have ploop: "pathfinish p = pathstart p" using loop by auto - have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast - have [simp]: "z \ path_image \" using pasz by blast - have paps: "path_image p \ S - {z}" and cint_eq: "(\f. f holomorphic_on S - {z} \ contour_integral p f = contour_integral \ f)" - using pf ps led d [OF vpg vpp] \d > 0\ by auto - have wn_eq: "winding_number p z = winding_number \ z" - using vpp paps - by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) - have "winding_number p w = winding_number \ w" if "w \ S" for w - proof - - have hol: "(\v. 1 / (v - w)) holomorphic_on S - {z}" - using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) - have "w \ path_image p" "w \ path_image \" using paps pasz that by auto - then show ?thesis - using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) - qed - then have wn0: "\w. w \ S \ winding_number p w = 0" - by (simp add: zero) - show ?thesis - using Cauchy_integral_formula_global_weak [OF S holf z polyp paps ploop wn0] hols - by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) -qed - -theorem Cauchy_theorem_global: - assumes S: "open S" and holf: "f holomorphic_on S" - and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" - and pas: "path_image \ \ S" - and zero: "\w. w \ S \ winding_number \ w = 0" - shows "(f has_contour_integral 0) \" -proof - - obtain z where "z \ S" and znot: "z \ path_image \" - proof - - have "compact (path_image \)" - using compact_valid_path_image vpg by blast - then have "path_image \ \ S" - by (metis (no_types) compact_open path_image_nonempty S) - with pas show ?thesis by (blast intro: that) - qed - then have pasz: "path_image \ \ S - {z}" using pas by blast - have hol: "(\w. (w - z) * f w) holomorphic_on S" - by (rule holomorphic_intros holf)+ - show ?thesis - using Cauchy_integral_formula_global [OF S hol \z \ S\ vpg pasz loop zero] - by (auto simp: znot elim!: has_contour_integral_eq) -qed - -corollary Cauchy_theorem_global_outside: - assumes "open S" "f holomorphic_on S" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ S" - "\w. w \ S \ w \ outside(path_image \)" - shows "(f has_contour_integral 0) \" -by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) - -lemma simply_connected_imp_winding_number_zero: - assumes "simply_connected S" "path g" - "path_image g \ S" "pathfinish g = pathstart g" "z \ S" - shows "winding_number g z = 0" -proof - - have hom: "homotopic_loops S g (linepath (pathstart g) (pathstart g))" - by (meson assms homotopic_paths_imp_homotopic_loops pathfinish_linepath simply_connected_eq_contractible_path) - then have "homotopic_paths (- {z}) g (linepath (pathstart g) (pathstart g))" - by (meson \z \ S\ homotopic_loops_imp_homotopic_paths_null homotopic_paths_subset subset_Compl_singleton) - then have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" - by (rule winding_number_homotopic_paths) - also have "\ = 0" - using assms by (force intro: winding_number_trivial) - finally show ?thesis . -qed - -lemma Cauchy_theorem_simply_connected: - assumes "open S" "simply_connected S" "f holomorphic_on S" "valid_path g" - "path_image g \ S" "pathfinish g = pathstart g" - shows "(f has_contour_integral 0) g" -using assms -apply (simp add: simply_connected_eq_contractible_path) -apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"] - homotopic_paths_imp_homotopic_loops) -using valid_path_imp_path by blast - -proposition\<^marker>\tag unimportant\ holomorphic_logarithm_exists: - assumes A: "convex A" "open A" - and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" - and z0: "z0 \ A" - obtains g where "g holomorphic_on A" and "\x. x \ A \ exp (g x) = f x" -proof - - note f' = holomorphic_derivI [OF f(1) A(2)] - obtain g where g: "\x. x \ A \ (g has_field_derivative deriv f x / f x) (at x)" - proof (rule holomorphic_convex_primitive' [OF A]) - show "(\x. deriv f x / f x) holomorphic_on A" - by (intro holomorphic_intros f A) - qed (auto simp: A at_within_open[of _ A]) - define h where "h = (\x. -g z0 + ln (f z0) + g x)" - from g and A have g_holo: "g holomorphic_on A" - by (auto simp: holomorphic_on_def at_within_open[of _ A] field_differentiable_def) - hence h_holo: "h holomorphic_on A" - by (auto simp: h_def intro!: holomorphic_intros) - have "\c. \x\A. f x / exp (h x) - 1 = c" - proof (rule has_field_derivative_zero_constant, goal_cases) - case (2 x) - note [simp] = at_within_open[OF _ \open A\] - from 2 and z0 and f show ?case - by (auto simp: h_def exp_diff field_simps intro!: derivative_eq_intros g f') - qed fact+ - then obtain c where c: "\x. x \ A \ f x / exp (h x) - 1 = c" - by blast - from c[OF z0] and z0 and f have "c = 0" - by (simp add: h_def) - with c have "\x. x \ A \ exp (h x) = f x" by simp - from that[OF h_holo this] show ?thesis . -qed - -end diff --git a/src/HOL/Analysis/Change_Of_Vars.thy b/src/HOL/Analysis/Change_Of_Vars.thy --- a/src/HOL/Analysis/Change_Of_Vars.thy +++ b/src/HOL/Analysis/Change_Of_Vars.thy @@ -1,3474 +1,3474 @@ (* Title: HOL/Analysis/Change_Of_Vars.thy Authors: LC Paulson, based on material from HOL Light *) section\Change of Variables Theorems\ theory Change_Of_Vars imports Vitali_Covering_Theorem Determinants begin subsection \Measurable Shear and Stretch\ proposition fixes a :: "real^'n" assumes "m \ n" and ab_ne: "cbox a b \ {}" and an: "0 \ a$n" shows measurable_shear_interval: "(\x. \ i. if i = m then x$m + x$n else x$i) ` (cbox a b) \ lmeasurable" (is "?f ` _ \ _") and measure_shear_interval: "measure lebesgue ((\x. \ i. if i = m then x$m + x$n else x$i) ` cbox a b) = measure lebesgue (cbox a b)" (is "?Q") proof - have lin: "linear ?f" by (rule linearI) (auto simp: plus_vec_def scaleR_vec_def algebra_simps) show fab: "?f ` cbox a b \ lmeasurable" by (simp add: lin measurable_linear_image_interval) let ?c = "\ i. if i = m then b$m + b$n else b$i" let ?mn = "axis m 1 - axis n (1::real)" have eq1: "measure lebesgue (cbox a ?c) = measure lebesgue (?f ` cbox a b) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a$m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m})" proof (rule measure_Un3_negligible) show "cbox a ?c \ {x. ?mn \ x \ a$m} \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "negligible {x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ a $ m}) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible ((?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ a $ m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(?f ` cbox a b) \ (cbox a ?c \ {x. ?mn \ x \ b$m}) \ {x. ?mn \ x = b$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (?f ` cbox a b \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) have "negligible {x. ?mn \ x = b$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m})) \ {x. ?mn \ x = b$m}" using \m \ n\ ab_ne apply (auto simp: algebra_simps mem_box_cart inner_axis') apply (drule_tac x=m in spec)+ apply simp done ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (cbox a ?c \ {x. ?mn \ x \ b$m}))" by (rule negligible_subset) show "?f ` cbox a b \ cbox a ?c \ {x. ?mn \ x \ a $ m} \ cbox a ?c \ {x. ?mn \ x \ b$m} = cbox a ?c" (is "?lhs = _") proof show "?lhs \ cbox a ?c" by (auto simp: mem_box_cart add_mono) (meson add_increasing2 an order_trans) show "cbox a ?c \ ?lhs" apply (auto simp: algebra_simps image_iff inner_axis' lambda_add_Galois [OF \m \ n\]) apply (auto simp: mem_box_cart split: if_split_asm) done qed qed (fact fab) let ?d = "\ i. if i = m then a $ m - b $ m else 0" have eq2: "measure lebesgue (cbox a ?c \ {x. ?mn \ x \ a $ m}) + measure lebesgue (cbox a ?c \ {x. ?mn \ x \ b$m}) = measure lebesgue (cbox a (\ i. if i = m then a $ m + b $ n else b $ i))" proof (rule measure_translate_add[of "cbox a ?c \ {x. ?mn \ x \ a$m}" "cbox a ?c \ {x. ?mn \ x \ b$m}" "(\ i. if i = m then a$m - b$m else 0)" "cbox a (\ i. if i = m then a$m + b$n else b$i)"]) show "(cbox a ?c \ {x. ?mn \ x \ a$m}) \ lmeasurable" "cbox a ?c \ {x. ?mn \ x \ b$m} \ lmeasurable" by (auto simp: convex_Int convex_halfspace_le convex_halfspace_ge bounded_Int measurable_convex) have "\x. \x $ n + a $ m \ x $ m\ \ x \ (+) (\ i. if i = m then a $ m - b $ m else 0) ` {x. x $ n + b $ m \ x $ m}" using \m \ n\ by (rule_tac x="x - (\ i. if i = m then a$m - b$m else 0)" in image_eqI) (simp_all add: mem_box_cart) then have imeq: "(+) ?d ` {x. b $ m \ ?mn \ x} = {x. a $ m \ ?mn \ x}" using \m \ n\ by (auto simp: mem_box_cart inner_axis' algebra_simps) have "\x. \0 \ a $ n; x $ n + a $ m \ x $ m; \i. i \ m \ a $ i \ x $ i \ x $ i \ b $ i\ \ a $ m \ x $ m" using \m \ n\ by force then have "(+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {x. a $ m \ ?mn \ x}" using an ab_ne apply (simp add: cbox_translation [symmetric] translation_Int interval_ne_empty_cart imeq) apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib) by (metis (full_types) add_mono mult_2_right) then show "cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}) = cbox a (\ i. if i = m then a $ m + b $ n else b $ i)" (is "?lhs = ?rhs") using an \m \ n\ apply (auto simp: mem_box_cart inner_axis' algebra_simps if_distrib all_if_distrib, force) apply (drule_tac x=n in spec)+ by (meson ab_ne add_mono_thms_linordered_semiring(3) dual_order.trans interval_ne_empty_cart(1)) have "negligible{x. ?mn \ x = a$m}" by (metis \m \ n\ axis_index_axis eq_iff_diff_eq_0 negligible_hyperplane) moreover have "(cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x})) \ {x. ?mn \ x = a$m}" using \m \ n\ antisym_conv by (fastforce simp: algebra_simps mem_box_cart inner_axis') ultimately show "negligible (cbox a ?c \ {x. ?mn \ x \ a $ m} \ (+) ?d ` (cbox a ?c \ {x. b $ m \ ?mn \ x}))" by (rule negligible_subset) qed have ac_ne: "cbox a ?c \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have ax_ne: "cbox a (\ i. if i = m then a $ m + b $ n else b $ i) \ {}" using ab_ne an by (clarsimp simp: interval_eq_empty_cart) (meson add_less_same_cancel1 le_less_linear less_le_trans) have eq3: "measure lebesgue (cbox a ?c) = measure lebesgue (cbox a (\ i. if i = m then a$m + b$n else b$i)) + measure lebesgue (cbox a b)" by (simp add: content_cbox_if_cart ab_ne ac_ne ax_ne algebra_simps prod.delta_remove if_distrib [of "\u. u - z" for z] prod.remove) show ?Q using eq1 eq2 eq3 by (simp add: algebra_simps) qed proposition fixes S :: "(real^'n) set" assumes "S \ lmeasurable" shows measurable_stretch: "((\x. \ k. m k * x$k) ` S) \ lmeasurable" (is "?f ` S \ _") and measure_stretch: "measure lebesgue ((\x. \ k. m k * x$k) ` S) = \prod m UNIV\ * measure lebesgue S" (is "?MEQ") proof - have "(?f ` S) \ lmeasurable \ ?MEQ" proof (cases "\k. m k \ 0") case True have m0: "0 < \prod m UNIV\" using True by simp have "(indicat_real (?f ` S) has_integral \prod m UNIV\ * measure lebesgue S) UNIV" proof (clarsimp simp add: has_integral_alt [where i=UNIV]) fix e :: "real" assume "e > 0" have "(indicat_real S has_integral (measure lebesgue S)) UNIV" using assms lmeasurable_iff_has_integral by blast then obtain B where "B>0" and B: "\a b. ball 0 B \ cbox a b \ \z. (indicat_real S has_integral z) (cbox a b) \ \z - measure lebesgue S\ < e / \prod m UNIV\" by (simp add: has_integral_alt [where i=UNIV]) (metis (full_types) divide_pos_pos m0 m0 \e > 0\) show "\B>0. \a b. ball 0 B \ cbox a b \ (\z. (indicat_real (?f ` S) has_integral z) (cbox a b) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" proof (intro exI conjI allI) let ?C = "Max (range (\k. \m k\)) * B" show "?C > 0" using True \B > 0\ by (simp add: Max_gr_iff) show "ball 0 ?C \ cbox u v \ (\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e)" for u v proof assume uv: "ball 0 ?C \ cbox u v" with \?C > 0\ have cbox_ne: "cbox u v \ {}" using centre_in_ball by blast let ?\ = "\k. u$k / m k" let ?\ = "\k. v$k / m k" have invm0: "\k. inverse (m k) \ 0" using True by auto have "ball 0 B \ (\x. \ k. x $ k / m k) ` ball 0 ?C" proof clarsimp fix x :: "real^'n" assume x: "norm x < B" have [simp]: "\Max (range (\k. \m k\))\ = Max (range (\k. \m k\))" by (meson Max_ge abs_ge_zero abs_of_nonneg finite finite_imageI order_trans rangeI) have "norm (\ k. m k * x $ k) \ norm (Max (range (\k. \m k\)) *\<^sub>R x)" by (rule norm_le_componentwise_cart) (auto simp: abs_mult intro: mult_right_mono) also have "\ < ?C" using x by simp (metis \B > 0\ \?C > 0\ mult.commute real_mult_less_iff1 zero_less_mult_pos) finally have "norm (\ k. m k * x $ k) < ?C" . then show "x \ (\x. \ k. x $ k / m k) ` ball 0 ?C" using stretch_Galois [of "inverse \ m"] True by (auto simp: image_iff field_simps) qed then have Bsub: "ball 0 B \ cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))" using cbox_ne uv image_stretch_interval_cart [of "inverse \ m" u v, symmetric] by (force simp: field_simps) obtain z where zint: "(indicat_real S has_integral z) (cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" and zless: "\z - measure lebesgue S\ < e / \prod m UNIV\" using B [OF Bsub] by blast have ind: "indicat_real (?f ` S) = (\x. indicator S (\ k. x$k / m k))" using True stretch_Galois [of m] by (force simp: indicator_def) show "\z. (indicat_real (?f ` S) has_integral z) (cbox u v) \ \z - \prod m UNIV\ * measure lebesgue S\ < e" proof (simp add: ind, intro conjI exI) have "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) ((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k)))" using True has_integral_stretch_cart [OF zint, of "inverse \ m"] by (simp add: field_simps prod_dividef) moreover have "((\x. \ k. x $ k * m k) ` cbox (\ k. min (?\ k) (?\ k)) (\ k. max (?\ k) (?\ k))) = cbox u v" using True image_stretch_interval_cart [of "inverse \ m" u v, symmetric] image_stretch_interval_cart [of "\k. 1" u v, symmetric] \cbox u v \ {}\ by (simp add: field_simps image_comp o_def) ultimately show "((\x. indicat_real S (\ k. x $ k/ m k)) has_integral z *\<^sub>R \prod m UNIV\) (cbox u v)" by simp have "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ = \prod m UNIV\ * \z - measure lebesgue S\" by (metis (no_types, hide_lams) abs_abs abs_scaleR mult.commute real_scaleR_def right_diff_distrib') also have "\ < e" using zless True by (simp add: field_simps) finally show "\z *\<^sub>R \prod m UNIV\ - \prod m UNIV\ * measure lebesgue S\ < e" . qed qed qed qed then show ?thesis by (auto simp: has_integral_integrable integral_unique lmeasure_integral_UNIV measurable_integrable) next case False then obtain k where "m k = 0" and prm: "prod m UNIV = 0" by auto have nfS: "negligible (?f ` S)" by (rule negligible_subset [OF negligible_standard_hyperplane_cart]) (use \m k = 0\ in auto) then have "(?f ` S) \ lmeasurable" by (simp add: negligible_iff_measure) with nfS show ?thesis by (simp add: prm negligible_iff_measure0) qed then show "(?f ` S) \ lmeasurable" ?MEQ by metis+ qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "linear f" "S \ lmeasurable" shows measurable_linear_image: "(f ` S) \ lmeasurable" and measure_linear_image: "measure lebesgue (f ` S) = \det (matrix f)\ * measure lebesgue S" (is "?Q f S") proof - have "\S \ lmeasurable. (f ` S) \ lmeasurable \ ?Q f S" proof (rule induct_linear_elementary [OF \linear f\]; intro ballI) fix f g and S :: "(real,'n) vec set" assume "linear f" and "linear g" and f [rule_format]: "\S \ lmeasurable. f ` S \ lmeasurable \ ?Q f S" and g [rule_format]: "\S \ lmeasurable. g ` S \ lmeasurable \ ?Q g S" and S: "S \ lmeasurable" then have gS: "g ` S \ lmeasurable" by blast show "(f \ g) ` S \ lmeasurable \ ?Q (f \ g) S" using f [OF gS] g [OF S] matrix_compose [OF \linear g\ \linear f\] by (simp add: o_def image_comp abs_mult det_mul) next fix f :: "real^'n::_ \ real^'n::_" and i and S :: "(real^'n::_) set" assume "linear f" and 0: "\x. f x $ i = 0" and "S \ lmeasurable" then have "\ inj f" by (metis (full_types) linear_injective_imp_surjective one_neq_zero surjE vec_component) have detf: "det (matrix f) = 0" using \\ inj f\ det_nz_iff_inj[OF \linear f\] by blast show "f ` S \ lmeasurable \ ?Q f S" proof show "f ` S \ lmeasurable" using lmeasurable_iff_indicator_has_integral \linear f\ \\ inj f\ negligible_UNIV negligible_linear_singular_image by blast have "measure lebesgue (f ` S) = 0" by (meson \\ inj f\ \linear f\ negligible_imp_measure0 negligible_linear_singular_image) also have "\ = \det (matrix f)\ * measure lebesgue S" by (simp add: detf) finally show "?Q f S" . qed next fix c and S :: "(real^'n::_) set" assume "S \ lmeasurable" show "(\a. \ i. c i * a $ i) ` S \ lmeasurable \ ?Q (\a. \ i. c i * a $ i) S" proof show "(\a. \ i. c i * a $ i) ` S \ lmeasurable" by (simp add: \S \ lmeasurable\ measurable_stretch) show "?Q (\a. \ i. c i * a $ i) S" by (simp add: measure_stretch [OF \S \ lmeasurable\, of c] axis_def matrix_def det_diagonal) qed next fix m :: "'n" and n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i" have lin: "linear ?h" by (rule linearI) (simp_all add: plus_vec_def scaleR_vec_def) have meq: "measure lebesgue ((\v::(real, 'n) vec. \ i. v $ Fun.swap m n id i) ` cbox a b) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have him: "?h ` (cbox a b) \ {}" by blast have eq: "?h ` (cbox a b) = cbox (?h a) (?h b)" by (auto simp: image_iff lambda_swap_Galois mem_box_cart) (metis swap_id_eq)+ show ?thesis using him prod.permute [OF permutes_swap_id, where S=UNIV and g="\i. (b - a)$i", symmetric] by (simp add: eq content_cbox_cart False) qed have "(\ i j. if Fun.swap m n id i = j then 1 else 0) = (\ i j. if j = Fun.swap m n id i then 1 else (0::real))" by (auto intro!: Cart_lambda_cong) then have "matrix ?h = transpose(\ i j. mat 1 $ i $ Fun.swap m n id j)" by (auto simp: matrix_eq transpose_def axis_def mat_def matrix_def) then have 1: "\det (matrix ?h)\ = 1" by (simp add: det_permute_columns permutes_swap_id sign_swap_id abs_mult) show "?h ` S \ lmeasurable \ ?Q ?h S" proof show "?h ` S \ lmeasurable" "?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force+ qed next fix m n :: "'n" and S :: "(real, 'n) vec set" assume "m \ n" and "S \ lmeasurable" let ?h = "\v::(real, 'n) vec. \ i. if i = m then v $ m + v $ n else v $ i" have lin: "linear ?h" by (rule linearI) (auto simp: algebra_simps plus_vec_def scaleR_vec_def vec_eq_iff) consider "m < n" | " n < m" using \m \ n\ less_linear by blast then have 1: "det(matrix ?h) = 1" proof cases assume "m < n" have *: "matrix ?h $ i $ j = (0::real)" if "j < i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j < i\ axis_def \m < n\ by auto with \m < n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_upperdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) next assume "n < m" have *: "matrix ?h $ i $ j = (0::real)" if "j > i" for i j :: 'n proof - have "axis j 1 = (\ n. if n = j then 1 else (0::real))" using axis_def by blast then have "(\ p q. if p = m then axis q 1 $ m + axis q 1 $ n else axis q 1 $ p) $ i $ j = (0::real)" using \j > i\ axis_def \m > n\ by auto with \m > n\ show ?thesis by (auto simp: matrix_def axis_def cong: if_cong) qed show ?thesis using \m \ n\ by (subst det_lowerdiagonal [OF *]) (auto simp: matrix_def axis_def cong: if_cong) qed have meq: "measure lebesgue (?h ` (cbox a b)) = measure lebesgue (cbox a b)" for a b proof (cases "cbox a b = {}") case True then show ?thesis by simp next case False then have ne: "(+) (\ i. if i = n then - a $ n else 0) ` cbox a b \ {}" by auto let ?v = "\ i. if i = n then - a $ n else 0" have "?h ` cbox a b = (+) (\ i. if i = m \ i = n then a $ n else 0) ` ?h ` (+) ?v ` (cbox a b)" using \m \ n\ unfolding image_comp o_def by (force simp: vec_eq_iff) then have "measure lebesgue (?h ` (cbox a b)) = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` (+) ?v ` cbox a b)" by (rule ssubst) (rule measure_translation) also have "\ = measure lebesgue ((\v. \ i. if i = m then v $ m + v $ n else v $ i) ` cbox (?v +a) (?v + b))" by (metis (no_types, lifting) cbox_translation) also have "\ = measure lebesgue ((+) (\ i. if i = n then - a $ n else 0) ` cbox a b)" apply (subst measure_shear_interval) using \m \ n\ ne apply auto apply (simp add: cbox_translation) by (metis cbox_borel cbox_translation measure_completion sets_lborel) also have "\ = measure lebesgue (cbox a b)" by (rule measure_translation) finally show ?thesis . qed show "?h ` S \ lmeasurable \ ?Q ?h S" using measure_linear_sufficient [OF lin \S \ lmeasurable\] meq 1 by force qed with assms show "(f ` S) \ lmeasurable" "?Q f S" by metis+ qed lemma fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes f: "orthogonal_transformation f" and S: "S \ lmeasurable" shows measurable_orthogonal_image: "f ` S \ lmeasurable" and measure_orthogonal_image: "measure lebesgue (f ` S) = measure lebesgue S" proof - have "linear f" by (simp add: f orthogonal_transformation_linear) then show "f ` S \ lmeasurable" by (metis S measurable_linear_image) show "measure lebesgue (f ` S) = measure lebesgue S" by (simp add: measure_linear_image \linear f\ S f) qed proposition measure_semicontinuous_with_hausdist_explicit: assumes "bounded S" and neg: "negligible(frontier S)" and "e > 0" obtains d where "d > 0" "\T. \T \ lmeasurable; \y. y \ T \ \x. x \ S \ dist x y < d\ \ measure lebesgue T < measure lebesgue S + e" proof (cases "S = {}") case True with that \e > 0\ show ?thesis by force next case False then have frS: "frontier S \ {}" using \bounded S\ frontier_eq_empty not_bounded_UNIV by blast have "S \ lmeasurable" by (simp add: \bounded S\ measurable_Jordan neg) have null: "(frontier S) \ null_sets lebesgue" by (metis neg negligible_iff_null_sets) have "frontier S \ lmeasurable" and mS0: "measure lebesgue (frontier S) = 0" using neg negligible_imp_measurable negligible_iff_measure by blast+ with \e > 0\ sets_lebesgue_outer_open obtain U where "open U" and U: "frontier S \ U" "U - frontier S \ lmeasurable" "emeasure lebesgue (U - frontier S) < e" by (metis fmeasurableD) with null have "U \ lmeasurable" by (metis borel_open measurable_Diff_null_set sets_completionI_sets sets_lborel) have "measure lebesgue (U - frontier S) = measure lebesgue U" using mS0 by (simp add: \U \ lmeasurable\ fmeasurableD measure_Diff_null_set null) with U have mU: "measure lebesgue U < e" by (simp add: emeasure_eq_measure2 ennreal_less_iff) show ?thesis proof have "U \ UNIV" using \U \ lmeasurable\ by auto then have "- U \ {}" by blast with \open U\ \frontier S \ U\ show "setdist (frontier S) (- U) > 0" by (auto simp: \bounded S\ open_closed compact_frontier_bounded setdist_gt_0_compact_closed frS) fix T assume "T \ lmeasurable" and T: "\t. t \ T \ \y. y \ S \ dist y t < setdist (frontier S) (- U)" then have "measure lebesgue T - measure lebesgue S \ measure lebesgue (T - S)" by (simp add: \S \ lmeasurable\ measure_diff_le_measure_setdiff) also have "\ \ measure lebesgue U" proof - have "T - S \ U" proof clarify fix x assume "x \ T" and "x \ S" then obtain y where "y \ S" and y: "dist y x < setdist (frontier S) (- U)" using T by blast have "closed_segment x y \ frontier S \ {}" using connected_Int_frontier \x \ S\ \y \ S\ by blast then obtain z where z: "z \ closed_segment x y" "z \ frontier S" by auto with y have "dist z x < setdist(frontier S) (- U)" by (auto simp: dist_commute dest!: dist_in_closed_segment) with z have False if "x \ -U" using setdist_le_dist [OF \z \ frontier S\ that] by auto then show "x \ U" by blast qed then show ?thesis by (simp add: \S \ lmeasurable\ \T \ lmeasurable\ \U \ lmeasurable\ fmeasurableD measure_mono_fmeasurable sets.Diff) qed finally have "measure lebesgue T - measure lebesgue S \ measure lebesgue U" . with mU show "measure lebesgue T < measure lebesgue S + e" by linarith qed qed proposition fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" and bounded: "\x. x \ S \ \det (matrix (f' x))\ \ B" shows measurable_bounded_differentiable_image: "f ` S \ lmeasurable" and measure_bounded_differentiable_image: "measure lebesgue (f ` S) \ B * measure lebesgue S" (is "?M") proof - have "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ B * measure lebesgue S" proof (cases "B < 0") case True then have "S = {}" by (meson abs_ge_zero bounded empty_iff equalityI less_le_trans linorder_not_less subsetI) then show ?thesis by auto next case False then have "B \ 0" by arith let ?\ = "measure lebesgue" have f_diff: "f differentiable_on S" using deriv by (auto simp: differentiable_on_def differentiable_def) have eps: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * ?\ S" (is "?ME") if "e > 0" for e proof - have eps_d: "f ` S \ lmeasurable" "?\ (f ` S) \ (B+e) * (?\ S + d)" (is "?MD") if "d > 0" for d proof - obtain T where T: "open T" "S \ T" and TS: "(T-S) \ lmeasurable" and "emeasure lebesgue (T-S) < ennreal d" using S \d > 0\ sets_lebesgue_outer_open by blast then have "?\ (T-S) < d" by (metis emeasure_eq_measure2 ennreal_leI not_less) with S T TS have "T \ lmeasurable" and Tless: "?\ T < ?\ S + d" by (auto simp: measurable_measure_Diff dest!: fmeasurable_Diff_D) have "\r. 0 < r \ r < d \ ball x r \ T \ f ` (S \ ball x r) \ lmeasurable \ ?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" if "x \ S" "d > 0" for x d proof - have lin: "linear (f' x)" and lim0: "((\y. (f y - (f x + f' x (y - x))) /\<^sub>R norm(y - x)) \ 0) (at x within S)" using deriv \x \ S\ by (auto simp: has_derivative_within bounded_linear.linear field_simps) have bo: "bounded (f' x ` ball 0 1)" by (simp add: bounded_linear_image linear_linear lin) have neg: "negligible (frontier (f' x ` ball 0 1))" using deriv has_derivative_linear \x \ S\ by (auto intro!: negligible_convex_frontier [OF convex_linear_image]) have 0: "0 < e * unit_ball_vol (real CARD('n))" using \e > 0\ by simp obtain k where "k > 0" and k: "\U. \U \ lmeasurable; \y. y \ U \ \z. z \ f' x ` ball 0 1 \ dist z y < k\ \ ?\ U < ?\ (f' x ` ball 0 1) + e * unit_ball_vol (CARD('n))" using measure_semicontinuous_with_hausdist_explicit [OF bo neg 0] by blast obtain l where "l > 0" and l: "ball x l \ T" using \x \ S\ \open T\ \S \ T\ openE by blast obtain \ where "0 < \" and \: "\y. \y \ S; y \ x; dist y x < \\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x) < k" using lim0 \k > 0\ by (simp add: Lim_within) (auto simp add: field_simps) define r where "r \ min (min l (\/2)) (min 1 (d/2))" show ?thesis proof (intro exI conjI) show "r > 0" "r < d" using \l > 0\ \\ > 0\ \d > 0\ by (auto simp: r_def) have "r \ l" by (auto simp: r_def) with l show "ball x r \ T" by auto have ex_lessK: "\x' \ ball 0 1. dist (f' x x') ((f y - f x) /\<^sub>R r) < k" if "y \ S" and "dist x y < r" for y proof (cases "y = x") case True with lin linear_0 \k > 0\ that show ?thesis by (rule_tac x=0 in bexI) (auto simp: linear_0) next case False then show ?thesis proof (rule_tac x="(y - x) /\<^sub>R r" in bexI) have "f' x ((y - x) /\<^sub>R r) = f' x (y - x) /\<^sub>R r" by (simp add: lin linear_scale) then have "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) = norm (f' x (y - x) /\<^sub>R r - (f y - f x) /\<^sub>R r)" by (simp add: dist_norm) also have "\ = norm (f' x (y - x) - (f y - f x)) / r" using \r > 0\ by (simp add: divide_simps scale_right_diff_distrib [symmetric]) also have "\ \ norm (f y - (f x + f' x (y - x))) / norm (y - x)" using that \r > 0\ False by (simp add: algebra_simps field_split_simps dist_norm norm_minus_commute mult_right_mono) also have "\ < k" using that \0 < \\ by (simp add: dist_commute r_def \ [OF \y \ S\ False]) finally show "dist (f' x ((y - x) /\<^sub>R r)) ((f y - f x) /\<^sub>R r) < k" . show "(y - x) /\<^sub>R r \ ball 0 1" using that \r > 0\ by (simp add: dist_norm divide_simps norm_minus_commute) qed qed let ?rfs = "(\x. x /\<^sub>R r) ` (+) (- f x) ` f ` (S \ ball x r)" have rfs_mble: "?rfs \ lmeasurable" proof (rule bounded_set_imp_lmeasurable) have "f differentiable_on S \ ball x r" using f_diff by (auto simp: fmeasurableD differentiable_on_subset) with S show "?rfs \ sets lebesgue" by (auto simp: sets.Int intro!: lebesgue_sets_translation differentiable_image_in_sets_lebesgue) let ?B = "(\(x, y). x + y) ` (f' x ` ball 0 1 \ ball 0 k)" have "bounded ?B" by (simp add: bounded_plus [OF bo]) moreover have "?rfs \ ?B" apply (auto simp: dist_norm image_iff dest!: ex_lessK) by (metis (no_types, hide_lams) add.commute diff_add_cancel dist_0_norm dist_commute dist_norm mem_ball) ultimately show "bounded (?rfs)" by (rule bounded_subset) qed then have "(\x. r *\<^sub>R x) ` ?rfs \ lmeasurable" by (simp add: measurable_linear_image) with \r > 0\ have "(+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) then have "(+) (f x) ` (+) (- f x) ` f ` (S \ ball x r) \ lmeasurable" using measurable_translation by blast then show fsb: "f ` (S \ ball x r) \ lmeasurable" by (simp add: image_comp o_def) have "?\ (f ` (S \ ball x r)) = ?\ (?rfs) * r ^ CARD('n)" using \r > 0\ fsb by (simp add: measure_linear_image measure_translation_subtract measurable_translation_subtract field_simps cong: image_cong_simp) also have "\ \ (\det (matrix (f' x))\ * unit_ball_vol (CARD('n)) + e * unit_ball_vol (CARD('n))) * r ^ CARD('n)" proof - have "?\ (?rfs) < ?\ (f' x ` ball 0 1) + e * unit_ball_vol (CARD('n))" using rfs_mble by (force intro: k dest!: ex_lessK) then have "?\ (?rfs) < \det (matrix (f' x))\ * unit_ball_vol (CARD('n)) + e * unit_ball_vol (CARD('n))" by (simp add: lin measure_linear_image [of "f' x"] content_ball) with \r > 0\ show ?thesis by auto qed also have "\ \ (B + e) * ?\ (ball x r)" using bounded [OF \x \ S\] \r > 0\ by (simp add: content_ball algebra_simps) finally show "?\ (f ` (S \ ball x r)) \ (B + e) * ?\ (ball x r)" . qed qed then obtain r where r0d: "\x d. \x \ S; d > 0\ \ 0 < r x d \ r x d < d" and rT: "\x d. \x \ S; d > 0\ \ ball x (r x d) \ T" and r: "\x d. \x \ S; d > 0\ \ (f ` (S \ ball x (r x d))) \ lmeasurable \ ?\ (f ` (S \ ball x (r x d))) \ (B + e) * ?\ (ball x (r x d))" by metis obtain C where "countable C" and Csub: "C \ {(x,r x t) |x t. x \ S \ 0 < t}" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible(S - (\i \ C. ball (fst i) (snd i)))" apply (rule Vitali_covering_theorem_balls [of S "{(x,r x t) |x t. x \ S \ 0 < t}" fst snd]) apply auto by (metis dist_eq_0_iff r0d) let ?UB = "(\(x,s) \ C. ball x s)" have eq: "f ` (S \ ?UB) = (\(x,s) \ C. f ` (S \ ball x s))" by auto have mle: "?\ (\(x,s) \ K. f ` (S \ ball x s)) \ (B + e) * (?\ S + d)" (is "?l \ ?r") if "K \ C" and "finite K" for K proof - have gt0: "b > 0" if "(a, b) \ K" for a b using Csub that \K \ C\ r0d by auto have inj: "inj_on (\(x, y). ball x y) K" by (force simp: inj_on_def ball_eq_ball_iff dest: gt0) have disjnt: "disjoint ((\(x, y). ball x y) ` K)" using pwC that apply (clarsimp simp: pairwise_def case_prod_unfold ball_eq_ball_iff) by (metis subsetD fst_conv snd_conv) have "?l \ (\i\K. ?\ (case i of (x, s) \ f ` (S \ ball x s)))" proof (rule measure_UNION_le [OF \finite K\], clarify) fix x r assume "(x,r) \ K" then have "x \ S" using Csub \K \ C\ by auto show "f ` (S \ ball x r) \ sets lebesgue" by (meson Int_lower1 S differentiable_on_subset f_diff fmeasurableD lmeasurable_ball order_refl sets.Int differentiable_image_in_sets_lebesgue) qed also have "\ \ (\(x,s) \ K. (B + e) * ?\ (ball x s))" apply (rule sum_mono) using Csub r \K \ C\ by auto also have "\ = (B + e) * (\(x,s) \ K. ?\ (ball x s))" by (simp add: prod.case_distrib sum_distrib_left) also have "\ = (B + e) * sum ?\ ((\(x, y). ball x y) ` K)" using \B \ 0\ \e > 0\ by (simp add: inj sum.reindex prod.case_distrib) also have "\ = (B + e) * ?\ (\(x,s) \ K. ball x s)" using \B \ 0\ \e > 0\ that by (subst measure_Union') (auto simp: disjnt measure_Union') also have "\ \ (B + e) * ?\ T" using \B \ 0\ \e > 0\ that apply simp apply (rule measure_mono_fmeasurable [OF _ _ \T \ lmeasurable\]) using Csub rT by force+ also have "\ \ (B + e) * (?\ S + d)" using \B \ 0\ \e > 0\ Tless by simp finally show ?thesis . qed have fSUB_mble: "(f ` (S \ ?UB)) \ lmeasurable" unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: fmeasurable_UN_bound [OF \countable C\ _ mle]) have fSUB_meas: "?\ (f ` (S \ ?UB)) \ (B + e) * (?\ S + d)" (is "?MUB") unfolding eq using Csub r False \e > 0\ that by (auto simp: intro!: measure_UN_bound [OF \countable C\ _ mle]) have neg: "negligible ((f ` (S \ ?UB) - f ` S) \ (f ` S - f ` (S \ ?UB)))" proof (rule negligible_subset [OF negligible_differentiable_image_negligible [OF order_refl negC, where f=f]]) show "f differentiable_on S - (\i\C. ball (fst i) (snd i))" by (meson DiffE differentiable_on_subset subsetI f_diff) qed force show "f ` S \ lmeasurable" by (rule lmeasurable_negligible_symdiff [OF fSUB_mble neg]) show ?MD using fSUB_meas measure_negligible_symdiff [OF fSUB_mble neg] by simp qed show "f ` S \ lmeasurable" using eps_d [of 1] by simp show ?ME proof (rule field_le_epsilon) fix \ :: real assume "0 < \" then show "?\ (f ` S) \ (B + e) * ?\ S + \" using eps_d [of "\ / (B+e)"] \e > 0\ \B \ 0\ by (auto simp: divide_simps mult_ac) qed qed show ?thesis proof (cases "?\ S = 0") case True with eps have "?\ (f ` S) = 0" by (metis mult_zero_right not_le zero_less_measure_iff) then show ?thesis using eps [of 1] by (simp add: True) next case False have "?\ (f ` S) \ B * ?\ S" proof (rule field_le_epsilon) fix e :: real assume "e > 0" then show "?\ (f ` S) \ B * ?\ S + e" using eps [of "e / ?\ S"] False by (auto simp: algebra_simps zero_less_measure_iff) qed with eps [of 1] show ?thesis by auto qed qed then show "f ` S \ lmeasurable" ?M by blast+ qed lemma m_diff_image_weak: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ lmeasurable" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" proof - let ?\ = "measure lebesgue" have aint_S: "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int unfolding absolutely_integrable_on_def by auto define m where "m \ integral S (\x. \det (matrix (f' x))\)" have *: "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" if "e > 0" for e proof - define T where "T \ \n. {x \ S. n * e \ \det (matrix (f' x))\ \ \det (matrix (f' x))\ < (Suc n) * e}" have meas_t: "T n \ lmeasurable" for n proof - have *: "(\x. \det (matrix (f' x))\) \ borel_measurable (lebesgue_on S)" using aint_S by (simp add: S borel_measurable_restrict_space_iff fmeasurableD set_integrable_def) have [intro]: "x \ sets (lebesgue_on S) \ x \ sets lebesgue" for x using S sets_restrict_space_subset by blast have "{x \ S. real n * e \ \det (matrix (f' x))\} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_ge space_restrict_space) then have 1: "{x \ S. real n * e \ \det (matrix (f' x))\} \ lmeasurable" using S by (simp add: fmeasurableI2) have "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ sets lebesgue" using * by (auto simp: borel_measurable_iff_halfspace_less space_restrict_space) then have 2: "{x \ S. \det (matrix (f' x))\ < (1 + real n) * e} \ lmeasurable" using S by (simp add: fmeasurableI2) show ?thesis using fmeasurable.Int [OF 1 2] by (simp add: T_def Int_def cong: conj_cong) qed have aint_T: "\k. (\x. \det (matrix (f' x))\) absolutely_integrable_on T k" using set_integrable_subset [OF aint_S] meas_t T_def by blast have Seq: "S = (\n. T n)" apply (auto simp: T_def) apply (rule_tac x="nat(floor(abs(det(matrix(f' x))) / e))" in exI) using that apply auto using of_int_floor_le pos_le_divide_eq apply blast by (metis add.commute pos_divide_less_eq real_of_int_floor_add_one_gt) have meas_ft: "f ` T n \ lmeasurable" for n proof (rule measurable_bounded_differentiable_image) show "T n \ lmeasurable" by (simp add: meas_t) next fix x :: "(real,'n) vec" assume "x \ T n" show "(f has_derivative f' x) (at x within T n)" by (metis (no_types, lifting) \x \ T n\ deriv has_derivative_within_subset mem_Collect_eq subsetI T_def) show "\det (matrix (f' x))\ \ (Suc n) * e" using \x \ T n\ T_def by auto next show "(\x. \det (matrix (f' x))\) integrable_on T n" using aint_T absolutely_integrable_on_def by blast qed have disT: "disjoint (range T)" unfolding disjoint_def proof clarsimp show "T m \ T n = {}" if "T m \ T n" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) with \e > 0\ show ?case unfolding T_def proof (clarsimp simp add: Collect_conj_eq [symmetric]) fix x assume "e > 0" "m < n" "n * e \ \det (matrix (f' x))\" "\det (matrix (f' x))\ < (1 + real m) * e" then have "n < 1 + real m" by (metis (no_types, hide_lams) less_le_trans mult.commute not_le real_mult_le_cancel_iff2) then show "False" using less.hyps by linarith qed qed auto qed have injT: "inj_on T ({n. T n \ {}})" unfolding inj_on_def proof clarsimp show "m = n" if "T m = T n" "T n \ {}" for m n using that proof (induction m n rule: linorder_less_wlog) case (less m n) have False if "T n \ T m" "x \ T n" for x using \e > 0\ \m < n\ that apply (auto simp: T_def mult.commute intro: less_le_trans dest!: subsetD) by (metis add.commute less_le_trans nat_less_real_le not_le real_mult_le_cancel_iff2) then show ?case using less.prems by blast qed auto qed have sum_eq_Tim: "(\k\n. f (T k)) = sum f (T ` {..n})" if "f {} = 0" for f :: "_ \ real" and n proof (subst sum.reindex_nontrivial) fix i j assume "i \ {..n}" "j \ {..n}" "i \ j" "T i = T j" with that injT [unfolded inj_on_def] show "f (T i) = 0" by simp metis qed (use atMost_atLeast0 in auto) let ?B = "m + e * ?\ S" have "(\k\n. ?\ (f ` T k)) \ ?B" for n proof - have "(\k\n. ?\ (f ` T k)) \ (\k\n. ((k+1) * e) * ?\(T k))" proof (rule sum_mono [OF measure_bounded_differentiable_image]) show "(f has_derivative f' x) (at x within T k)" if "x \ T k" for k x using that unfolding T_def by (blast intro: deriv has_derivative_within_subset) show "(\x. \det (matrix (f' x))\) integrable_on T k" for k using absolutely_integrable_on_def aint_T by blast show "\det (matrix (f' x))\ \ real (k + 1) * e" if "x \ T k" for k x using T_def that by auto qed (use meas_t in auto) also have "\ \ (\k\n. (k * e) * ?\(T k)) + (\k\n. e * ?\(T k))" by (simp add: algebra_simps sum.distrib) also have "\ \ ?B" proof (rule add_mono) have "(\k\n. real k * e * ?\ (T k)) = (\k\n. integral (T k) (\x. k * e))" by (simp add: lmeasure_integral [OF meas_t] flip: integral_mult_right integral_mult_left) also have "\ \ (\k\n. integral (T k) (\x. (abs (det (matrix (f' x))))))" proof (rule sum_mono) fix k assume "k \ {..n}" show "integral (T k) (\x. k * e) \ integral (T k) (\x. \det (matrix (f' x))\)" proof (rule integral_le [OF integrable_on_const [OF meas_t]]) show "(\x. \det (matrix (f' x))\) integrable_on T k" using absolutely_integrable_on_def aint_T by blast next fix x assume "x \ T k" show "k * e \ \det (matrix (f' x))\" using \x \ T k\ T_def by blast qed qed also have "\ = sum (\T. integral T (\x. \det (matrix (f' x))\)) (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = integral (\k\n. T k) (\x. \det (matrix (f' x))\)" proof (rule integral_unique [OF has_integral_Union, symmetric]) fix S assume "S \ T ` {..n}" then show "((\x. \det (matrix (f' x))\) has_integral integral S (\x. \det (matrix (f' x))\)) S" using absolutely_integrable_on_def aint_T by blast next show "pairwise (\S S'. negligible (S \ S')) (T ` {..n})" using disT unfolding disjnt_iff by (auto simp: pairwise_def intro!: empty_imp_negligible) qed auto also have "\ \ m" unfolding m_def proof (rule integral_subset_le) have "(\x. \det (matrix (f' x))\) absolutely_integrable_on (\k\n. T k)" apply (rule set_integrable_subset [OF aint_S]) apply (intro measurable meas_t fmeasurableD) apply (force simp: Seq) done then show "(\x. \det (matrix (f' x))\) integrable_on (\k\n. T k)" using absolutely_integrable_on_def by blast qed (use Seq int in auto) finally show "(\k\n. real k * e * ?\ (T k)) \ m" . next have "(\k\n. ?\ (T k)) = sum ?\ (T ` {..n})" by (auto intro: sum_eq_Tim) also have "\ = ?\ (\k\n. T k)" using S disT by (auto simp: pairwise_def meas_t intro: measure_Union' [symmetric]) also have "\ \ ?\ S" using S by (auto simp: Seq intro: meas_t fmeasurableD measure_mono_fmeasurable) finally have "(\k\n. ?\ (T k)) \ ?\ S" . then show "(\k\n. e * ?\ (T k)) \ e * ?\ S" by (metis less_eq_real_def ordered_comm_semiring_class.comm_mult_left_mono sum_distrib_left that) qed finally show "(\k\n. ?\ (f ` T k)) \ ?B" . qed moreover have "measure lebesgue (\k\n. f ` T k) \ (\k\n. ?\ (f ` T k))" for n by (simp add: fmeasurableD meas_ft measure_UNION_le) ultimately have B_ge_m: "?\ (\k\n. (f ` T k)) \ ?B" for n by (meson order_trans) have "(\n. f ` T n) \ lmeasurable" by (rule fmeasurable_countable_Union [OF meas_ft B_ge_m]) moreover have "?\ (\n. f ` T n) \ m + e * ?\ S" by (rule measure_countable_Union_le [OF meas_ft B_ge_m]) ultimately show "f ` S \ lmeasurable" "?\ (f ` S) \ m + e * ?\ S" by (auto simp: Seq image_Union) qed show ?thesis proof show "f ` S \ lmeasurable" using * linordered_field_no_ub by blast let ?x = "m - ?\ (f ` S)" have False if "?\ (f ` S) > integral S (\x. \det (matrix (f' x))\)" proof - have ml: "m < ?\ (f ` S)" using m_def that by blast then have "?\ S \ 0" using "*"(2) bgauge_existence_lemma by fastforce with ml have 0: "0 < - (m - ?\ (f ` S))/2 / ?\ S" using that zero_less_measure_iff by force then show ?thesis using * [OF 0] that by (auto simp: field_split_simps m_def split: if_split_asm) qed then show "?\ (f ` S) \ integral S (\x. \det (matrix (f' x))\)" by fastforce qed qed theorem fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and deriv: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and int: "(\x. \det (matrix (f' x))\) integrable_on S" shows measurable_differentiable_image: "f ` S \ lmeasurable" and measure_differentiable_image: "measure lebesgue (f ` S) \ integral S (\x. \det (matrix (f' x))\)" (is "?M") proof - let ?I = "\n::nat. cbox (vec (-n)) (vec n) \ S" let ?\ = "measure lebesgue" have "x \ cbox (vec (- real (nat \norm x\))) (vec (real (nat \norm x\)))" for x :: "real^'n::_" apply (auto simp: mem_box_cart) apply (metis abs_le_iff component_le_norm_cart minus_le_iff of_nat_ceiling order.trans) by (meson abs_le_D1 norm_bound_component_le_cart real_nat_ceiling_ge) then have Seq: "S = (\n. ?I n)" by auto have fIn: "f ` ?I n \ lmeasurable" and mfIn: "?\ (f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" (is ?MN) for n proof - have In: "?I n \ lmeasurable" by (simp add: S bounded_Int bounded_set_imp_lmeasurable sets.Int) moreover have "\x. x \ ?I n \ (f has_derivative f' x) (at x within ?I n)" by (meson Int_iff deriv has_derivative_within_subset subsetI) moreover have int_In: "(\x. \det (matrix (f' x))\) integrable_on ?I n" proof - have "(\x. \det (matrix (f' x))\) absolutely_integrable_on S" using int absolutely_integrable_integrable_bound by force then have "(\x. \det (matrix (f' x))\) absolutely_integrable_on ?I n" by (metis (no_types) Int_lower1 In fmeasurableD inf_commute set_integrable_subset) then show ?thesis using absolutely_integrable_on_def by blast qed ultimately have "f ` ?I n \ lmeasurable" "?\ (f ` ?I n) \ integral (?I n) (\x. \det (matrix (f' x))\)" using m_diff_image_weak by metis+ moreover have "integral (?I n) (\x. \det (matrix (f' x))\) \ integral S (\x. \det (matrix (f' x))\)" by (simp add: int_In int integral_subset_le) ultimately show "f ` ?I n \ lmeasurable" ?MN by auto qed have "?I k \ ?I n" if "k \ n" for k n by (rule Int_mono) (use that in \auto simp: subset_interval_imp_cart\) then have "(\k\n. f ` ?I k) = f ` ?I n" for n by (fastforce simp add:) with mfIn have "?\ (\k\n. f ` ?I k) \ integral S (\x. \det (matrix (f' x))\)" for n by simp then have "(\n. f ` ?I n) \ lmeasurable" "?\ (\n. f ` ?I n) \ integral S (\x. \det (matrix (f' x))\)" by (rule fmeasurable_countable_Union [OF fIn] measure_countable_Union_le [OF fIn])+ then show "f ` S \ lmeasurable" ?M by (metis Seq image_UN)+ qed lemma borel_measurable_simple_function_limit_increasing: fixes f :: "'a::euclidean_space \ real" shows "(f \ borel_measurable lebesgue \ (\x. 0 \ f x)) \ (\g. (\n x. 0 \ g n x \ g n x \ f x) \ (\n x. g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \ (\x. (\n. g n x) \ f x))" (is "?lhs = ?rhs") proof assume f: ?lhs have leb_f: "{x. a \ f x \ f x < b} \ sets lebesgue" for a b proof - have "{x. a \ f x \ f x < b} = {x. f x < b} - {x. f x < a}" by auto also have "\ \ sets lebesgue" using borel_measurable_vimage_halfspace_component_lt [of f UNIV] f by auto finally show ?thesis . qed have "g n x \ f x" if inc_g: "\n x. 0 \ g n x \ g n x \ g (Suc n) x" and meas_g: "\n. g n \ borel_measurable lebesgue" and fin: "\n. finite(range (g n))" and lim: "\x. (\n. g n x) \ f x" for g n x proof - have "\r>0. \N. \n\N. dist (g n x) (f x) \ r" if "g n x > f x" proof - have g: "g n x \ g (N + n) x" for N by (rule transitive_stepwise_le) (use inc_g in auto) have "\na\N. g n x - f x \ dist (g na x) (f x)" for N apply (rule_tac x="N+n" in exI) using g [of N] by (auto simp: dist_norm) with that show ?thesis using diff_gt_0_iff_gt by blast qed with lim show ?thesis apply (auto simp: lim_sequentially) by (meson less_le_not_le not_le_imp_less) qed moreover let ?\ = "\n k. indicator {y. k/2^n \ f y \ f y < (k+1)/2^n}" let ?g = "\n x. (\k::real | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x)" have "\g. (\n x. 0 \ g n x \ g n x \ (g(Suc n) x)) \ (\n. g n \ borel_measurable lebesgue) \ (\n. finite(range (g n))) \(\x. (\n. g n x) \ f x)" proof (intro exI allI conjI) show "0 \ ?g n x" for n x proof (clarify intro!: ordered_comm_monoid_add_class.sum_nonneg) fix k::real assume "k \ \" and k: "\k\ \ 2 ^ (2*n)" show "0 \ k/2^n * ?\ n k x" using f \k \ \\ apply (auto simp: indicator_def field_split_simps Ints_def) apply (drule spec [where x=x]) using zero_le_power [of "2::real" n] mult_nonneg_nonneg [of "f x" "2^n"] by linarith qed show "?g n x \ ?g (Suc n) x" for n x proof - have "?g n x = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * (indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x + indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x))" by (rule sum.cong [OF refl]) (simp add: indicator_def field_split_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. k/2^n \ f y \ f y < (k+1/2)/2^n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * indicator {y. (k+1/2)/2^n \ f y \ f y < (k+1)/2^n} x)" by (simp add: comm_monoid_add_class.sum.distrib algebra_simps) also have "\ = (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k)/2 ^ Suc n \ f y \ f y < (2 * k+1)/2 ^ Suc n} x) + (\k | k \ \ \ \k\ \ 2 ^ (2*n). (2 * k)/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2 * k+1) + 1)/2 ^ Suc n} x)" by (force simp: field_simps indicator_def intro: sum.cong) also have "\ \ (\k | k \ \ \ \k\ \ 2 ^ (2 * Suc n). k/2 ^ Suc n * (indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x))" (is "?a + _ \ ?b") proof - have *: "\sum f I \ sum h I; a + sum h I \ b\ \ a + sum f I \ b" for I a b f and h :: "real\real" by linarith let ?h = "\k. (2*k+1)/2 ^ Suc n * (indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < ((2*k+1) + 1)/2 ^ Suc n} x)" show ?thesis proof (rule *) show "(\k | k \ \ \ \k\ \ 2 ^ (2*n). 2 * k/2 ^ Suc n * indicator {y. (2 * k+1)/2 ^ Suc n \ f y \ f y < (2 * k+1 + 1)/2 ^ Suc n} x) \ sum ?h {k \ \. \k\ \ 2 ^ (2*n)}" by (rule sum_mono) (simp add: indicator_def field_split_simps) next have \: "?a = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have \: "sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" by (auto simp: inj_on_def field_simps comm_monoid_add_class.sum.reindex) have 0: "(*) 2 ` {k \ \. P k} \ (\x. 2 * x + 1) ` {k \ \. P k} = {}" for P :: "real \ bool" proof - have "2 * i \ 2 * j + 1" for i j :: int by arith thus ?thesis unfolding Ints_def by auto (use of_int_eq_iff in fastforce) qed have "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} = (\k \ (*)2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)}. k/2 ^ Suc n * indicator {y. k/2 ^ Suc n \ f y \ f y < (k+1)/2 ^ Suc n} x)" unfolding \ \ using finite_abs_int_segment [of "2 ^ (2*n)"] by (subst sum_Un) (auto simp: 0) also have "\ \ ?b" proof (rule sum_mono2) show "finite {k::real. k \ \ \ \k\ \ 2 ^ (2 * Suc n)}" by (rule finite_abs_int_segment) show "(*) 2 ` {k::real. k \ \ \ \k\ \ 2^(2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2^(2*n)} \ {k \ \. \k\ \ 2 ^ (2 * Suc n)}" apply auto using one_le_power [of "2::real" "2*n"] by linarith have *: "\x \ (S \ T) - U; \x. x \ S \ x \ U; \x. x \ T \ x \ U\ \ P x" for S T U P by blast have "0 \ b" if "b \ \" "f x * (2 * 2^n) < b + 1" for b proof - have "0 \ f x * (2 * 2^n)" by (simp add: f) also have "\ < b+1" by (simp add: that) finally show "0 \ b" using \b \ \\ by (auto simp: elim!: Ints_cases) qed then show "0 \ b/2 ^ Suc n * indicator {y. b/2 ^ Suc n \ f y \ f y < (b + 1)/2 ^ Suc n} x" if "b \ {k \ \. \k\ \ 2 ^ (2 * Suc n)} - ((*) 2 ` {k \ \. \k\ \ 2 ^ (2*n)} \ (\x. 2*x + 1) ` {k \ \. \k\ \ 2 ^ (2*n)})" for b using that by (simp add: indicator_def divide_simps) qed finally show "?a + sum ?h {k \ \. \k\ \ 2 ^ (2*n)} \ ?b" . qed qed finally show ?thesis . qed show "?g n \ borel_measurable lebesgue" for n apply (intro borel_measurable_indicator borel_measurable_times borel_measurable_sum) using leb_f sets_restrict_UNIV by auto show "finite (range (?g n))" for n proof - have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) \ (\k. k/2^n) ` {k \ \. \k\ \ 2 ^ (2*n)}" for x proof (cases "\k. k \ \ \ \k\ \ 2 ^ (2*n) \ k/2^n \ f x \ f x < (k+1)/2^n") case True then show ?thesis by (blast intro: indicator_sum_eq) next case False then have "(\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) = 0" by auto then show ?thesis by force qed then have "range (?g n) \ ((\k. (k/2^n)) ` {k. k \ \ \ \k\ \ 2 ^ (2*n)})" by auto moreover have "finite ((\k::real. (k/2^n)) ` {k \ \. \k\ \ 2 ^ (2*n)})" by (intro finite_imageI finite_abs_int_segment) ultimately show ?thesis by (rule finite_subset) qed show "(\n. ?g n x) \ f x" for x proof (clarsimp simp add: lim_sequentially) fix e::real assume "e > 0" obtain N1 where N1: "2 ^ N1 > abs(f x)" using real_arch_pow by fastforce obtain N2 where N2: "(1/2) ^ N2 < e" using real_arch_pow_inv \e > 0\ by fastforce have "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) < e" if "N1 + N2 \ n" for n proof - let ?m = "real_of_int \2^n * f x\" have "\?m\ \ 2^n * 2^N1" using N1 apply (simp add: f) by (meson floor_mono le_floor_iff less_le_not_le mult_le_cancel_left_pos zero_less_numeral zero_less_power) also have "\ \ 2 ^ (2*n)" by (metis that add_leD1 add_le_cancel_left mult.commute mult_2_right one_less_numeral_iff power_add power_increasing_iff semiring_norm(76)) finally have m_le: "\?m\ \ 2 ^ (2*n)" . have "?m/2^n \ f x" "f x < (?m + 1)/2^n" by (auto simp: mult.commute pos_divide_le_eq mult_imp_less_div_pos) then have eq: "dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k/2^n * ?\ n k x) (f x) = dist (?m/2^n) (f x)" by (subst indicator_sum_eq [of ?m]) (auto simp: m_le) have "\2^n\ * \?m/2^n - f x\ = \2^n * (?m/2^n - f x)\" by (simp add: abs_mult) also have "\ < 2 ^ N2 * e" using N2 by (simp add: divide_simps mult.commute) linarith also have "\ \ \2^n\ * e" using that \e > 0\ by auto finally have "dist (?m/2^n) (f x) < e" by (simp add: dist_norm) then show ?thesis using eq by linarith qed then show "\no. \n\no. dist (\k | k \ \ \ \k\ \ 2 ^ (2*n). k * ?\ n k x/2^n) (f x) < e" by force qed qed ultimately show ?rhs by metis next assume RHS: ?rhs with borel_measurable_simple_function_limit [of f UNIV, unfolded lebesgue_on_UNIV_eq] show ?lhs by (blast intro: order_trans) qed subsection\Borel measurable Jacobian determinant\ lemma lemma_partial_derivatives0: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim0: "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x. x \ S - {0} \ norm x < e \ k * norm x \ \v \ x\)" shows "f x = 0" proof - interpret linear f by fact have "dim {x. f x = 0} \ DIM('a)" by (rule dim_subset_UNIV) moreover have False if less: "dim {x. f x = 0} < DIM('a)" proof - obtain d where "d \ 0" and d: "\y. f y = 0 \ d \ y = 0" using orthogonal_to_subspace_exists [OF less] orthogonal_def by (metis (mono_tags, lifting) mem_Collect_eq span_base) then obtain k where "k > 0" and k: "\e. e > 0 \ \y. y \ S - {0} \ norm y < e \ k * norm y \ \d \ y\" using lb by blast have "\h. \n. ((h n \ S \ h n \ 0 \ k * norm (h n) \ \d \ h n\) \ norm (h n) < 1 / real (Suc n)) \ norm (h (Suc n)) < norm (h n)" proof (rule dependent_nat_choice) show "\y. (y \ S \ y \ 0 \ k * norm y \ \d \ y\) \ norm y < 1 / real (Suc 0)" by simp (metis DiffE insertCI k not_less not_one_le_zero) qed (use k [of "min (norm x) (1/(Suc n + 1))" for x n] in auto) then obtain \ where \: "\n. \ n \ S - {0}" and kd: "\n. k * norm(\ n) \ \d \ \ n\" and norm_lt: "\n. norm(\ n) < 1/(Suc n)" by force let ?\ = "\n. \ n /\<^sub>R norm (\ n)" have com: "\g. (\n. g n \ sphere (0::'a) 1) \ \l \ sphere 0 1. \\::nat\nat. strict_mono \ \ (g \ \) \ l" using compact_sphere compact_def by metis moreover have "\n. ?\ n \ sphere 0 1" using \ by auto ultimately obtain l::'a and \::"nat\nat" where l: "l \ sphere 0 1" and "strict_mono \" and to_l: "(?\ \ \) \ l" by meson moreover have "continuous (at l) (\x. (\d \ x\ - k))" by (intro continuous_intros) ultimately have lim_dl: "((\x. (\d \ x\ - k)) \ (?\ \ \)) \ (\d \ l\ - k)" by (meson continuous_imp_tendsto) have "\\<^sub>F i in sequentially. 0 \ ((\x. \d \ x\ - k) \ ((\n. \ n /\<^sub>R norm (\ n)) \ \)) i" using \ kd by (auto simp: field_split_simps) then have "k \ \d \ l\" using tendsto_lowerbound [OF lim_dl, of 0] by auto moreover have "d \ l = 0" proof (rule d) show "f l = 0" proof (rule LIMSEQ_unique [of "f \ ?\ \ \"]) have "isCont f l" using \linear f\ linear_continuous_at linear_conv_bounded_linear by blast then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ f l" unfolding comp_assoc using to_l continuous_imp_tendsto by blast have "\ \ 0" using norm_lt LIMSEQ_norm_0 by metis with \strict_mono \\ have "(\ \ \) \ 0" by (metis LIMSEQ_subseq_LIMSEQ) with lim0 \ have "((\x. f x /\<^sub>R norm x) \ (\ \ \)) \ 0" by (force simp: tendsto_at_iff_sequentially) then show "(f \ (\n. \ n /\<^sub>R norm (\ n)) \ \) \ 0" by (simp add: o_def scale) qed qed ultimately show False using \k > 0\ by auto qed ultimately have dim: "dim {x. f x = 0} = DIM('a)" by force then show ?thesis using dim_eq_full by (metis (mono_tags, lifting) eq_0_on_span eucl.span_Basis linear_axioms linear_eq_stdbasis mem_Collect_eq module_hom_zero span_base span_raw_def) qed lemma lemma_partial_derivatives: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" and lim: "((\x. f (x - a) /\<^sub>R norm (x - a)) \ 0) (at a within S)" and lb: "\v. v \ 0 \ (\k>0. \e>0. \x \ S - {a}. norm(a - x) < e \ k * norm(a - x) \ \v \ (x - a)\)" shows "f x = 0" proof - have "((\x. f x /\<^sub>R norm x) \ 0) (at 0 within (\x. x-a) ` S)" using lim by (simp add: Lim_within dist_norm) then show ?thesis proof (rule lemma_partial_derivatives0 [OF \linear f\]) fix v :: "'a" assume v: "v \ 0" show "\k>0. \e>0. \x. x \ (\x. x - a) ` S - {0} \ norm x < e \ k * norm x \ \v \ x\" using lb [OF v] by (force simp: norm_minus_commute) qed qed proposition borel_measurable_partial_derivatives: fixes f :: "real^'m::{finite,wellorder} \ real^'n" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. (matrix(f' x)$m$n)) \ borel_measurable (lebesgue_on S)" proof - have contf: "continuous_on S f" using continuous_on_eq_continuous_within f has_derivative_continuous by blast have "{x \ S. (matrix (f' x)$m$n) \ b} \ sets lebesgue" for b proof (rule sets_negligible_symdiff) let ?T = "{x \ S. \e>0. \d>0. \A. A$m$n < b \ (\i j. A$i$j \ \) \ (\y \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x))}" let ?U = "S \ (\e \ {e \ \. e > 0}. \A \ {A. A$m$n < b \ (\i j. A$i$j \ \)}. \d \ {d \ \. 0 < d}. S \ (\y \ S. {x \ S. norm(y - x) < d \ norm(f y - f x - A *v (y - x)) \ e * norm(y - x)}))" have "?T = ?U" proof (intro set_eqI iffI) fix x assume xT: "x \ ?T" then show "x \ ?U" proof (clarsimp simp add:) fix q :: real assume "q \ \" "q > 0" then obtain d A where "d > 0" and A: "A $ m $ n < b" "\i j. A $ i $ j \ \" "\y. \y\S; norm (y - x) < d\ \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)" using xT by auto then obtain \ where "d > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with A show "\A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\s. s \ \ \ 0 < s \ (\y\S. norm (y - x) < s \ norm (f y - f x - A *v (y - x)) \ q * norm (y - x)))" by force qed next fix x assume xU: "x \ ?U" then show "x \ ?T" proof clarsimp fix e :: "real" assume "e > 0" then obtain \ where \: "e > \" "\ > 0" "\ \ \" using Rats_dense_in_real by blast with xU obtain A r where "x \ S" and Ar: "A $ m $ n < b" "\i j. A $ i $ j \ \" "r \ \" "r > 0" and "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ \ * norm (y - x)" by (auto simp: split: if_split_asm) then have "\y\S. norm (y - x) < r \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)" by (meson \e > \\ less_eq_real_def mult_right_mono norm_ge_zero order_trans) then show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" using \x \ S\ Ar by blast qed qed moreover have "?U \ sets lebesgue" proof - have coQ: "countable {e \ \. 0 < e}" using countable_Collect countable_rat by blast have ne: "{e \ \. (0::real) < e} \ {}" using zero_less_one Rats_1 by blast have coA: "countable {A. A $ m $ n < b \ (\i j. A $ i $ j \ \)}" proof (rule countable_subset) show "countable {A. \i j. A $ i $ j \ \}" using countable_vector [OF countable_vector, of "\i j. \"] by (simp add: countable_rat) qed blast have *: "\U \ {} \ closedin (top_of_set S) (S \ \ U)\ \ closedin (top_of_set S) (S \ \ U)" for U by fastforce have eq: "{x::(real,'m)vec. P x \ (Q x \ R x)} = {x. P x \ \ Q x} \ {x. P x \ R x}" for P Q R by auto have sets: "S \ (\y\S. {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}) \ sets lebesgue" for e A d proof - have clo: "closedin (top_of_set S) {x \ S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x)}" for y proof - have cont1: "continuous_on S (\x. norm (y - x))" and cont2: "continuous_on S (\x. e * norm (y - x) - norm (f y - f x - (A *v y - A *v x)))" by (force intro: contf continuous_intros)+ have clo1: "closedin (top_of_set S) {x \ S. d \ norm (y - x)}" using continuous_closedin_preimage [OF cont1, of "{d..}"] by (simp add: vimage_def Int_def) have clo2: "closedin (top_of_set S) {x \ S. norm (f y - f x - (A *v y - A *v x)) \ e * norm (y - x)}" using continuous_closedin_preimage [OF cont2, of "{0..}"] by (simp add: vimage_def Int_def) show ?thesis by (auto simp: eq not_less matrix_vector_mult_diff_distrib intro: clo1 clo2) qed show ?thesis by (rule lebesgue_closedin [of S]) (force intro: * S clo)+ qed show ?thesis by (intro sets sets.Int S sets.countable_UN'' sets.countable_INT'' coQ coA) auto qed ultimately show "?T \ sets lebesgue" by simp let ?M = "(?T - {x \ S. matrix (f' x) $ m $ n \ b} \ ({x \ S. matrix (f' x) $ m $ n \ b} - ?T))" let ?\ = "\x v. \\>0. \e>0. \y \ S-{x}. norm (x - y) < e \ \v \ (y - x)\ < \ * norm (x - y)" have nN: "negligible {x \ S. \v\0. ?\ x v}" unfolding negligible_eq_zero_density proof clarsimp fix x v and r e :: "real" assume "x \ S" "v \ 0" "r > 0" "e > 0" and Theta [rule_format]: "?\ x v" moreover have "(norm v * e / 2) / CARD('m) ^ CARD('m) > 0" by (simp add: \v \ 0\ \e > 0\) ultimately obtain d where "d > 0" and dless: "\y. \y \ S - {x}; norm (x - y) < d\ \ \v \ (y - x)\ < ((norm v * e / 2) / CARD('m) ^ CARD('m)) * norm (x - y)" by metis let ?W = "ball x (min d r) \ {y. \v \ (y - x)\ < (norm v * e/2 * min d r) / CARD('m) ^ CARD('m)}" have "open {x. \v \ (x - a)\ < b}" for a b by (intro open_Collect_less continuous_intros) show "\d>0. d \ r \ (\U. {x' \ S. \v\0. ?\ x' v} \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d))" proof (intro exI conjI) show "0 < min d r" "min d r \ r" using \r > 0\ \d > 0\ by auto show "{x' \ S. \v. v \ 0 \ (\\>0. \e>0. \z\S - {x'}. norm (x' - z) < e \ \v \ (z - x')\ < \ * norm (x' - z))} \ ball x (min d r) \ ?W" proof (clarsimp simp: dist_norm norm_minus_commute) fix y w assume "y \ S" "w \ 0" and less [rule_format]: "\\>0. \e>0. \z\S - {y}. norm (y - z) < e \ \w \ (z - y)\ < \ * norm (y - z)" and d: "norm (y - x) < d" and r: "norm (y - x) < r" show "\v \ (y - x)\ < norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" proof (cases "y = x") case True with \r > 0\ \d > 0\ \e > 0\ \v \ 0\ show ?thesis by simp next case False have "\v \ (y - x)\ < norm v * e / 2 / real (CARD('m) ^ CARD('m)) * norm (x - y)" apply (rule dless) using False \y \ S\ d by (auto simp: norm_minus_commute) also have "\ \ norm v * e * min d r / (2 * real CARD('m) ^ CARD('m))" using d r \e > 0\ by (simp add: field_simps norm_minus_commute mult_left_mono) finally show ?thesis . qed qed show "?W \ lmeasurable" by (simp add: fmeasurable_Int_fmeasurable borel_open) obtain k::'m where True by metis obtain T where T: "orthogonal_transformation T" and v: "v = T(norm v *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis define b where "b \ norm v" have "b > 0" using \v \ 0\ by (auto simp: b_def) obtain eqb: "inv T v = b *\<^sub>R axis k (1::real)" and "inj T" "bij T" and invT: "orthogonal_transformation (inv T)" by (metis UNIV_I b_def T v bij_betw_inv_into_left orthogonal_transformation_inj orthogonal_transformation_bij orthogonal_transformation_inv) let ?v = "\ i. min d r / CARD('m)" let ?v' = "\ i. if i = k then (e/2 * min d r) / CARD('m) ^ CARD('m) else min d r" let ?x' = "inv T x" let ?W' = "(ball ?x' (min d r) \ {y. \(y - ?x')$k\ < e * min d r / (2 * CARD('m) ^ CARD('m))})" have abs: "x - e \ y \ y \ x + e \ abs(y - x) \ e" for x y e::real by auto have "?W = T ` ?W'" proof - have 1: "T ` (ball (inv T x) (min d r)) = ball x (min d r)" by (simp add: T image_orthogonal_transformation_ball orthogonal_transformation_surj surj_f_inv_f) have 2: "{y. \v \ (y - x)\ < b * e * min d r / (2 * real CARD('m) ^ CARD('m))} = T ` {y. \y $ k - ?x' $ k\ < e * min d r / (2 * real CARD('m) ^ CARD('m))}" proof - have *: "\T (b *\<^sub>R axis k 1) \ (y - x)\ = b * \inv T y $ k - ?x' $ k\" for y proof - have "\T (b *\<^sub>R axis k 1) \ (y - x)\ = \(b *\<^sub>R axis k 1) \ inv T (y - x)\" by (metis (no_types, hide_lams) b_def eqb invT orthogonal_transformation_def v) also have "\ = b * \(axis k 1) \ inv T (y - x)\" using \b > 0\ by (simp add: abs_mult) also have "\ = b * \inv T y $ k - ?x' $ k\" using orthogonal_transformation_linear [OF invT] by (simp add: inner_axis' linear_diff) finally show ?thesis by simp qed show ?thesis using v b_def [symmetric] using \b > 0\ by (simp add: * bij_image_Collect_eq [OF \bij T\] mult_less_cancel_left_pos times_divide_eq_right [symmetric] del: times_divide_eq_right) qed show ?thesis using \b > 0\ by (simp add: image_Int \inj T\ 1 2 b_def [symmetric]) qed moreover have "?W' \ lmeasurable" by (auto intro: fmeasurable_Int_fmeasurable) ultimately have "measure lebesgue ?W = measure lebesgue ?W'" by (metis measure_orthogonal_image T) also have "\ \ measure lebesgue (cbox (?x' - ?v') (?x' + ?v'))" proof (rule measure_mono_fmeasurable) show "?W' \ cbox (?x' - ?v') (?x' + ?v')" apply (clarsimp simp add: mem_box_cart abs dist_norm norm_minus_commute simp del: min_less_iff_conj min.bounded_iff) by (metis component_le_norm_cart less_eq_real_def le_less_trans vector_minus_component) qed auto also have "\ \ e/2 * measure lebesgue (cbox (?x' - ?v) (?x' + ?v))" proof - have "cbox (?x' - ?v) (?x' + ?v) \ {}" using \r > 0\ \d > 0\ by (auto simp: interval_eq_empty_cart divide_less_0_iff) with \r > 0\ \d > 0\ \e > 0\ show ?thesis apply (simp add: content_cbox_if_cart mem_box_cart) apply (auto simp: prod_nonneg) apply (simp add: abs if_distrib prod.delta_remove field_simps power_diff split: if_split_asm) done qed also have "\ \ e/2 * measure lebesgue (cball ?x' (min d r))" proof (rule mult_left_mono [OF measure_mono_fmeasurable]) have *: "norm (?x' - y) \ min d r" if y: "\i. \?x' $ i - y $ i\ \ min d r / real CARD('m)" for y proof - have "norm (?x' - y) \ (\i\UNIV. \(?x' - y) $ i\)" by (rule norm_le_l1_cart) also have "\ \ real CARD('m) * (min d r / real CARD('m))" by (rule sum_bounded_above) (use y in auto) finally show ?thesis by simp qed show "cbox (?x' - ?v) (?x' + ?v) \ cball ?x' (min d r)" apply (clarsimp simp only: mem_box_cart dist_norm mem_cball intro!: *) by (simp add: abs_diff_le_iff abs_minus_commute) qed (use \e > 0\ in auto) also have "\ < e * content (cball ?x' (min d r))" using \r > 0\ \d > 0\ \e > 0\ by auto also have "\ = e * content (ball x (min d r))" using \r > 0\ \d > 0\ by (simp add: content_cball content_ball) finally show "measure lebesgue ?W < e * content (ball x (min d r))" . qed qed have *: "(\x. (x \ S) \ (x \ T \ x \ U)) \ (T - U) \ (U - T) \ S" for S T U :: "(real,'m) vec set" by blast have MN: "?M \ {x \ S. \v\0. ?\ x v}" proof (rule *) fix x assume x: "x \ {x \ S. \v\0. ?\ x v}" show "(x \ ?T) \ (x \ {x \ S. matrix (f' x) $ m $ n \ b})" proof (cases "x \ S") case True then have x: "\ ?\ x v" if "v \ 0" for v using x that by force show ?thesis proof (rule iffI; clarsimp) assume b: "\e>0. \d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" (is "\e>0. \d>0. \A. ?\ e d A") then have "\k. \d>0. \A. ?\ (1 / Suc k) d A" by (metis (no_types, hide_lams) less_Suc_eq_0_disj of_nat_0_less_iff zero_less_divide_1_iff) then obtain \ A where \: "\k. \ k > 0" and Ab: "\k. A k $ m $ n < b" and A: "\k y. \y \ S; norm (y - x) < \ k\ \ norm (f y - f x - A k *v (y - x)) \ 1/(Suc k) * norm (y - x)" by metis have "\i j. \a. (\n. A n $ i $ j) \ a" proof (intro allI) fix i j have vax: "(A n *v axis j 1) $ i = A n $ i $ j" for n by (metis cart_eq_inner_axis matrix_vector_mul_component) let ?CA = "{x. Cauchy (\n. (A n) *v x)}" have "subspace ?CA" unfolding subspace_def convergent_eq_Cauchy [symmetric] by (force simp: algebra_simps intro: tendsto_intros) then have CA_eq: "?CA = span ?CA" by (metis span_eq_iff) also have "\ = UNIV" proof - have "dim ?CA \ CARD('m)" using dim_subset_UNIV[of ?CA] by auto moreover have "False" if less: "dim ?CA < CARD('m)" proof - obtain d where "d \ 0" and d: "\y. y \ span ?CA \ orthogonal d y" using less by (force intro: orthogonal_to_subspace_exists [of ?CA]) with x [OF \d \ 0\] obtain \ where "\ > 0" and \: "\e. e > 0 \ \y \ S - {x}. norm (x - y) < e \ \ * norm (x - y) \ \d \ (y - x)\" by (fastforce simp: not_le Bex_def) obtain \ z where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \x: "\ \ x" and z: "(\n. (\ n - x) /\<^sub>R norm (\ n - x)) \ z" proof - have "\\. (\i. (\ i \ S - {x} \ \ * norm(\ i - x) \ \d \ (\ i - x)\ \ norm(\ i - x) < 1/Suc i) \ norm(\(Suc i) - x) < norm(\ i - x))" proof (rule dependent_nat_choice) show "\y. y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1 / Suc 0" using \ [of 1] by (auto simp: dist_norm norm_minus_commute) next fix y i assume "y \ S - {x} \ \ * norm (y - x) \ \d \ (y - x)\ \ norm (y - x) < 1/Suc i" then have "min (norm(y - x)) (1/((Suc i) + 1)) > 0" by auto then obtain y' where "y' \ S - {x}" and y': "norm (x - y') < min (norm (y - x)) (1/((Suc i) + 1))" "\ * norm (x - y') \ \d \ (y' - x)\" using \ by metis with \ show "\y'. (y' \ S - {x} \ \ * norm (y' - x) \ \d \ (y' - x)\ \ norm (y' - x) < 1/(Suc (Suc i))) \ norm (y' - x) < norm (y - x)" by (auto simp: dist_norm norm_minus_commute) qed then obtain \ where \Sx: "\i. \ i \ S - {x}" and \le: "\i. \ * norm(\ i - x) \ \d \ (\ i - x)\" and \conv: "\i. norm(\ i - x) < 1/(Suc i)" by blast let ?f = "\i. (\ i - x) /\<^sub>R norm (\ i - x)" have "?f i \ sphere 0 1" for i using \Sx by auto then obtain l \ where "l \ sphere 0 1" "strict_mono \" and l: "(?f \ \) \ l" using compact_sphere [of "0::(real,'m) vec" 1] unfolding compact_def by meson show thesis proof show "(\ \ \) i \ S - {x}" "\ * norm ((\ \ \) i - x) \ \d \ ((\ \ \) i - x)\" for i using \Sx \le by auto have "\ \ x" proof (clarsimp simp add: LIMSEQ_def dist_norm) fix r :: "real" assume "r > 0" with real_arch_invD obtain no where "no \ 0" "real no > 1/r" by (metis divide_less_0_1_iff not_less_iff_gr_or_eq of_nat_0_eq_iff reals_Archimedean2) with \conv show "\no. \n\no. norm (\ n - x) < r" by (metis \r > 0\ add.commute divide_inverse inverse_inverse_eq inverse_less_imp_less less_trans mult.left_neutral nat_le_real_less of_nat_Suc) qed with \strict_mono \\ show "(\ \ \) \ x" by (metis LIMSEQ_subseq_LIMSEQ) show "(\n. ((\ \ \) n - x) /\<^sub>R norm ((\ \ \) n - x)) \ l" using l by (auto simp: o_def) qed qed have "isCont (\x. (\d \ x\ - \)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\y. \d \ ((\ y - x) /\<^sub>R norm (\ y - x))\ - \) \ \d \ z\ - \" by auto moreover have "\\<^sub>F i in sequentially. 0 \ \d \ ((\ i - x) /\<^sub>R norm (\ i - x))\ - \" proof (rule eventuallyI) fix n show "0 \ \d \ ((\ n - x) /\<^sub>R norm (\ n - x))\ - \" using \le [of n] \Sx by (auto simp: abs_mult divide_simps) qed ultimately have "\ \ \d \ z\" using tendsto_lowerbound [where a=0] by fastforce have "Cauchy (\n. (A n) *v z)" proof (clarsimp simp add: Cauchy_def) fix \ :: "real" assume "0 < \" then obtain N::nat where "N > 0" and N: "\/2 > 1/N" by (metis half_gt_zero inverse_eq_divide neq0_conv real_arch_inverse) show "\M. \m\M. \n\M. dist (A m *v z) (A n *v z) < \" proof (intro exI allI impI) fix i j assume ij: "N \ i" "N \ j" let ?V = "\i k. A i *v ((\ k - x) /\<^sub>R norm (\ k - x))" have "\\<^sub>F k in sequentially. dist (\ k) x < min (\ i) (\ j)" using \x [unfolded tendsto_iff] by (meson min_less_iff_conj \) then have even: "\\<^sub>F k in sequentially. norm (?V i k - ?V j k) - 2 / N \ 0" proof (rule eventually_mono, clarsimp) fix p assume p: "dist (\ p) x < \ i" "dist (\ p) x < \ j" let ?C = "\k. f (\ p) - f x - A k *v (\ p - x)" have "norm ((A i - A j) *v (\ p - x)) = norm (?C j - ?C i)" by (simp add: algebra_simps) also have "\ \ norm (?C j) + norm (?C i)" using norm_triangle_ineq4 by blast also have "\ \ 1/(Suc j) * norm (\ p - x) + 1/(Suc i) * norm (\ p - x)" by (metis A Diff_iff \Sx dist_norm p add_mono) also have "\ \ 1/N * norm (\ p - x) + 1/N * norm (\ p - x)" apply (intro add_mono mult_right_mono) using ij \N > 0\ by (auto simp: field_simps) also have "\ = 2 / N * norm (\ p - x)" by simp finally have no_le: "norm ((A i - A j) *v (\ p - x)) \ 2 / N * norm (\ p - x)" . have "norm (?V i p - ?V j p) = norm ((A i - A j) *v ((\ p - x) /\<^sub>R norm (\ p - x)))" by (simp add: algebra_simps) also have "\ = norm ((A i - A j) *v (\ p - x)) / norm (\ p - x)" by (simp add: divide_inverse matrix_vector_mult_scaleR) also have "\ \ 2 / N" using no_le by (auto simp: field_split_simps) finally show "norm (?V i p - ?V j p) \ 2 / N" . qed have "isCont (\w. (norm(A i *v w - A j *v w) - 2 / N)) z" by (intro continuous_intros) from isCont_tendsto_compose [OF this z] have lim: "(\w. norm (A i *v ((\ w - x) /\<^sub>R norm (\ w - x)) - A j *v ((\ w - x) /\<^sub>R norm (\ w - x))) - 2 / N) \ norm (A i *v z - A j *v z) - 2 / N" by auto have "dist (A i *v z) (A j *v z) \ 2 / N" using tendsto_upperbound [OF lim even] by (auto simp: dist_norm) with N show "dist (A i *v z) (A j *v z) < \" by linarith qed qed then have "d \ z = 0" using CA_eq d orthogonal_def by auto then show False using \0 < \\ \\ \ \d \ z\\ by auto qed ultimately show ?thesis using dim_eq_full by fastforce qed finally have "?CA = UNIV" . then have "Cauchy (\n. (A n) *v axis j 1)" by auto then obtain L where "(\n. A n *v axis j 1) \ L" by (auto simp: Cauchy_convergent_iff convergent_def) then have "(\x. (A x *v axis j 1) $ i) \ L $ i" by (rule tendsto_vec_nth) then show "\a. (\n. A n $ i $ j) \ a" by (force simp: vax) qed then obtain B where B: "\i j. (\n. A n $ i $ j) \ B $ i $ j" by (auto simp: lambda_skolem) have lin_df: "linear (f' x)" and lim_df: "((\y. (1 / norm (y - x)) *\<^sub>R (f y - (f x + f' x (y - x)))) \ 0) (at x within S)" using \x \ S\ assms by (auto simp: has_derivative_within linear_linear) moreover interpret linear "f' x" by fact have "(matrix (f' x) - B) *v w = 0" for w proof (rule lemma_partial_derivatives [of "(*v) (matrix (f' x) - B)"]) show "linear ((*v) (matrix (f' x) - B))" by (rule matrix_vector_mul_linear) have "((\y. ((f x + f' x (y - x)) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" using tendsto_minus [OF lim_df] by (simp add: algebra_simps field_split_simps) then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x)) \ 0) (at x within S)" proof (rule Lim_transform) have "((\y. ((f y + B *v x - (f x + B *v y)) /\<^sub>R norm (y - x))) \ 0) (at x within S)" proof (clarsimp simp add: Lim_within dist_norm) fix e :: "real" assume "e > 0" then obtain q::nat where "q \ 0" and qe2: "1/q < e/2" by (metis divide_pos_pos inverse_eq_divide real_arch_inverse zero_less_numeral) let ?g = "\p. sum (\i. sum (\j. abs((A p - B)$i$j)) UNIV) UNIV" have "(\k. onorm (\y. (A k - B) *v y)) \ 0" proof (rule Lim_null_comparison) show "\\<^sub>F k in sequentially. norm (onorm (\y. (A k - B) *v y)) \ ?g k" proof (rule eventually_sequentiallyI) fix k :: "nat" assume "0 \ k" have "0 \ onorm ((*v) (A k - B))" using matrix_vector_mul_bounded_linear by (rule onorm_pos_le) then show "norm (onorm ((*v) (A k - B))) \ (\i\UNIV. \j\UNIV. \(A k - B) $ i $ j\)" by (simp add: onorm_le_matrix_component_sum del: vector_minus_component) qed next show "?g \ 0" using B Lim_null tendsto_rabs_zero_iff by (fastforce intro!: tendsto_null_sum) qed with \e > 0\ obtain p where "\n. n \ p \ \onorm ((*v) (A n - B))\ < e/2" unfolding lim_sequentially by (metis diff_zero dist_real_def divide_pos_pos zero_less_numeral) then have pqe2: "\onorm ((*v) (A (p + q) - B))\ < e/2" (*17 [`abs (onorm (\y. A (p + q) ** y - B ** y)) < e / &2`]*) using le_add1 by blast show "\d>0. \y\S. y \ x \ norm (y - x) < d \ inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" proof (intro exI, safe) show "0 < \(p + q)" by (simp add: \) next fix y assume y: "y \ S" "norm (y - x) < \(p + q)" and "y \ x" have *: "\norm(b - c) < e - d; norm(y - x - b) \ d\ \ norm(y - x - c) < e" for b c d e x and y:: "real^'n" using norm_triangle_ineq2 [of "y - x - c" "y - x - b"] by simp have "norm (f y - f x - B *v (y - x)) < e * norm (y - x)" proof (rule *) show "norm (f y - f x - A (p + q) *v (y - x)) \ norm (y - x) / (Suc (p + q))" using A [OF y] by simp have "norm (A (p + q) *v (y - x) - B *v (y - x)) \ onorm(\x. (A(p + q) - B) *v x) * norm(y - x)" by (metis linear_linear matrix_vector_mul_linear matrix_vector_mult_diff_rdistrib onorm) also have "\ < (e/2) * norm (y - x)" using \y \ x\ pqe2 by auto also have "\ \ (e - 1 / (Suc (p + q))) * norm (y - x)" proof (rule mult_right_mono) have "1 / Suc (p + q) \ 1 / q" using \q \ 0\ by (auto simp: field_split_simps) also have "\ < e/2" using qe2 by auto finally show "e / 2 \ e - 1 / real (Suc (p + q))" by linarith qed auto finally show "norm (A (p + q) *v (y - x) - B *v (y - x)) < e * norm (y - x) - norm (y - x) / real (Suc (p + q))" by (simp add: algebra_simps) qed then show "inverse (norm (y - x)) * norm (f y + B *v x - (f x + B *v y)) < e" using \y \ x\ by (simp add: field_split_simps algebra_simps) qed qed then show "((\y. (matrix (f' x) - B) *v (y - x) /\<^sub>R norm (y - x) - (f x + f' x (y - x) - f y) /\<^sub>R norm (y - x)) \ 0) (at x within S)" by (simp add: algebra_simps diff lin_df matrix_vector_mul_linear scalar_mult_eq_scaleR) qed qed (use x in \simp; auto simp: not_less\) ultimately have "f' x = (*v) B" by (force simp: algebra_simps scalar_mult_eq_scaleR) show "matrix (f' x) $ m $ n \ b" proof (rule tendsto_upperbound [of "\i. (A i $ m $ n)" _ sequentially]) show "(\i. A i $ m $ n) \ matrix (f' x) $ m $ n" by (simp add: B \f' x = (*v) B\) show "\\<^sub>F i in sequentially. A i $ m $ n \ b" by (simp add: Ab less_eq_real_def) qed auto next fix e :: "real" assume "x \ S" and b: "matrix (f' x) $ m $ n \ b" and "e > 0" then obtain d where "d>0" and d: "\y. y\S \ 0 < dist y x \ dist y x < d \ norm (f y - f x - f' x (y - x)) / (norm (y - x)) < e/2" using f [OF \x \ S\] by (simp add: Deriv.has_derivative_at_within Lim_within) (auto simp add: field_simps dest: spec [of _ "e/2"]) let ?A = "matrix(f' x) - (\ i j. if i = m \ j = n then e / 4 else 0)" obtain B where BRats: "\i j. B$i$j \ \" and Bo_e6: "onorm((*v) (?A - B)) < e/6" using matrix_rational_approximation \e > 0\ by (metis zero_less_divide_iff zero_less_numeral) show "\d>0. \A. A $ m $ n < b \ (\i j. A $ i $ j \ \) \ (\y\S. norm (y - x) < d \ norm (f y - f x - A *v (y - x)) \ e * norm (y - x))" proof (intro exI conjI ballI allI impI) show "d>0" by (rule \d>0\) show "B $ m $ n < b" proof - have "\matrix ((*v) (?A - B)) $ m $ n\ \ onorm ((*v) (?A - B))" using component_le_onorm [OF matrix_vector_mul_linear, of _ m n] by metis then show ?thesis using b Bo_e6 by simp qed show "B $ i $ j \ \" for i j using BRats by auto show "norm (f y - f x - B *v (y - x)) \ e * norm (y - x)" if "y \ S" and y: "norm (y - x) < d" for y proof (cases "y = x") case True then show ?thesis by simp next case False have *: "norm(d' - d) \ e/2 \ norm(y - (x + d')) < e/2 \ norm(y - x - d) \ e" for d d' e and x y::"real^'n" using norm_triangle_le [of "d' - d" "y - (x + d')"] by simp show ?thesis proof (rule *) have split246: "\norm y \ e / 6; norm(x - y) \ e / 4\ \ norm x \ e/2" if "e > 0" for e and x y :: "real^'n" using norm_triangle_le [of y "x-y" "e/2"] \e > 0\ by simp have "linear (f' x)" using True f has_derivative_linear by blast then have "norm (f' x (y - x) - B *v (y - x)) = norm ((matrix (f' x) - B) *v (y - x))" by (simp add: matrix_vector_mult_diff_rdistrib) also have "\ \ (e * norm (y - x)) / 2" proof (rule split246) have "norm ((?A - B) *v (y - x)) / norm (y - x) \ onorm(\x. (?A - B) *v x)" by (rule le_onorm) auto also have "\ < e/6" by (rule Bo_e6) finally have "norm ((?A - B) *v (y - x)) / norm (y - x) < e / 6" . then show "norm ((?A - B) *v (y - x)) \ e * norm (y - x) / 6" by (simp add: field_split_simps False) have "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) = norm ((\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x))" by (simp add: algebra_simps) also have "\ = norm((e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real))" proof - have "(\j\UNIV. (if i = m \ j = n then e / 4 else 0) * (y $ j - x $ j)) * 4 = e * (y $ n - x $ n) * axis m 1 $ i" for i proof (cases "i=m") case True then show ?thesis by (auto simp: if_distrib [of "\z. z * _"] cong: if_cong) next case False then show ?thesis by (simp add: axis_def) qed then have "(\ i j. if i = m \ j = n then e / 4 else 0) *v (y - x) = (e/4) *\<^sub>R (y - x)$n *\<^sub>R axis m (1::real)" by (auto simp: vec_eq_iff matrix_vector_mult_def) then show ?thesis by metis qed also have "\ \ e * norm (y - x) / 4" using \e > 0\ apply (simp add: norm_mult abs_mult) by (metis component_le_norm_cart vector_minus_component) finally show "norm ((matrix (f' x) - B) *v (y - x) - ((?A - B) *v (y - x))) \ e * norm (y - x) / 4" . show "0 < e * norm (y - x)" by (simp add: False \e > 0\) qed finally show "norm (f' x (y - x) - B *v (y - x)) \ (e * norm (y - x)) / 2" . show "norm (f y - (f x + f' x (y - x))) < (e * norm (y - x)) / 2" using False d [OF \y \ S\] y by (simp add: dist_norm field_simps) qed qed qed qed qed auto qed show "negligible ?M" using negligible_subset [OF nN MN] . qed then show ?thesis by (simp add: borel_measurable_vimage_halfspace_component_le sets_restrict_space_iff assms) qed theorem borel_measurable_det_Jacobian: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" shows "(\x. det(matrix(f' x))) \ borel_measurable (lebesgue_on S)" unfolding det_def by (intro measurable) (auto intro: f borel_measurable_partial_derivatives [OF S]) text\The localisation wrt S uses the same argument for many similar results.\ (*See HOL Light's MEASURABLE_ON_LEBESGUE_MEASURABLE_PREIMAGE_BOREL, etc.*) theorem borel_measurable_lebesgue_on_preimage_borel: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "S \ sets lebesgue" shows "f \ borel_measurable (lebesgue_on S) \ (\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" proof - have "{x. (if x \ S then f x else 0) \ T} \ sets lebesgue \ {x \ S. f x \ T} \ sets lebesgue" if "T \ sets borel" for T proof (cases "0 \ T") case True then have "{x \ S. f x \ T} = {x. (if x \ S then f x else 0) \ T} \ S" "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T} \ -S" by auto then show ?thesis by (metis (no_types, lifting) Compl_in_sets_lebesgue assms sets.Int sets.Un) next case False then have "{x. (if x \ S then f x else 0) \ T} = {x \ S. f x \ T}" by auto then show ?thesis by auto qed then show ?thesis unfolding borel_measurable_lebesgue_preimage_borel borel_measurable_if [OF assms, symmetric] by blast qed lemma sets_lebesgue_almost_borel: assumes "S \ sets lebesgue" obtains B N where "B \ sets borel" "negligible N" "B \ N = S" proof - obtain T N N' where "S = T \ N" "N \ N'" "N' \ null_sets lborel" "T \ sets borel" using sets_completionE [OF assms] by auto then show thesis by (metis negligible_iff_null_sets negligible_subset null_sets_completionI that) qed lemma double_lebesgue_sets: assumes S: "S \ sets lebesgue" and T: "T \ sets lebesgue" and fim: "f ` S \ T" shows "(\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue) \ f \ borel_measurable (lebesgue_on S) \ (\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue)" (is "?lhs \ _ \ ?rhs") unfolding borel_measurable_lebesgue_on_preimage_borel [OF S] proof (intro iffI allI conjI impI, safe) fix V :: "'b set" assume *: "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "V \ sets borel" then have V: "V \ sets lebesgue" by simp have "{x \ S. f x \ V} = {x \ S. f x \ T \ V}" using fim by blast also have "{x \ S. f x \ T \ V} \ sets lebesgue" using T V * le_inf_iff by blast finally show "{x \ S. f x \ V} \ sets lebesgue" . next fix U :: "'b set" assume "\U. U \ sets lebesgue \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" "negligible U" "U \ T" then show "{x \ S. f x \ U} \ sets lebesgue" using negligible_imp_sets by blast next fix U :: "'b set" assume 1 [rule_format]: "(\T. T \ sets borel \ {x \ S. f x \ T} \ sets lebesgue)" and 2 [rule_format]: "\U. negligible U \ U \ T \ {x \ S. f x \ U} \ sets lebesgue" and "U \ sets lebesgue" "U \ T" then obtain C N where C: "C \ sets borel \ negligible N \ C \ N = U" using sets_lebesgue_almost_borel by metis then have "{x \ S. f x \ C} \ sets lebesgue" by (blast intro: 1) moreover have "{x \ S. f x \ N} \ sets lebesgue" using C \U \ T\ by (blast intro: 2) moreover have "{x \ S. f x \ C \ N} = {x \ S. f x \ C} \ {x \ S. f x \ N}" by auto ultimately show "{x \ S. f x \ U} \ sets lebesgue" using C by auto qed subsection\Simplest case of Sard's theorem (we don't need continuity of derivative)\ lemma Sard_lemma00: fixes P :: "'b::euclidean_space set" assumes "a \ 0" and a: "a *\<^sub>R i \ 0" and i: "i \ Basis" and P: "P \ {x. a *\<^sub>R i \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (DIM('b) - 1)" proof - have "a > 0" using assms by simp let ?v = "(\j\Basis. (if j = i then e else m) *\<^sub>R j)" show thesis proof have "- e \ x \ i" "x \ i \ e" if "t \ P" "norm (x - t) \ e" for x t using \a > 0\ that Basis_le_norm [of i "x-t"] P i by (auto simp: inner_commute algebra_simps) moreover have "- m \ x \ j" "x \ j \ m" if "norm x \ m" "t \ P" "norm (x - t) \ e" "j \ Basis" and "j \ i" for x t j using that Basis_le_norm [of j x] by auto ultimately show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ cbox (-?v) ?v" by (auto simp: mem_box) have *: "\k\Basis. - ?v \ k \ ?v \ k" using \0 \ m\ \0 \ e\ by (auto simp: inner_Basis) have 2: "2 ^ DIM('b) = 2 * 2 ^ (DIM('b) - Suc 0)" by (metis DIM_positive Suc_pred power_Suc) show "measure lebesgue (cbox (-?v) ?v) \ 2 * e * (2 * m) ^ (DIM('b) - 1)" using \i \ Basis\ by (simp add: content_cbox [OF *] prod.distrib prod.If_cases Diff_eq [symmetric] 2) qed blast qed text\As above, but reorienting the vector (HOL Light's @text{GEOM\_BASIS\_MULTIPLE\_TAC})\ lemma Sard_lemma0: fixes P :: "(real^'n::{finite,wellorder}) set" assumes "a \ 0" and P: "P \ {x. a \ x = 0}" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain T and k::'n where T: "orthogonal_transformation T" and a: "a = T (norm a *\<^sub>R axis k (1::real))" using rotation_rightward_line by metis have Tinv [simp]: "T (inv T x) = x" for x by (simp add: T orthogonal_transformation_surj surj_f_inv_f) obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ T-`P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof (rule Sard_lemma00 [of "norm a" "axis k (1::real)" "T-`P" m e]) have "norm a *\<^sub>R axis k 1 \ x = 0" if "T x \ P" for x proof - have "a \ T x = 0" using P that by blast then show ?thesis by (metis (no_types, lifting) T a orthogonal_orthogonal_transformation orthogonal_def) qed then show "T -` P \ {x. norm a *\<^sub>R axis k 1 \ x = 0}" by auto qed (use assms T in auto) show thesis proof show "T ` S \ lmeasurable" using S measurable_orthogonal_image T by blast have "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" proof clarsimp fix x t assume "norm x \ m" "t \ P" "norm (x - t) \ e" then have "norm (inv T x) \ m" using orthogonal_transformation_inv [OF T] by (simp add: orthogonal_transformation_norm) moreover have "\t\T -` P. norm (inv T x - t) \ e" proof have "T (inv T x - inv T t) = x - t" using T linear_diff orthogonal_transformation_def by (metis (no_types, hide_lams) Tinv) then have "norm (inv T x - inv T t) = norm (x - t)" by (metis T orthogonal_transformation_norm) then show "norm (inv T x - inv T t) \ e" using \norm (x - t) \ e\ by linarith next show "inv T t \ T -` P" using \t \ P\ by force qed ultimately show "x \ T ` {z. norm z \ m \ (\t\T -` P. norm (z - t) \ e)}" by force qed then show "{z. norm z \ m \ (\t\P. norm (z - t) \ e)} \ T ` S" using image_mono [OF subS] by (rule order_trans) show "measure lebesgue (T ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" using mS T by (simp add: S measure_orthogonal_image) qed qed text\As above, but translating the sets (HOL Light's @text{GEN\_GEOM\_ORIGIN\_TAC})\ lemma Sard_lemma1: fixes P :: "(real^'n::{finite,wellorder}) set" assumes P: "dim P < CARD('n)" and "0 \ m" "0 \ e" obtains S where "S \ lmeasurable" and "{z. norm(z - w) \ m \ (\t \ P. norm(z - w - t) \ e)} \ S" and "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" proof - obtain a where "a \ 0" "P \ {x. a \ x = 0}" using lowdim_subset_hyperplane [of P] P span_base by auto then obtain S where S: "S \ lmeasurable" and subS: "{z. norm z \ m \ (\t \ P. norm(z - t) \ e)} \ S" and mS: "measure lebesgue S \ (2 * e) * (2 * m) ^ (CARD('n) - 1)" by (rule Sard_lemma0 [OF _ _ \0 \ m\ \0 \ e\]) show thesis proof show "(+)w ` S \ lmeasurable" by (metis measurable_translation S) show "{z. norm (z - w) \ m \ (\t\P. norm (z - w - t) \ e)} \ (+)w ` S" using subS by force show "measure lebesgue ((+)w ` S) \ 2 * e * (2 * m) ^ (CARD('n) - 1)" by (metis measure_translation mS) qed qed lemma Sard_lemma2: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" (is "?m \ ?n") and "B > 0" "bounded S" and derS: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" and B: "\x. x \ S \ onorm(f' x) \ B" shows "negligible(f ` S)" proof - have lin_f': "\x. x \ S \ linear(f' x)" using derS has_derivative_linear by blast show ?thesis proof (clarsimp simp add: negligible_outer_le) fix e :: "real" assume "e > 0" obtain c where csub: "S \ cbox (- (vec c)) (vec c)" and "c > 0" proof - obtain b where b: "\x. x \ S \ norm x \ b" using \bounded S\ by (auto simp: bounded_iff) show thesis proof have "- \b\ - 1 \ x $ i \ x $ i \ \b\ + 1" if "x \ S" for x i using component_le_norm_cart [of x i] b [OF that] by auto then show "S \ cbox (- vec (\b\ + 1)) (vec (\b\ + 1))" by (auto simp: mem_box_cart) qed auto qed then have box_cc: "box (- (vec c)) (vec c) \ {}" and cbox_cc: "cbox (- (vec c)) (vec c) \ {}" by (auto simp: interval_eq_empty_cart) obtain d where "d > 0" "d \ B" and d: "(d * 2) * (4 * B) ^ (?n - 1) \ e / (2*c) ^ ?m / ?m ^ ?m" apply (rule that [of "min B (e / (2*c) ^ ?m / ?m ^ ?m / (4 * B) ^ (?n - 1) / 2)"]) using \B > 0\ \c > 0\ \e > 0\ by (simp_all add: divide_simps min_mult_distrib_right) have "\r. 0 < r \ r \ 1/2 \ (x \ S \ (\y. y \ S \ norm(y - x) < r \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)))" for x proof (cases "x \ S") case True then obtain r where "r > 0" and "\y. \y \ S; norm (y - x) < r\ \ norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using derS \d > 0\ by (force simp: has_derivative_within_alt) then show ?thesis by (rule_tac x="min r (1/2)" in exI) simp next case False then show ?thesis by (rule_tac x="1/2" in exI) simp qed then obtain r where r12: "\x. 0 < r x \ r x \ 1/2" and r: "\x y. \x \ S; y \ S; norm(y - x) < r x\ \ norm(f y - f x - f' x (y - x)) \ d * norm(y - x)" by metis then have ga: "gauge (\x. ball x (r x))" by (auto simp: gauge_def) obtain \ where \: "countable \" and sub_cc: "\\ \ cbox (- vec c) (vec c)" and cbox: "\K. K \ \ \ interior K \ {} \ (\u v. K = cbox u v)" and djointish: "pairwise (\A B. interior A \ interior B = {}) \" and covered: "\K. K \ \ \ \x \ S \ K. K \ ball x (r x)" and close: "\u v. cbox u v \ \ \ \n. \i::'m. v $ i - u $ i = 2*c / 2^n" and covers: "S \ \\" apply (rule covering_lemma [OF csub box_cc ga]) apply (auto simp: Basis_vec_def cart_eq_inner_axis [symmetric]) done let ?\ = "measure lebesgue" have "\T. T \ lmeasurable \ f ` (K \ S) \ T \ ?\ T \ e / (2*c) ^ ?m * ?\ K" if "K \ \" for K proof - obtain u v where uv: "K = cbox u v" using cbox \K \ \\ by blast then have uv_ne: "cbox u v \ {}" using cbox that by fastforce obtain x where x: "x \ S \ cbox u v" "cbox u v \ ball x (r x)" using \K \ \\ covered uv by blast then have "dim (range (f' x)) < ?n" using rank_dim_range [of "matrix (f' x)"] x rank[of x] by (auto simp: matrix_works scalar_mult_eq_scaleR lin_f') then obtain T where T: "T \ lmeasurable" and subT: "{z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))} \ T" and measT: "?\ T \ (2 * (d * norm(v - u))) * (2 * ((2 * B) * norm(v - u))) ^ (?n - 1)" (is "_ \ ?DVU") apply (rule Sard_lemma1 [of "range (f' x)" "(2 * B) * norm(v - u)" "d * norm(v - u)" "f x"]) using \B > 0\ \d > 0\ by simp_all show ?thesis proof (intro exI conjI) have "f ` (K \ S) \ {z. norm(z - f x) \ (2 * B) * norm(v - u) \ (\t \ range (f' x). norm(z - f x - t) \ d * norm(v - u))}" unfolding uv proof (clarsimp simp: mult.assoc, intro conjI) fix y assume y: "y \ cbox u v" and "y \ S" then have "norm (y - x) < r x" by (metis dist_norm mem_ball norm_minus_commute subsetCE x(2)) then have le_dyx: "norm (f y - f x - f' x (y - x)) \ d * norm (y - x)" using r [of x y] x \y \ S\ by blast have yx_le: "norm (y - x) \ norm (v - u)" proof (rule norm_le_componentwise_cart) show "norm ((y - x) $ i) \ norm ((v - u) $ i)" for i using x y by (force simp: mem_box_cart dest!: spec [where x=i]) qed have *: "\norm(y - x - z) \ d; norm z \ B; d \ B\ \ norm(y - x) \ 2 * B" for x y z :: "real^'n::_" and d B using norm_triangle_ineq2 [of "y - x" z] by auto show "norm (f y - f x) \ 2 * (B * norm (v - u))" proof (rule * [OF le_dyx]) have "norm (f' x (y - x)) \ onorm (f' x) * norm (y - x)" using onorm [of "f' x" "y-x"] by (meson IntE lin_f' linear_linear x(1)) also have "\ \ B * norm (v - u)" proof (rule mult_mono) show "onorm (f' x) \ B" using B x by blast qed (use \B > 0\ yx_le in auto) finally show "norm (f' x (y - x)) \ B * norm (v - u)" . show "d * norm (y - x) \ B * norm (v - u)" using \B > 0\ by (auto intro: mult_mono [OF \d \ B\ yx_le]) qed show "\t. norm (f y - f x - f' x t) \ d * norm (v - u)" apply (rule_tac x="y-x" in exI) using \d > 0\ yx_le le_dyx mult_left_mono [where c=d] by (meson order_trans real_mult_le_cancel_iff2) qed with subT show "f ` (K \ S) \ T" by blast show "?\ T \ e / (2*c) ^ ?m * ?\ K" proof (rule order_trans [OF measT]) have "?DVU = (d * 2 * (4 * B) ^ (?n - 1)) * norm (v - u)^?n" using \c > 0\ apply (simp add: algebra_simps power_mult_distrib) by (metis Suc_pred power_Suc zero_less_card_finite) also have "\ \ (e / (2*c) ^ ?m / (?m ^ ?m)) * norm(v - u) ^ ?n" by (rule mult_right_mono [OF d]) auto also have "\ \ e / (2*c) ^ ?m * ?\ K" proof - have "u \ ball (x) (r x)" "v \ ball x (r x)" using box_ne_empty(1) contra_subsetD [OF x(2)] mem_box(2) uv_ne by fastforce+ moreover have "r x \ 1/2" using r12 by auto ultimately have "norm (v - u) \ 1" using norm_triangle_half_r [of x u 1 v] by (metis (no_types, hide_lams) dist_commute dist_norm less_eq_real_def less_le_trans mem_ball) then have "norm (v - u) ^ ?n \ norm (v - u) ^ ?m" by (simp add: power_decreasing [OF mlen]) also have "\ \ ?\ K * real (?m ^ ?m)" proof - obtain n where n: "\i. v$i - u$i = 2 * c / 2^n" using close [of u v] \K \ \\ uv by blast have "norm (v - u) ^ ?m \ (\i\UNIV. \(v - u) $ i\) ^ ?m" by (intro norm_le_l1_cart power_mono) auto also have "\ \ (\i\UNIV. v $ i - u $ i) * real CARD('m) ^ CARD('m)" by (simp add: n field_simps \c > 0\ less_eq_real_def) also have "\ = ?\ K * real (?m ^ ?m)" by (simp add: uv uv_ne content_cbox_cart) finally show ?thesis . qed finally have *: "1 / real (?m ^ ?m) * norm (v - u) ^ ?n \ ?\ K" by (simp add: field_split_simps) show ?thesis using mult_left_mono [OF *, of "e / (2*c) ^ ?m"] \c > 0\ \e > 0\ by auto qed finally show "?DVU \ e / (2*c) ^ ?m * ?\ K" . qed qed (use T in auto) qed then obtain g where meas_g: "\K. K \ \ \ g K \ lmeasurable" and sub_g: "\K. K \ \ \ f ` (K \ S) \ g K" and le_g: "\K. K \ \ \ ?\ (g K) \ e / (2*c)^?m * ?\ K" by metis have le_e: "?\ (\i\\. g i) \ e" if "\ \ \" "finite \" for \ proof - have "?\ (\i\\. g i) \ (\i\\. ?\ (g i))" using meas_g \\ \ \\ by (auto intro: measure_UNION_le [OF \finite \\]) also have "\ \ (\K\\. e / (2*c) ^ ?m * ?\ K)" using \\ \ \\ sum_mono [OF le_g] by (meson le_g subsetCE sum_mono) also have "\ = e / (2*c) ^ ?m * (\K\\. ?\ K)" by (simp add: sum_distrib_left) also have "\ \ e" proof - have "\ division_of \\" proof (rule division_ofI) show "K \ \\" "K \ {}" "\a b. K = cbox a b" if "K \ \" for K using \K \ \\ covered cbox \\ \ \\ by (auto simp: Union_upper) show "interior K \ interior L = {}" if "K \ \" and "L \ \" and "K \ L" for K L by (metis (mono_tags, lifting) \\ \ \\ pairwiseD djointish pairwise_subset that) qed (use that in auto) then have "sum ?\ \ \ ?\ (\\)" by (simp add: content_division) also have "\ \ ?\ (cbox (- vec c) (vec c) :: (real, 'm) vec set)" proof (rule measure_mono_fmeasurable) show "\\ \ cbox (- vec c) (vec c)" by (meson Sup_subset_mono sub_cc order_trans \\ \ \\) qed (use \\ division_of \\\ lmeasurable_division in auto) also have "\ = content (cbox (- vec c) (vec c) :: (real, 'm) vec set)" by simp also have "\ \ (2 ^ ?m * c ^ ?m)" using \c > 0\ by (simp add: content_cbox_if_cart) finally have "sum ?\ \ \ (2 ^ ?m * c ^ ?m)" . then show ?thesis using \e > 0\ \c > 0\ by (auto simp: field_split_simps mult_less_0_iff) qed finally show ?thesis . qed show "\T. f ` S \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "f ` S \ \ (g ` \)" using covers sub_g by force show "\ (g ` \) \ lmeasurable" by (rule fmeasurable_UN_bound [OF \countable \\ meas_g le_e]) show "?\ (\ (g ` \)) \ e" by (rule measure_UN_bound [OF \countable \\ meas_g le_e]) qed qed qed theorem baby_Sard: fixes f :: "real^'m::{finite,wellorder} \ real^'n::{finite,wellorder}" assumes mlen: "CARD('m) \ CARD('n)" and der: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and rank: "\x. x \ S \ rank(matrix(f' x)) < CARD('n)" shows "negligible(f ` S)" proof - let ?U = "\n. {x \ S. norm(x) \ n \ onorm(f' x) \ real n}" have "\x. x \ S \ \n. norm x \ real n \ onorm (f' x) \ real n" by (meson linear order_trans real_arch_simple) then have eq: "S = (\n. ?U n)" by auto have "negligible (f ` ?U n)" for n proof (rule Sard_lemma2 [OF mlen]) show "0 < real n + 1" by auto show "bounded (?U n)" using bounded_iff by blast show "(f has_derivative f' x) (at x within ?U n)" if "x \ ?U n" for x using der that by (force intro: has_derivative_subset) qed (use rank in auto) then show ?thesis by (subst eq) (simp add: image_Union negligible_Union_nat) qed subsection\A one-way version of change-of-variables not assuming injectivity. \ lemma integral_on_image_ubound_weak: fixes f :: "real^'n::{finite,wellorder} \ real" assumes S: "S \ sets lebesgue" and f: "f \ borel_measurable (lebesgue_on (g ` S))" and nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and det_int_fg: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" and meas_gim: "\T. \T \ g ` S; T \ sets lebesgue\ \ {x \ S. g x \ T} \ sets lebesgue" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. \det (matrix (g' x))\" have cont_g: "continuous_on S g" using der_g has_derivative_continuous_on by blast have [simp]: "space (lebesgue_on S) = S" by (simp add: S) have gS_in_sets_leb: "g ` S \ sets lebesgue" apply (rule differentiable_image_in_sets_lebesgue) using der_g by (auto simp: S differentiable_def differentiable_on_def) obtain h where nonneg_h: "\n x. 0 \ h n x" and h_le_f: "\n x. x \ S \ h n (g x) \ f (g x)" and h_inc: "\n x. h n x \ h (Suc n) x" and h_meas: "\n. h n \ borel_measurable lebesgue" and fin_R: "\n. finite(range (h n))" and lim: "\x. x \ g ` S \ (\n. h n x) \ f x" proof - let ?f = "\x. if x \ g ` S then f x else 0" have "?f \ borel_measurable lebesgue \ (\x. 0 \ ?f x)" by (auto simp: gS_in_sets_leb f nonneg_fg measurable_restrict_space_iff [symmetric]) then show ?thesis apply (clarsimp simp add: borel_measurable_simple_function_limit_increasing) apply (rename_tac h) by (rule_tac h=h in that) (auto split: if_split_asm) qed have h_lmeas: "{t. h n (g t) = y} \ S \ sets lebesgue" for y n proof - have "space (lebesgue_on (UNIV::(real,'n) vec set)) = UNIV" by simp then have "((h n) -`{y} \ g ` S) \ sets (lebesgue_on (g ` S))" by (metis Int_commute borel_measurable_vimage h_meas image_eqI inf_top.right_neutral sets_restrict_space space_borel space_completion space_lborel) then have "({u. h n u = y} \ g ` S) \ sets lebesgue" using gS_in_sets_leb by (simp add: integral_indicator fmeasurableI2 sets_restrict_space_iff vimage_def) then have "{x \ S. g x \ ({u. h n u = y} \ g ` S)} \ sets lebesgue" using meas_gim[of "({u. h n u = y} \ g ` S)"] by force moreover have "{t. h n (g t) = y} \ S = {x \ S. g x \ ({u. h n u = y} \ g ` S)}" by blast ultimately show ?thesis by auto qed have hint: "h n integrable_on g ` S \ integral (g ` S) (h n) \ integral S (\x. ?D x * h n (g x))" (is "?INT \ ?lhs \ ?rhs") for n proof - let ?R = "range (h n)" have hn_eq: "h n = (\x. \y\?R. y * indicat_real {x. h n x = y} x)" by (simp add: indicator_def if_distrib fin_R cong: if_cong) have yind: "(\t. y * indicator{x. h n x = y} t) integrable_on (g ` S) \ (integral (g ` S) (\t. y * indicator {x. h n x = y} t)) \ integral S (\t. \det (matrix (g' t))\ * y * indicator {x. h n x = y} (g t))" if y: "y \ ?R" for y::real proof (cases "y=0") case True then show ?thesis using gS_in_sets_leb integrable_0 by force next case False with that have "y > 0" using less_eq_real_def nonneg_h by fastforce have "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. ?D x) \ borel_measurable (lebesgue_on ({t. h n (g t) = y} \ S))" apply (intro borel_measurable_abs borel_measurable_det_Jacobian [OF h_lmeas, where f=g]) by (meson der_g IntD2 has_derivative_within_subset inf_le2) then have "(\x. if x \ {t. h n (g t) = y} \ S then ?D x else 0) \ borel_measurable lebesgue" by (rule borel_measurable_if_I [OF _ h_lmeas]) then show "(\x. if x \ {t. h n (g t) = y} then ?D x else 0) \ borel_measurable (lebesgue_on S)" by (simp add: if_if_eq_conj Int_commute borel_measurable_if [OF S, symmetric]) show "(\x. ?D x *\<^sub>R f (g x) /\<^sub>R y) integrable_on S" by (rule integrable_cmul) (use det_int_fg in auto) show "norm (if x \ {t. h n (g t) = y} then ?D x else 0) \ ?D x *\<^sub>R f (g x) /\<^sub>R y" if "x \ S" for x using nonneg_h [of n x] \y > 0\ nonneg_fg [of x] h_le_f [of x n] that by (auto simp: divide_simps mult_left_mono) qed (use S in auto) then have int_det: "(\t. \det (matrix (g' t))\) integrable_on ({t. h n (g t) = y} \ S)" using integrable_restrict_Int by force have "(g ` ({t. h n (g t) = y} \ S)) \ lmeasurable" apply (rule measurable_differentiable_image [OF h_lmeas]) apply (blast intro: has_derivative_within_subset [OF der_g]) apply (rule int_det) done moreover have "g ` ({t. h n (g t) = y} \ S) = {x. h n x = y} \ g ` S" by blast moreover have "measure lebesgue (g ` ({t. h n (g t) = y} \ S)) \ integral ({t. h n (g t) = y} \ S) (\t. \det (matrix (g' t))\)" apply (rule measure_differentiable_image [OF h_lmeas _ int_det]) apply (blast intro: has_derivative_within_subset [OF der_g]) done ultimately show ?thesis using \y > 0\ integral_restrict_Int [of S "{t. h n (g t) = y}" "\t. \det (matrix (g' t))\ * y"] apply (simp add: integrable_on_indicator integrable_on_cmult_iff integral_indicator) apply (simp add: indicator_def if_distrib cong: if_cong) done qed have hn_int: "h n integrable_on g ` S" apply (subst hn_eq) using yind by (force intro: integrable_sum [OF fin_R]) then show ?thesis proof have "?lhs = integral (g ` S) (\x. \y\range (h n). y * indicat_real {x. h n x = y} x)" by (metis hn_eq) also have "\ = (\y\range (h n). integral (g ` S) (\x. y * indicat_real {x. h n x = y} x))" by (rule integral_sum [OF fin_R]) (use yind in blast) also have "\ \ (\y\range (h n). integral S (\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)))" using yind by (force intro: sum_mono) also have "\ = integral S (\u. \y\range (h n). \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u))" proof (rule integral_sum [OF fin_R, symmetric]) fix y assume y: "y \ ?R" with nonneg_h have "y \ 0" by auto show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\x. indicat_real {x. h n x = y} (g x)) \ borel_measurable (lebesgue_on S)" using h_lmeas S by (auto simp: indicator_vimage [symmetric] borel_measurable_indicator_iff sets_restrict_space_iff) then show "(\u. \det (matrix (g' u))\ * y * indicat_real {x. h n x = y} (g u)) \ borel_measurable (lebesgue_on S)" by (intro borel_measurable_times borel_measurable_abs borel_measurable_const borel_measurable_det_Jacobian [OF S der_g]) next fix x assume "x \ S" have "y * indicat_real {x. h n x = y} (g x) \ f (g x)" by (metis (full_types) \x \ S\ h_le_f indicator_def mem_Collect_eq mult.right_neutral mult_zero_right nonneg_fg) with \y \ 0\ show "norm (?D x * y * indicat_real {x. h n x = y} (g x)) \ ?D x * f(g x)" by (simp add: abs_mult mult.assoc mult_left_mono) qed (use S det_int_fg in auto) qed also have "\ = integral S (\T. \det (matrix (g' T))\ * (\y\range (h n). y * indicat_real {x. h n x = y} (g T)))" by (simp add: sum_distrib_left mult.assoc) also have "\ = ?rhs" by (metis hn_eq) finally show "integral (g ` S) (h n) \ ?rhs" . qed qed have le: "integral S (\T. \det (matrix (g' T))\ * h n (g T)) \ ?b" for n proof (rule integral_le) show "(\T. \det (matrix (g' T))\ * h n (g T)) integrable_on S" proof (rule measurable_bounded_by_integrable_imp_integrable) have "(\T. \det (matrix (g' T))\ *\<^sub>R h n (g T)) \ borel_measurable (lebesgue_on S)" proof (intro borel_measurable_scaleR borel_measurable_abs borel_measurable_det_Jacobian \S \ sets lebesgue\) have eq: "{x \ S. f x \ a} = (\b \ (f ` S) \ atMost a. {x. f x = b} \ S)" for f and a::real by auto have "finite ((\x. h n (g x)) ` S \ {..a})" for a by (force intro: finite_subset [OF _ fin_R]) with h_lmeas [of n] show "(\x. h n (g x)) \ borel_measurable (lebesgue_on S)" apply (simp add: borel_measurable_vimage_halfspace_component_le \S \ sets lebesgue\ sets_restrict_space_iff eq) by (metis (mono_tags) SUP_inf sets.finite_UN) qed (use der_g in blast) then show "(\T. \det (matrix (g' T))\ * h n (g T)) \ borel_measurable (lebesgue_on S)" by simp show "norm (?D x * h n (g x)) \ ?D x *\<^sub>R f (g x)" if "x \ S" for x by (simp add: h_le_f mult_left_mono nonneg_h that) qed (use S det_int_fg in auto) show "?D x * h n (g x) \ ?D x * f (g x)" if "x \ S" for x by (simp add: \x \ S\ h_le_f mult_left_mono) show "(\x. ?D x * f (g x)) integrable_on S" using det_int_fg by blast qed have "f integrable_on g ` S \ (\k. integral (g ` S) (h k)) \ integral (g ` S) f" proof (rule monotone_convergence_increasing) have "\integral (g ` S) (h n)\ \ integral S (\x. ?D x * f (g x))" for n proof - have "\integral (g ` S) (h n)\ = integral (g ` S) (h n)" using hint by (simp add: integral_nonneg nonneg_h) also have "\ \ integral S (\x. ?D x * f (g x))" using hint le by (meson order_trans) finally show ?thesis . qed then show "bounded (range (\k. integral (g ` S) (h k)))" by (force simp: bounded_iff) qed (use h_inc lim hint in auto) moreover have "integral (g ` S) (h n) \ integral S (\x. ?D x * f (g x))" for n using hint by (blast intro: le order_trans) ultimately show ?thesis by (auto intro: Lim_bounded) qed lemma integral_on_image_ubound_nonneg: fixes f :: "real^'n::{finite,wellorder} \ real" assumes nonneg_fg: "\x. x \ S \ 0 \ f(g x)" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "f integrable_on (g ` S) \ integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" (is "_ \ _ \ ?b") proof - let ?D = "\x. det (matrix (g' x))" define S' where "S' \ {x \ S. ?D x * f(g x) \ 0}" then have der_gS': "\x. x \ S' \ (g has_derivative g' x) (at x within S')" by (metis (mono_tags, lifting) der_g has_derivative_within_subset mem_Collect_eq subset_iff) have "(\x. if x \ S then \?D x\ * f (g x) else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV intS) then have Df_borel: "(\x. if x \ S then \?D x\ * f (g x) else 0) \ borel_measurable lebesgue" using integrable_imp_measurable lebesgue_on_UNIV_eq by force have S': "S' \ sets lebesgue" proof - from Df_borel borel_measurable_vimage_open [of _ UNIV] have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ T} \ sets lebesgue" if "open T" for T using that unfolding lebesgue_on_UNIV_eq by (fastforce simp add: dest!: spec) then have "{x. (if x \ S then \?D x\ * f (g x) else 0) \ -{0}} \ sets lebesgue" using open_Compl by blast then show ?thesis by (simp add: S'_def conj_ac split: if_split_asm cong: conj_cong) qed then have gS': "g ` S' \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) show "g differentiable_on S'" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_within_subset) qed auto have f: "f \ borel_measurable (lebesgue_on (g ` S'))" proof (clarsimp simp add: borel_measurable_vimage_open) fix T :: "real set" assume "open T" have "{x \ g ` S'. f x \ T} = g ` {x \ S'. f(g x) \ T}" by blast moreover have "g ` {x \ S'. f(g x) \ T} \ sets lebesgue" proof (rule differentiable_image_in_sets_lebesgue) let ?h = "\x. \?D x\ * f (g x) /\<^sub>R \?D x\" have "(\x. if x \ S' then \?D x\ * f (g x) else 0) = (\x. if x \ S then \?D x\ * f (g x) else 0)" by (auto simp: S'_def) also have "\ \ borel_measurable lebesgue" by (rule Df_borel) finally have *: "(\x. \?D x\ * f (g x)) \ borel_measurable (lebesgue_on S')" by (simp add: borel_measurable_if_D) have "?h \ borel_measurable (lebesgue_on S')" by (intro * S' der_gS' borel_measurable_det_Jacobian measurable) (blast intro: der_gS') moreover have "?h x = f(g x)" if "x \ S'" for x using that by (auto simp: S'_def) ultimately have "(\x. f(g x)) \ borel_measurable (lebesgue_on S')" by (metis (no_types, lifting) measurable_lebesgue_cong) then show "{x \ S'. f (g x) \ T} \ sets lebesgue" by (simp add: \S' \ sets lebesgue\ \open T\ borel_measurable_vimage_open sets_restrict_space_iff) show "g differentiable_on {x \ S'. f (g x) \ T}" using der_g unfolding S'_def differentiable_def differentiable_on_def by (blast intro: has_derivative_within_subset) qed auto ultimately have "{x \ g ` S'. f x \ T} \ sets lebesgue" by metis then show "{x \ g ` S'. f x \ T} \ sets (lebesgue_on (g ` S'))" by (simp add: \g ` S' \ sets lebesgue\ sets_restrict_space_iff) qed have intS': "(\x. \?D x\ * f (g x)) integrable_on S'" using intS by (rule integrable_spike_set) (auto simp: S'_def intro: empty_imp_negligible) have lebS': "{x \ S'. g x \ T} \ sets lebesgue" if "T \ g ` S'" "T \ sets lebesgue" for T proof - have "g \ borel_measurable (lebesgue_on S')" using der_gS' has_derivative_continuous_on S' by (blast intro: continuous_imp_measurable_on_sets_lebesgue) moreover have "{x \ S'. g x \ U} \ sets lebesgue" if "negligible U" "U \ g ` S'" for U proof (intro negligible_imp_sets negligible_differentiable_vimage that) fix x assume x: "x \ S'" then have "linear (g' x)" using der_gS' has_derivative_linear by blast with x show "inj (g' x)" by (auto simp: S'_def det_nz_iff_inj) qed (use der_gS' in auto) ultimately show ?thesis using double_lebesgue_sets [OF S' gS' order_refl] that by blast qed have int_gS': "f integrable_on g ` S' \ integral (g ` S') f \ integral S' (\x. \?D x\ * f(g x))" using integral_on_image_ubound_weak [OF S' f nonneg_fg der_gS' intS' lebS'] S'_def by blast have "negligible (g ` {x \ S. det(matrix(g' x)) = 0})" proof (rule baby_Sard, simp_all) fix x assume x: "x \ S \ det (matrix (g' x)) = 0" then show "(g has_derivative g' x) (at x within {x \ S. det (matrix (g' x)) = 0})" by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI) then show "rank (matrix (g' x)) < CARD('n)" using det_nz_iff_inj matrix_vector_mul_linear x by (fastforce simp add: less_rank_noninjective) qed then have negg: "negligible (g ` S - g ` {x \ S. ?D x \ 0})" by (rule negligible_subset) (auto simp: S'_def) have null: "g ` {x \ S. ?D x \ 0} - g ` S = {}" by (auto simp: S'_def) let ?F = "{x \ S. f (g x) \ 0}" have eq: "g ` S' = g ` ?F \ g ` {x \ S. ?D x \ 0}" by (auto simp: S'_def image_iff) show ?thesis proof have "((\x. if x \ g ` ?F then f x else 0) integrable_on g ` {x \ S. ?D x \ 0})" using int_gS' eq integrable_restrict_Int [where f=f] by simp then have "f integrable_on g ` {x \ S. ?D x \ 0}" by (auto simp: image_iff elim!: integrable_eq) then show "f integrable_on g ` S" apply (rule integrable_spike_set [OF _ empty_imp_negligible negligible_subset]) using negg null by auto have "integral (g ` S) f = integral (g ` {x \ S. ?D x \ 0}) f" using negg by (auto intro: negligible_subset integral_spike_set) also have "\ = integral (g ` {x \ S. ?D x \ 0}) (\x. if x \ g ` ?F then f x else 0)" by (auto simp: image_iff intro!: integral_cong) also have "\ = integral (g ` S') f" using eq integral_restrict_Int by simp also have "\ \ integral S' (\x. \?D x\ * f(g x))" by (metis int_gS') also have "\ \ ?b" by (rule integral_subset_le [OF _ intS' intS]) (use nonneg_fg S'_def in auto) finally show "integral (g ` S) f \ ?b" . qed qed lemma absolutely_integrable_on_image_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" proof - let ?D = "\x. \det (matrix (g' x))\ * f (g x)" let ?N = "{x \ S. f (g x) < 0}" and ?P = "{x \ S. f (g x) > 0}" have eq: "{x. (if x \ S then ?D x else 0) > 0} = {x \ S. ?D x > 0}" "{x. (if x \ S then ?D x else 0) < 0} = {x \ S. ?D x < 0}" by auto have "?D integrable_on S" using intS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?D x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?D x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have Dlt: "{x \ S. ?D x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have Dgt: "{x \ S. ?D x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have dfgbm: "?D \ borel_measurable (lebesgue_on S)" using intS absolutely_integrable_on_def integrable_imp_measurable by blast have der_gN: "(g has_derivative g' x) (at x within ?N)" if "x \ ?N" for x using der_g has_derivative_within_subset that by force have "(\x. - f x) integrable_on g ` ?N \ integral (g ` ?N) (\x. - f x) \ integral ?N (\x. \det (matrix (g' x))\ * - f (g x))" proof (rule integral_on_image_ubound_nonneg [OF _ der_gN]) have 1: "?D integrable_on {x \ S. ?D x < 0}" using Dlt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) have "uminus \ (\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: o_def mult_less_0_iff empty_imp_negligible integrable_spike_set [OF 1]) then show "(\x. \det (matrix (g' x))\ * - f (g x)) integrable_on ?N" by (simp add: integrable_neg_iff o_def) qed auto then have "f integrable_on g ` ?N" by (simp add: integrable_neg_iff) moreover have "g ` ?N = {y \ g ` S. f y < 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y < 0}" by simp then have N: "f absolutely_integrable_on {y \ g ` S. f y < 0}" by (rule absolutely_integrable_absolutely_integrable_ubound) auto have der_gP: "(g has_derivative g' x) (at x within ?P)" if "x \ ?P" for x using der_g has_derivative_within_subset that by force have "f integrable_on g ` ?P \ integral (g ` ?P) f \ integral ?P ?D" proof (rule integral_on_image_ubound_nonneg [OF _ der_gP]) have "?D integrable_on {x \ S. 0 < ?D x}" using Dgt by (auto intro: set_lebesgue_integral_eq_integral [OF set_integrable_subset] intS) then show "?D integrable_on ?P" apply (rule integrable_spike_set) by (auto simp: zero_less_mult_iff empty_imp_negligible) qed auto then have "f integrable_on g ` ?P" by metis moreover have "g ` ?P = {y \ g ` S. f y > 0}" by auto ultimately have "f integrable_on {y \ g ` S. f y > 0}" by simp then have P: "f absolutely_integrable_on {y \ g ` S. f y > 0}" by (rule absolutely_integrable_absolutely_integrable_lbound) auto have "(\x. if x \ g ` S \ f x < 0 \ x \ g ` S \ 0 < f x then f x else 0) = (\x. if x \ g ` S then f x else 0)" by auto then show ?thesis using absolutely_integrable_Un [OF N P] absolutely_integrable_restrict_UNIV [symmetric, where f=f] by simp qed proposition absolutely_integrable_on_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and intS: "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" shows "f absolutely_integrable_on (g ` S)" apply (rule absolutely_integrable_componentwise [OF absolutely_integrable_on_image_real [OF der_g]]) using absolutely_integrable_component [OF intS] by auto proposition integral_on_image_ubound: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes"\x. x \ S \ 0 \ f(g x)" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "(\x. \det (matrix (g' x))\ * f(g x)) integrable_on S" shows "integral (g ` S) f \ integral S (\x. \det (matrix (g' x))\ * f(g x))" using integral_on_image_ubound_nonneg [OF assms] by simp subsection\Change-of-variables theorem\ text\The classic change-of-variables theorem. We have two versions with quite general hypotheses, the first that the transforming function has a continuous inverse, the second that the base set is Lebesgue measurable.\ lemma cov_invertible_nonneg_le: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and f0: "\y. y \ T \ 0 \ f y" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "f integrable_on T \ (integral T f) \ b \ (\x. \det (matrix (g' x))\ * f(g x)) integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) \ b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ have gS: "g differentiable_on S" by (meson der_g differentiable_def differentiable_on_def) let ?D = "\x. \det (matrix (g' x))\ * f (g x)" show ?thesis proof assume ?lhs then have fT: "f integrable_on T" and intf: "integral T f \ b" by blast+ show ?rhs proof let ?fgh = "\x. \det (matrix (h' x))\ * (\det (matrix (g' (h x)))\ * f (g (h x)))" have ddf: "?fgh x = f x" if "x \ T" for x proof - have "matrix (h' x) ** matrix (g' (h x)) = mat 1" using that id[OF that] der_g[of "h x"] gh[OF that] left_inverse_linear has_derivative_linear by (subst matrix_compose[symmetric]) (force simp: matrix_id_mat_1 has_derivative_linear)+ then have "\det (matrix (h' x))\ * \det (matrix (g' (h x)))\ = 1" by (metis abs_1 abs_mult det_I det_mul) then show ?thesis by (simp add: gh that) qed have "?D integrable_on (h ` T)" proof (intro set_lebesgue_integral_eq_integral absolutely_integrable_on_image_real) show "(\x. ?fgh x) absolutely_integrable_on T" proof (subst absolutely_integrable_on_iff_nonneg) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (simp add: zero_le_mult_iff f0 gh) qed (use der_h in auto) with Seq show "(\x. ?D x) integrable_on S" by simp have "integral S (\x. ?D x) \ integral T (\x. ?fgh x)" unfolding Seq proof (rule integral_on_image_ubound) show "(\x. ?fgh x) integrable_on T" using ddf fT integrable_eq by force qed (use f0 gh der_h in auto) also have "\ = integral T f" by (force simp: ddf intro: integral_cong) also have "\ \ b" by (rule intf) finally show "integral S (\x. ?D x) \ b" . qed next assume R: ?rhs then have "f integrable_on g ` S" using der_g f0 hg integral_on_image_ubound_nonneg by blast moreover have "integral (g ` S) f \ integral S (\x. ?D x)" by (rule integral_on_image_ubound [OF f0 der_g]) (use R Teq in auto) ultimately show ?lhs using R by (simp add: Teq) qed qed lemma cov_invertible_nonneg_eq: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "\y. y \ T \ (h has_derivative h' y) (at y within T)" and "\y. y \ T \ 0 \ f y" and "\x. x \ S \ g x \ T \ h(g x) = x" and "\y. y \ T \ h y \ S \ g(h y) = y" and "\y. y \ T \ h' y \ g'(h y) = id" shows "((\x. \det (matrix (g' x))\ * f(g x)) has_integral b) S \ (f has_integral b) T" using cov_invertible_nonneg_le [OF assms] by (simp add: has_integral_iff) (meson eq_iff) lemma cov_invertible_real: fixes f :: "real^'n::{finite,wellorder} \ real" and g :: "real^'n::_ \ real^'n::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ * f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" (is "?lhs = ?rhs") proof - have Teq: "T = g`S" and Seq: "S = h`T" using hg gh image_iff by fastforce+ let ?DP = "\x. \det (matrix (g' x))\ * f(g x)" and ?DN = "\x. \det (matrix (g' x))\ * -f(g x)" have "+": "(?DP has_integral b) {x \ S. f (g x) > 0} \ (f has_integral b) {y \ T. f y > 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. f (g x)) -` Y \ {x \ S. f (g x) > 0} = ((\x. f (g x)) -` Y \ S) \ {x \ S. f (g x) > 0}" for Y by auto show "(g has_derivative g' x) (at x within {x \ S. f (g x) > 0})" if "x \ {x \ S. f (g x) > 0}" for x using that der_g has_derivative_within_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y > 0})" if "y \ {y \ T. f y > 0}" for y using that der_h has_derivative_within_subset by fastforce qed (use gh hg id in auto) have "-": "(?DN has_integral b) {x \ S. f (g x) < 0} \ ((\x. - f x) has_integral b) {y \ T. f y < 0}" for b proof (rule cov_invertible_nonneg_eq) have *: "(\x. - f (g x)) -` y \ {x \ S. f (g x) < 0} = ((\x. f (g x)) -` uminus ` y \ S) \ {x \ S. f (g x) < 0}" for y using image_iff by fastforce show "(g has_derivative g' x) (at x within {x \ S. f (g x) < 0})" if "x \ {x \ S. f (g x) < 0}" for x using that der_g has_derivative_within_subset by fastforce show "(h has_derivative h' y) (at y within {y \ T. f y < 0})" if "y \ {y \ T. f y < 0}" for y using that der_h has_derivative_within_subset by fastforce qed (use gh hg id in auto) show ?thesis proof assume LHS: ?lhs have eq: "{x. (if x \ S then ?DP x else 0) > 0} = {x \ S. ?DP x > 0}" "{x. (if x \ S then ?DP x else 0) < 0} = {x \ S. ?DP x < 0}" by auto have "?DP integrable_on S" using LHS absolutely_integrable_on_def by blast then have "(\x. if x \ S then ?DP x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ S then ?DP x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have SN: "{x \ S. ?DP x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have SP: "{x \ S. ?DP x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have "?DP absolutely_integrable_on {x \ S. ?DP x > 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SP) then have aP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" by (rule absolutely_integrable_spike_set) (auto simp: zero_less_mult_iff empty_imp_negligible) have "?DP absolutely_integrable_on {x \ S. ?DP x < 0}" using LHS by (fast intro!: set_integrable_subset [OF _, of _ S] SN) then have aN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" by (rule absolutely_integrable_spike_set) (auto simp: mult_less_0_iff empty_imp_negligible) have fN: "f integrable_on {y \ T. f y < 0}" "integral {y \ T. f y < 0} f = integral {x \ S. f (g x) < 0} ?DP" using "-" [of "integral {x \ S. f(g x) < 0} ?DN"] aN by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faN: "f absolutely_integrable_on {y \ T. f y < 0}" apply (rule absolutely_integrable_integrable_bound [where g = "\x. - f x"]) using fN by (auto simp: integrable_neg_iff) have fP: "f integrable_on {y \ T. f y > 0}" "integral {y \ T. f y > 0} f = integral {x \ S. f (g x) > 0} ?DP" using "+" [of "integral {x \ S. f(g x) > 0} ?DP"] aP by (auto simp: set_lebesgue_integral_eq_integral has_integral_iff integrable_neg_iff) have faP: "f absolutely_integrable_on {y \ T. f y > 0}" apply (rule absolutely_integrable_integrable_bound [where g = f]) using fP by auto have fa: "f absolutely_integrable_on ({y \ T. f y < 0} \ {y \ T. f y > 0})" by (rule absolutely_integrable_Un [OF faN faP]) show ?rhs proof have eq: "((if x \ T \ f x < 0 \ x \ T \ 0 < f x then 1 else 0) * f x) = (if x \ T then 1 else 0) * f x" for x by auto show "f absolutely_integrable_on T" using fa by (simp add: indicator_def set_integrable_def eq) have [simp]: "{y \ T. f y < 0} \ {y \ T. 0 < f y} = {}" for T and f :: "(real^'n::_) \ real" by auto have "integral T f = integral ({y \ T. f y < 0} \ {y \ T. f y > 0}) f" by (intro empty_imp_negligible integral_spike_set) (auto simp: eq) also have "\ = integral {y \ T. f y < 0} f + integral {y \ T. f y > 0} f" using fN fP by simp also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" by (simp add: fN fP) also have "\ = integral ({x \ S. f (g x) < 0} \ {x \ S. 0 < f (g x)}) ?DP" using aP aN by (simp add: set_lebesgue_integral_eq_integral) also have "\ = integral S ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using LHS by simp finally show "integral T f = b" . qed next assume RHS: ?rhs have eq: "{x. (if x \ T then f x else 0) > 0} = {x \ T. f x > 0}" "{x. (if x \ T then f x else 0) < 0} = {x \ T. f x < 0}" by auto have "f integrable_on T" using RHS absolutely_integrable_on_def by blast then have "(\x. if x \ T then f x else 0) integrable_on UNIV" by (simp add: integrable_restrict_UNIV) then have D_borel: "(\x. if x \ T then f x else 0) \ borel_measurable (lebesgue_on UNIV)" using integrable_imp_measurable lebesgue_on_UNIV_eq by blast then have TN: "{x \ T. f x < 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_lt by (drule_tac x=0 in spec) (auto simp: eq) from D_borel have TP: "{x \ T. f x > 0} \ sets lebesgue" unfolding borel_measurable_vimage_halfspace_component_gt by (drule_tac x=0 in spec) (auto simp: eq) have aint: "f absolutely_integrable_on {y. y \ T \ 0 < (f y)}" "f absolutely_integrable_on {y. y \ T \ (f y) < 0}" and intT: "integral T f = b" using set_integrable_subset [of _ T] TP TN RHS by blast+ show ?lhs proof have fN: "f integrable_on {v \ T. f v < 0}" using absolutely_integrable_on_def aint by blast then have DN: "(?DN has_integral integral {y \ T. f y < 0} (\x. - f x)) {x \ S. f (g x) < 0}" using "-" [of "integral {y \ T. f y < 0} (\x. - f x)"] by (simp add: has_integral_neg_iff integrable_integral) have aDN: "?DP absolutely_integrable_on {x \ S. f (g x) < 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DN]) using DN hg by (fastforce simp: abs_mult integrable_neg_iff)+ have fP: "f integrable_on {v \ T. f v > 0}" using absolutely_integrable_on_def aint by blast then have DP: "(?DP has_integral integral {y \ T. f y > 0} f) {x \ S. f (g x) > 0}" using "+" [of "integral {y \ T. f y > 0} f"] by (simp add: has_integral_neg_iff integrable_integral) have aDP: "?DP absolutely_integrable_on {x \ S. f (g x) > 0}" apply (rule absolutely_integrable_integrable_bound [where g = ?DP]) using DP hg by (fastforce simp: integrable_neg_iff)+ have eq: "(if x \ S then 1 else 0) * ?DP x = (if x \ S \ f (g x) < 0 \ x \ S \ f (g x) > 0 then 1 else 0) * ?DP x" for x by force have "?DP absolutely_integrable_on ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0})" by (rule absolutely_integrable_Un [OF aDN aDP]) then show I: "?DP absolutely_integrable_on S" by (simp add: indicator_def eq set_integrable_def) have [simp]: "{y \ S. f y < 0} \ {y \ S. 0 < f y} = {}" for S and f :: "(real^'n::_) \ real" by auto have "integral S ?DP = integral ({x \ S. f (g x) < 0} \ {x \ S. f (g x) > 0}) ?DP" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = integral {x \ S. f (g x) < 0} ?DP + integral {x \ S. 0 < f (g x)} ?DP" using aDN aDP by (simp add: set_lebesgue_integral_eq_integral) also have "\ = - integral {y \ T. f y < 0} (\x. - f x) + integral {y \ T. f y > 0} f" using DN DP by (auto simp: has_integral_iff) also have "\ = integral ({x \ T. f x < 0} \ {x \ T. 0 < f x}) f" by (simp add: fN fP) also have "\ = integral T f" by (intro empty_imp_negligible integral_spike_set) auto also have "\ = b" using intT by simp finally show "integral S ?DP = b" . qed qed qed lemma cv_inv_version3: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and der_h: "\y. y \ T \ (h has_derivative h' y) (at y within T)" and hg: "\x. x \ S \ g x \ T \ h(g x) = x" and gh: "\y. y \ T \ h y \ S \ g(h y) = y" and id: "\y. y \ T \ h' y \ g'(h y) = id" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on T \ integral T f = b" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have "((\x. \det (matrix (g' x))\ * f(g x) $ i) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ * (f(g x) $ i)) = b $ i) \ ((\x. f x $ i) absolutely_integrable_on T \ integral T (\x. f x $ i) = b $ i)" for i by (rule cov_invertible_real [OF der_g der_h hg gh id]) then have "?D absolutely_integrable_on S \ (?D has_integral b) S \ f absolutely_integrable_on T \ (f has_integral b) T" unfolding absolutely_integrable_componentwise_iff [where f=f] has_integral_componentwise_iff [of f] absolutely_integrable_componentwise_iff [where f="?D"] has_integral_componentwise_iff [of ?D] by (auto simp: all_conj_distrib Basis_vec_def cart_eq_inner_axis [symmetric] has_integral_iff set_lebesgue_integral_eq_integral) then show ?thesis using absolutely_integrable_on_def by blast qed lemma cv_inv_version4: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S) \ invertible(matrix(g' x))" and hg: "\x. x \ S \ continuous_on (g ` S) h \ h(g x) = x" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - have "\x. \h'. x \ S \ (g has_derivative g' x) (at x within S) \ linear h' \ g' x \ h' = id \ h' \ g' x = id" using der_g matrix_invertible has_derivative_linear by blast then obtain h' where h': "\x. x \ S \ (g has_derivative g' x) (at x within S) \ linear (h' x) \ g' x \ (h' x) = id \ (h' x) \ g' x = id" by metis show ?thesis proof (rule cv_inv_version3) show "\y. y \ g ` S \ (h has_derivative h' (h y)) (at y within g ` S)" using h' hg by (force simp: continuous_on_eq_continuous_within intro!: has_derivative_inverse_within) qed (use h' hg in auto) qed theorem has_absolute_integral_change_of_variables_invertible: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and hg: "\x. x \ S \ h(g x) = x" and conth: "continuous_on (g ` S) h" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" (is "?lhs = ?rhs") proof - let ?S = "{x \ S. invertible (matrix (g' x))}" and ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f(g x)" have *: "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b" proof (rule cv_inv_version4) show "(g has_derivative g' x) (at x within ?S) \ invertible (matrix (g' x))" if "x \ ?S" for x using der_g that has_derivative_within_subset that by fastforce show "continuous_on (g ` ?S) h \ h (g x) = x" if "x \ ?S" for x using that continuous_on_subset [OF conth] by (simp add: hg image_mono) qed have "(g has_derivative g' x) (at x within {x \ S. rank (matrix (g' x)) < CARD('m)})" if "x \ S" for x by (metis (no_types, lifting) der_g has_derivative_within_subset mem_Collect_eq subsetI that) then have "negligible (g ` {x \ S. \ invertible (matrix (g' x))})" by (auto simp: invertible_det_nz det_eq_0_rank intro: baby_Sard) then have neg: "negligible {x \ g ` S. x \ g ` ?S \ f x \ 0}" by (auto intro: negligible_subset) have [simp]: "{x \ g ` ?S. x \ g ` S \ f x \ 0} = {}" by auto have "?D absolutely_integrable_on ?S \ integral ?S ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" apply (intro conj_cong absolutely_integrable_spike_set_eq) apply(auto simp: integral_spike_set invertible_det_nz empty_imp_negligible neg) done moreover have "f absolutely_integrable_on (g ` ?S) \ integral (g ` ?S) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" by (auto intro!: conj_cong absolutely_integrable_spike_set_eq integral_spike_set neg) ultimately show ?thesis using "*" by blast qed theorem has_absolute_integral_change_of_variables_compact: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "compact S" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b)" proof - obtain h where hg: "\x. x \ S \ h(g x) = x" using inj by (metis the_inv_into_f_f) have conth: "continuous_on (g ` S) h" by (metis \compact S\ continuous_on_inv der_g has_derivative_continuous_on hg) show ?thesis by (rule has_absolute_integral_change_of_variables_invertible [OF der_g hg conth]) qed lemma has_absolute_integral_change_of_variables_compact_family: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes compact: "\n::nat. compact (F n)" and der_g: "\x. x \ (\n. F n) \ (g has_derivative g' x) (at x within (\n. F n))" and inj: "inj_on g (\n. F n)" shows "((\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on (\n. F n) \ integral (\n. F n) (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` (\n. F n)) \ integral (g ` (\n. F n)) f = b)" proof - let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" let ?U = "\n. \m\n. F m" let ?lift = "vec::real\real^1" have F_leb: "F m \ sets lebesgue" for m by (simp add: compact borel_compact) have iff: "(\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) absolutely_integrable_on (?U n) \ integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R f (g x)) = b \ f absolutely_integrable_on (g ` (?U n)) \ integral (g ` (?U n)) f = b" for n b and f :: "real^'m::_ \ real^'k" proof (rule has_absolute_integral_change_of_variables_compact) show "compact (?U n)" by (simp add: compact compact_UN) show "(g has_derivative g' x) (at x within (?U n))" if "x \ ?U n" for x using that by (blast intro!: has_derivative_within_subset [OF der_g]) show "inj_on g (?U n)" using inj by (auto simp: inj_on_def) qed show ?thesis unfolding image_UN proof safe assume DS: "?D absolutely_integrable_on (\n. F n)" and b: "b = integral (\n. F n) ?D" have DU: "\n. ?D absolutely_integrable_on (?U n)" "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" using integral_countable_UN [OF DS F_leb] by auto with iff have fag: "f absolutely_integrable_on g ` (?U n)" and fg_int: "integral (\m\n. g ` F m) f = integral (?U n) ?D" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\m. g ` F m) then norm(f x) else 0" have "(\x. if x \ (\m. g ` F m) then f x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ (\m\k. g ` F m) then f x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using fag by (simp add: image_UN) let ?nf = "\n x. if x \ (\m\n. g ` F m) then norm(f x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nf k integrable_on UNIV" for k using fag unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n have "(norm \ ?D) absolutely_integrable_on ?U n" by (intro absolutely_integrable_norm DU) then have "integral (g ` ?U n) (norm \ f) = integral (?U n) (norm \ ?D)" using iff [of n "vec \ norm \ f" "integral (?U n) (\x. \det (matrix (g' x))\ *\<^sub>R (?lift \ norm \ f) (g x))"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq by (auto simp: o_def) } moreover have "bounded (range (\k. integral (?U k) (norm \ ?D)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (?U k) (norm \ ?D)) \ integral (\n. F n) (norm \ ?D)" unfolding integral_restrict_UNIV [of _ "norm \ ?D", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ \ (F ` {..k}) then (norm \ ?D) x else 0) integrable_on UNIV" "(\x. if x \ (\n. F n) then (norm \ ?D) x else 0) integrable_on UNIV" using DU(1) DS unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by auto qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nf k)))" by (simp add: integral_restrict_UNIV image_UN [symmetric] o_def) next show "(\k. if x \ (\m\k. g ` F m) then norm (f x) else 0) \ (if x \ (\m. g ` F m) then norm (f x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ (\m\k. g ` F m) then f x else 0) \ (if x \ (\m. g ` F m) then f x else 0)" for x proof clarsimp fix m y assume "y \ F m" show "(\k. if \x\{..k}. g y \ g ` F x then f (g y) else 0) \ f (g y)" using \y \ F m\ by (force intro: tendsto_eventually eventually_sequentiallyI [of m]) qed qed auto then show fai: "f absolutely_integrable_on (\m. g ` F m)" using absolutely_integrable_restrict_UNIV by blast show "integral ((\x. g ` F x)) f = integral (\n. F n) ?D" proof (rule LIMSEQ_unique) show "(\n. integral (?U n) ?D) \ integral (\x. g ` F x) f" unfolding fg_int [symmetric] proof (rule integral_countable_UN [OF fai]) show "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" by (meson der_g differentiableI UnionI differentiable_on_def differentiable_on_subset rangeI subsetI) qed auto qed next show "(\n. integral (?U n) ?D) \ integral (\n. F n) ?D" by (rule DU) qed next assume fs: "f absolutely_integrable_on (\x. g ` F x)" and b: "b = integral ((\x. g ` F x)) f" have gF_leb: "g ` F m \ sets lebesgue" for m proof (rule differentiable_image_in_sets_lebesgue [OF F_leb]) show "g differentiable_on F m" using der_g unfolding differentiable_def differentiable_on_def by (meson Sup_upper UNIV_I UnionI has_derivative_within_subset image_eqI) qed auto have fgU: "\n. f absolutely_integrable_on (\m\n. g ` F m)" "(\n. integral (\m\n. g ` F m) f) \ integral (\m. g ` F m) f" using integral_countable_UN [OF fs gF_leb] by auto with iff have DUn: "?D absolutely_integrable_on ?U n" and D_int: "integral (?U n) ?D = integral (\m\n. g ` F m) f" for n by (auto simp: image_UN) let ?h = "\x. if x \ (\n. F n) then norm(?D x) else 0" have "(\x. if x \ (\n. F n) then ?D x else 0) absolutely_integrable_on UNIV" proof (rule dominated_convergence_absolutely_integrable) show "(\x. if x \ ?U k then ?D x else 0) absolutely_integrable_on UNIV" for k unfolding absolutely_integrable_restrict_UNIV using DUn by simp let ?nD = "\n x. if x \ ?U n then norm(?D x) else 0" show "?h integrable_on UNIV" proof (rule monotone_convergence_increasing [THEN conjunct1]) show "?nD k integrable_on UNIV" for k using DUn unfolding integrable_restrict_UNIV absolutely_integrable_on_def by (simp add: image_UN) { fix n::nat have "(norm \ f) absolutely_integrable_on (\m\n. g ` F m)" apply (rule absolutely_integrable_norm) using fgU by blast then have "integral (?U n) (norm \ ?D) = integral (g ` ?U n) (norm \ f)" using iff [of n "?lift \ norm \ f" "integral (g ` ?U n) (?lift \ norm \ f)"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq image_UN by (auto simp: o_def) } moreover have "bounded (range (\k. integral (g ` ?U k) (norm \ f)))" unfolding bounded_iff proof (rule exI, clarify) fix k show "norm (integral (g ` ?U k) (norm \ f)) \ integral (g ` (\n. F n)) (norm \ f)" unfolding integral_restrict_UNIV [of _ "norm \ f", symmetric] proof (rule integral_norm_bound_integral) show "(\x. if x \ g ` ?U k then (norm \ f) x else 0) integrable_on UNIV" "(\x. if x \ g ` (\n. F n) then (norm \ f) x else 0) integrable_on UNIV" using fgU fs unfolding absolutely_integrable_on_def o_def integrable_restrict_UNIV by (auto simp: image_UN) qed auto qed ultimately show "bounded (range (\k. integral UNIV (?nD k)))" unfolding integral_restrict_UNIV image_UN [symmetric] o_def by simp next show "(\k. if x \ ?U k then norm (?D x) else 0) \ (if x \ (\n. F n) then norm (?D x) else 0)" for x by (force intro: tendsto_eventually eventually_sequentiallyI) qed auto next show "(\k. if x \ ?U k then ?D x else 0) \ (if x \ (\n. F n) then ?D x else 0)" for x proof clarsimp fix n assume "x \ F n" show "(\m. if \j\{..m}. x \ F j then ?D x else 0) \ ?D x" using \x \ F n\ by (auto intro!: tendsto_eventually eventually_sequentiallyI [of n]) qed qed auto then show Dai: "?D absolutely_integrable_on (\n. F n)" unfolding absolutely_integrable_restrict_UNIV by simp show "integral (\n. F n) ?D = integral ((\x. g ` F x)) f" proof (rule LIMSEQ_unique) show "(\n. integral (\m\n. g ` F m) f) \ integral (\x. g ` F x) f" by (rule fgU) show "(\n. integral (\m\n. g ` F m) f) \ integral (\n. F n) ?D" unfolding D_int [symmetric] by (rule integral_countable_UN [OF Dai F_leb]) qed qed qed theorem has_absolute_integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - obtain C N where "fsigma C" and N: "N \ null_sets lebesgue" and CNS: "C \ N = S" and "disjnt C N" using lebesgue_set_almost_fsigma [OF S] . then obtain F :: "nat \ (real^'m::_) set" where F: "range F \ Collect compact" and Ceq: "C = Union(range F)" using fsigma_Union_compact by metis have "negligible N" using N by (simp add: negligible_iff_null_sets) let ?D = "\x. \det (matrix (g' x))\ *\<^sub>R f (g x)" have "?D absolutely_integrable_on C \ integral C ?D = b \ f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b" unfolding Ceq proof (rule has_absolute_integral_change_of_variables_compact_family) fix n x assume "x \ \(F ` UNIV)" then show "(g has_derivative g' x) (at x within \(F ` UNIV))" using Ceq \C \ N = S\ der_g has_derivative_within_subset by blast next have "\(F ` UNIV) \ S" using Ceq \C \ N = S\ by blast then show "inj_on g (\(F ` UNIV))" using inj by (meson inj_on_subset) qed (use F in auto) moreover have "?D absolutely_integrable_on C \ integral C ?D = b \ ?D absolutely_integrable_on S \ integral S ?D = b" proof (rule conj_cong) have neg: "negligible {x \ C - S. ?D x \ 0}" "negligible {x \ S - C. ?D x \ 0}" using CNS by (blast intro: negligible_subset [OF \negligible N\])+ then show "(?D absolutely_integrable_on C) = (?D absolutely_integrable_on S)" by (rule absolutely_integrable_spike_set_eq) show "(integral C ?D = b) \ (integral S ?D = b)" using integral_spike_set [OF neg] by simp qed moreover have "f absolutely_integrable_on (g ` C) \ integral (g ` C) f = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (rule conj_cong) have "g differentiable_on N" by (metis CNS der_g differentiable_def differentiable_on_def differentiable_on_subset sup.cobounded2) with \negligible N\ have neg_gN: "negligible (g ` N)" by (blast intro: negligible_differentiable_image_negligible) have neg: "negligible {x \ g ` C - g ` S. f x \ 0}" "negligible {x \ g ` S - g ` C. f x \ 0}" using CNS by (blast intro: negligible_subset [OF neg_gN])+ then show "(f absolutely_integrable_on g ` C) = (f absolutely_integrable_on g ` S)" by (rule absolutely_integrable_spike_set_eq) show "(integral (g ` C) f = b) \ (integral (g ` S) f = b)" using integral_spike_set [OF neg] by simp qed ultimately show ?thesis by simp qed corollary absolutely_integrable_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (g has_derivative g' x) (at x within S)" and "inj_on g S" shows "f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S" using assms has_absolute_integral_change_of_variables by blast corollary integral_change_of_variables: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_derivative g' x) (at x within S)" and inj: "inj_on g S" and disj: "(f absolutely_integrable_on (g ` S) \ (\x. \det (matrix (g' x))\ *\<^sub>R f(g x)) absolutely_integrable_on S)" shows "integral (g ` S) f = integral S (\x. \det (matrix (g' x))\ *\<^sub>R f(g x))" using has_absolute_integral_change_of_variables [OF S der_g inj] disj by blast lemma has_absolute_integral_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \g' x\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof - let ?lift = "vec :: real \ real^1" let ?drop = "(\x::real^1. x $ 1)" have S': "?lift ` S \ sets lebesgue" by (auto intro: differentiable_image_in_sets_lebesgue [OF S] differentiable_vec) have "((\x. vec (g (x $ 1))) has_derivative (*\<^sub>R) (g' z)) (at (vec z) within ?lift ` S)" if "z \ S" for z using der_g [OF that] by (simp add: has_vector_derivative_def has_derivative_vector_1) then have der': "\x. x \ ?lift ` S \ (?lift \ g \ ?drop has_derivative (*\<^sub>R) (g' (?drop x))) (at x within ?lift ` S)" by (auto simp: o_def) have inj': "inj_on (vec \ g \ ?drop) (vec ` S)" using inj by (simp add: inj_on_def) let ?fg = "\x. \g' x\ *\<^sub>R f(g x)" have "((\x. ?fg x $ i) absolutely_integrable_on S \ ((\x. ?fg x $ i) has_integral b $ i) S \ (\x. f x $ i) absolutely_integrable_on g ` S \ ((\x. f x $ i) has_integral b $ i) (g ` S))" for i using has_absolute_integral_change_of_variables [OF S' der' inj', of "\x. ?lift(f (?drop x) $ i)" "?lift (b$i)"] unfolding integrable_on_1_iff integral_on_1_eq absolutely_integrable_on_1_iff absolutely_integrable_drop absolutely_integrable_on_def by (auto simp: image_comp o_def integral_vec1_eq has_integral_iff) then have "?fg absolutely_integrable_on S \ (?fg has_integral b) S \ f absolutely_integrable_on (g ` S) \ (f has_integral b) (g ` S)" unfolding has_integral_componentwise_iff [where y=b] absolutely_integrable_componentwise_iff [where f=f] absolutely_integrable_componentwise_iff [where f = ?fg] by (force simp: Basis_vec_def cart_eq_inner_axis) then show ?thesis using absolutely_integrable_on_def by blast qed corollary absolutely_integrable_change_of_variables_1: fixes f :: "real \ real^'n::{finite,wellorder}" and g :: "real \ real" assumes S: "S \ sets lebesgue" and der_g: "\x. x \ S \ (g has_vector_derivative g' x) (at x within S)" and inj: "inj_on g S" shows "(f absolutely_integrable_on g ` S \ (\x. \g' x\ *\<^sub>R f(g x)) absolutely_integrable_on S)" using has_absolute_integral_change_of_variables_1 [OF assms] by auto subsection\Change of variables for integrals: special case of linear function\ lemma has_absolute_integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ integral S (\x. \det (matrix g)\ *\<^sub>R f(g x)) = b \ f absolutely_integrable_on (g ` S) \ integral (g ` S) f = b" proof (cases "det(matrix g) = 0") case True then have "negligible(g ` S)" using assms det_nz_iff_inj negligible_linear_singular_image by blast with True show ?thesis by (auto simp: absolutely_integrable_on_def integrable_negligible integral_negligible) next case False then obtain h where h: "\x. x \ S \ h (g x) = x" "linear h" - using assms det_nz_iff_inj linear_injective_isomorphism by blast + using assms det_nz_iff_inj linear_injective_isomorphism by metis show ?thesis proof (rule has_absolute_integral_change_of_variables_invertible) show "(g has_derivative g) (at x within S)" for x by (simp add: assms linear_imp_has_derivative) show "continuous_on (g ` S) h" using continuous_on_eq_continuous_within has_derivative_continuous linear_imp_has_derivative h by blast qed (use h in auto) qed lemma absolutely_integrable_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "(\x. \det (matrix g)\ *\<^sub>R f(g x)) absolutely_integrable_on S \ f absolutely_integrable_on (g ` S)" using assms has_absolute_integral_change_of_variables_linear by blast lemma absolutely_integrable_on_linear_image: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" shows "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S \ det(matrix g) = 0" unfolding assms absolutely_integrable_change_of_variables_linear [OF assms, symmetric] absolutely_integrable_on_scaleR_iff by (auto simp: set_integrable_def) lemma integral_change_of_variables_linear: fixes f :: "real^'m::{finite,wellorder} \ real^'n" and g :: "real^'m::_ \ real^'m::_" assumes "linear g" and "f absolutely_integrable_on (g ` S) \ (f \ g) absolutely_integrable_on S" shows "integral (g ` S) f = \det (matrix g)\ *\<^sub>R integral S (f \ g)" proof - have "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S) \ (f absolutely_integrable_on g ` S)" using absolutely_integrable_on_linear_image assms by blast moreover have ?thesis if "((\x. \det (matrix g)\ *\<^sub>R f (g x)) absolutely_integrable_on S)" "(f absolutely_integrable_on g ` S)" using has_absolute_integral_change_of_variables_linear [OF \linear g\] that by (auto simp: o_def) ultimately show ?thesis using absolutely_integrable_change_of_variables_linear [OF \linear g\] by blast qed subsection\Change of variable for measure\ lemma has_measure_differentiable_image: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ measure lebesgue (f ` S) = m \ ((\x. \det (matrix (f' x))\) has_integral m) S" using has_absolute_integral_change_of_variables [OF assms, of "\x. (1::real^1)" "vec m"] unfolding absolutely_integrable_on_1_iff integral_on_1_eq integrable_on_1_iff absolutely_integrable_on_def by (auto simp: has_integral_iff lmeasurable_iff_integrable_on lmeasure_integral) lemma measurable_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) integrable_on S" using has_measure_differentiable_image [OF assms] by blast lemma measurable_differentiable_image_alt: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes "S \ sets lebesgue" and "\x. x \ S \ (f has_derivative f' x) (at x within S)" and "inj_on f S" shows "f ` S \ lmeasurable \ (\x. \det (matrix (f' x))\) absolutely_integrable_on S" using measurable_differentiable_image_eq [OF assms] by (simp only: absolutely_integrable_on_iff_nonneg) lemma measure_differentiable_image_eq: fixes f :: "real^'n::{finite,wellorder} \ real^'n::_" assumes S: "S \ sets lebesgue" and der_f: "\x. x \ S \ (f has_derivative f' x) (at x within S)" and inj: "inj_on f S" and intS: "(\x. \det (matrix (f' x))\) integrable_on S" shows "measure lebesgue (f ` S) = integral S (\x. \det (matrix (f' x))\)" using measurable_differentiable_image_eq [OF S der_f inj] assms has_measure_differentiable_image by blast end diff --git a/src/HOL/Analysis/Complex_Analysis_Basics.thy b/src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy @@ -1,966 +1,854 @@ (* Author: John Harrison, Marco Maggesi, Graziano Gentili, Gianni Ciolli, Valentina Bruno Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014) *) section \Complex Analysis Basics\ text \Definitions of analytic and holomorphic functions, limit theorems, complex differentiation\ theory Complex_Analysis_Basics imports Derivative "HOL-Library.Nonpos_Ints" begin subsection\<^marker>\tag unimportant\\General lemmas\ lemma nonneg_Reals_cmod_eq_Re: "z \ \\<^sub>\\<^sub>0 \ norm z = Re z" by (simp add: complex_nonneg_Reals_iff cmod_eq_Re) lemma fact_cancel: fixes c :: "'a::real_field" shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)" using of_nat_neq_0 by force lemma vector_derivative_cnj_within: assumes "at x within A \ bot" and "f differentiable at x within A" shows "vector_derivative (\z. cnj (f z)) (at x within A) = cnj (vector_derivative f (at x within A))" (is "_ = cnj ?D") proof - let ?D = "vector_derivative f (at x within A)" from assms have "(f has_vector_derivative ?D) (at x within A)" by (subst (asm) vector_derivative_works) hence "((\x. cnj (f x)) has_vector_derivative cnj ?D) (at x within A)" by (rule has_vector_derivative_cnj) thus ?thesis using assms by (auto dest: vector_derivative_within) qed lemma vector_derivative_cnj: assumes "f differentiable at x" shows "vector_derivative (\z. cnj (f z)) (at x) = cnj (vector_derivative f (at x))" using assms by (intro vector_derivative_cnj_within) auto lemma shows open_halfspace_Re_lt: "open {z. Re(z) < b}" and open_halfspace_Re_gt: "open {z. Re(z) > b}" and closed_halfspace_Re_ge: "closed {z. Re(z) \ b}" and closed_halfspace_Re_le: "closed {z. Re(z) \ b}" and closed_halfspace_Re_eq: "closed {z. Re(z) = b}" and open_halfspace_Im_lt: "open {z. Im(z) < b}" and open_halfspace_Im_gt: "open {z. Im(z) > b}" and closed_halfspace_Im_ge: "closed {z. Im(z) \ b}" and closed_halfspace_Im_le: "closed {z. Im(z) \ b}" and closed_halfspace_Im_eq: "closed {z. Im(z) = b}" by (intro open_Collect_less closed_Collect_le closed_Collect_eq continuous_on_Re continuous_on_Im continuous_on_id continuous_on_const)+ lemma closed_complex_Reals: "closed (\ :: complex set)" proof - have "(\ :: complex set) = {z. Im z = 0}" by (auto simp: complex_is_Real_iff) then show ?thesis by (metis closed_halfspace_Im_eq) qed lemma closed_Real_halfspace_Re_le: "closed (\ \ {w. Re w \ x})" by (simp add: closed_Int closed_complex_Reals closed_halfspace_Re_le) lemma closed_nonpos_Reals_complex [simp]: "closed (\\<^sub>\\<^sub>0 :: complex set)" proof - have "\\<^sub>\\<^sub>0 = \ \ {z. Re(z) \ 0}" using complex_nonpos_Reals_iff complex_is_Real_iff by auto then show ?thesis by (metis closed_Real_halfspace_Re_le) qed lemma closed_Real_halfspace_Re_ge: "closed (\ \ {w. x \ Re(w)})" using closed_halfspace_Re_ge by (simp add: closed_Int closed_complex_Reals) lemma closed_nonneg_Reals_complex [simp]: "closed (\\<^sub>\\<^sub>0 :: complex set)" proof - have "\\<^sub>\\<^sub>0 = \ \ {z. Re(z) \ 0}" using complex_nonneg_Reals_iff complex_is_Real_iff by auto then show ?thesis by (metis closed_Real_halfspace_Re_ge) qed lemma closed_real_abs_le: "closed {w \ \. \Re w\ \ r}" proof - have "{w \ \. \Re w\ \ r} = (\ \ {w. Re w \ r}) \ (\ \ {w. Re w \ -r})" by auto then show "closed {w \ \. \Re w\ \ r}" by (simp add: closed_Int closed_Real_halfspace_Re_ge closed_Real_halfspace_Re_le) qed lemma real_lim: fixes l::complex assumes "(f \ l) F" and "\ trivial_limit F" and "eventually P F" and "\a. P a \ f a \ \" shows "l \ \" proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)]) show "eventually (\x. f x \ \) F" using assms(3, 4) by (auto intro: eventually_mono) qed lemma real_lim_sequentially: fixes l::complex shows "(f \ l) sequentially \ (\N. \n\N. f n \ \) \ l \ \" by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially) lemma real_series: fixes l::complex shows "f sums l \ (\n. f n \ \) \ l \ \" unfolding sums_def by (metis real_lim_sequentially sum_in_Reals) lemma Lim_null_comparison_Re: assumes "eventually (\x. norm(f x) \ Re(g x)) F" "(g \ 0) F" shows "(f \ 0) F" by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp -lemma closed_segment_same_Re: - assumes "Re a = Re b" - shows "closed_segment a b = {z. Re z = Re a \ Im z \ closed_segment (Im a) (Im b)}" -proof safe - fix z assume "z \ closed_segment a b" - then obtain u where u: "u \ {0..1}" "z = a + of_real u * (b - a)" - by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) - from assms show "Re z = Re a" by (auto simp: u) - from u(1) show "Im z \ closed_segment (Im a) (Im b)" - by (force simp: u closed_segment_def algebra_simps) -next - fix z assume [simp]: "Re z = Re a" and "Im z \ closed_segment (Im a) (Im b)" - then obtain u where u: "u \ {0..1}" "Im z = Im a + of_real u * (Im b - Im a)" - by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) - from u(1) show "z \ closed_segment a b" using assms - by (force simp: u closed_segment_def algebra_simps scaleR_conv_of_real complex_eq_iff) -qed - -lemma closed_segment_same_Im: - assumes "Im a = Im b" - shows "closed_segment a b = {z. Im z = Im a \ Re z \ closed_segment (Re a) (Re b)}" -proof safe - fix z assume "z \ closed_segment a b" - then obtain u where u: "u \ {0..1}" "z = a + of_real u * (b - a)" - by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) - from assms show "Im z = Im a" by (auto simp: u) - from u(1) show "Re z \ closed_segment (Re a) (Re b)" - by (force simp: u closed_segment_def algebra_simps) -next - fix z assume [simp]: "Im z = Im a" and "Re z \ closed_segment (Re a) (Re b)" - then obtain u where u: "u \ {0..1}" "Re z = Re a + of_real u * (Re b - Re a)" - by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) - from u(1) show "z \ closed_segment a b" using assms - by (force simp: u closed_segment_def algebra_simps scaleR_conv_of_real complex_eq_iff) -qed - subsection\Holomorphic functions\ definition\<^marker>\tag important\ holomorphic_on :: "[complex \ complex, complex set] \ bool" (infixl "(holomorphic'_on)" 50) where "f holomorphic_on s \ \x\s. f field_differentiable (at x within s)" named_theorems\<^marker>\tag important\ holomorphic_intros "structural introduction rules for holomorphic_on" lemma holomorphic_onI [intro?]: "(\x. x \ s \ f field_differentiable (at x within s)) \ f holomorphic_on s" by (simp add: holomorphic_on_def) lemma holomorphic_onD [dest?]: "\f holomorphic_on s; x \ s\ \ f field_differentiable (at x within s)" by (simp add: holomorphic_on_def) lemma holomorphic_on_imp_differentiable_on: "f holomorphic_on s \ f differentiable_on s" unfolding holomorphic_on_def differentiable_on_def by (simp add: field_differentiable_imp_differentiable) lemma holomorphic_on_imp_differentiable_at: "\f holomorphic_on s; open s; x \ s\ \ f field_differentiable (at x)" using at_within_open holomorphic_on_def by fastforce lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}" by (simp add: holomorphic_on_def) lemma holomorphic_on_open: "open s \ f holomorphic_on s \ (\x \ s. \f'. DERIV f x :> f')" by (auto simp: holomorphic_on_def field_differentiable_def has_field_derivative_def at_within_open [of _ s]) lemma holomorphic_on_imp_continuous_on: "f holomorphic_on s \ continuous_on s f" by (metis field_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def) lemma holomorphic_on_subset [elim]: "f holomorphic_on s \ t \ s \ f holomorphic_on t" unfolding holomorphic_on_def by (metis field_differentiable_within_subset subsetD) lemma holomorphic_transform: "\f holomorphic_on s; \x. x \ s \ f x = g x\ \ g holomorphic_on s" by (metis field_differentiable_transform_within linordered_field_no_ub holomorphic_on_def) lemma holomorphic_cong: "s = t ==> (\x. x \ s \ f x = g x) \ f holomorphic_on s \ g holomorphic_on t" by (metis holomorphic_transform) lemma holomorphic_on_linear [simp, holomorphic_intros]: "((*) c) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_linear) lemma holomorphic_on_const [simp, holomorphic_intros]: "(\z. c) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_const) lemma holomorphic_on_ident [simp, holomorphic_intros]: "(\x. x) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_ident) lemma holomorphic_on_id [simp, holomorphic_intros]: "id holomorphic_on s" unfolding id_def by (rule holomorphic_on_ident) lemma holomorphic_on_compose: "f holomorphic_on s \ g holomorphic_on (f ` s) \ (g o f) holomorphic_on s" using field_differentiable_compose_within[of f _ s g] by (auto simp: holomorphic_on_def) lemma holomorphic_on_compose_gen: "f holomorphic_on s \ g holomorphic_on t \ f ` s \ t \ (g o f) holomorphic_on s" by (metis holomorphic_on_compose holomorphic_on_subset) lemma holomorphic_on_balls_imp_entire: assumes "\bdd_above A" "\r. r \ A \ f holomorphic_on ball c r" shows "f holomorphic_on B" proof (rule holomorphic_on_subset) show "f holomorphic_on UNIV" unfolding holomorphic_on_def proof fix z :: complex from \\bdd_above A\ obtain r where r: "r \ A" "r > norm (z - c)" by (meson bdd_aboveI not_le) with assms(2) have "f holomorphic_on ball c r" by blast moreover from r have "z \ ball c r" by (auto simp: dist_norm norm_minus_commute) ultimately show "f field_differentiable at z" by (auto simp: holomorphic_on_def at_within_open[of _ "ball c r"]) qed qed auto lemma holomorphic_on_balls_imp_entire': assumes "\r. r > 0 \ f holomorphic_on ball c r" shows "f holomorphic_on B" proof (rule holomorphic_on_balls_imp_entire) { fix M :: real have "\x. x > max M 0" by (intro gt_ex) hence "\x>0. x > M" by auto } thus "\bdd_above {(0::real)<..}" unfolding bdd_above_def by (auto simp: not_le) qed (insert assms, auto) lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \ (\z. -(f z)) holomorphic_on s" by (metis field_differentiable_minus holomorphic_on_def) lemma holomorphic_on_add [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z + g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_add) lemma holomorphic_on_diff [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z - g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_diff) lemma holomorphic_on_mult [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z * g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_mult) lemma holomorphic_on_inverse [holomorphic_intros]: "\f holomorphic_on s; \z. z \ s \ f z \ 0\ \ (\z. inverse (f z)) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_inverse) lemma holomorphic_on_divide [holomorphic_intros]: "\f holomorphic_on s; g holomorphic_on s; \z. z \ s \ g z \ 0\ \ (\z. f z / g z) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_divide) lemma holomorphic_on_power [holomorphic_intros]: "f holomorphic_on s \ (\z. (f z)^n) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_power) lemma holomorphic_on_sum [holomorphic_intros]: "(\i. i \ I \ (f i) holomorphic_on s) \ (\x. sum (\i. f i x) I) holomorphic_on s" unfolding holomorphic_on_def by (metis field_differentiable_sum) lemma holomorphic_on_prod [holomorphic_intros]: "(\i. i \ I \ (f i) holomorphic_on s) \ (\x. prod (\i. f i x) I) holomorphic_on s" by (induction I rule: infinite_finite_induct) (auto intro: holomorphic_intros) lemma holomorphic_pochhammer [holomorphic_intros]: "f holomorphic_on A \ (\s. pochhammer (f s) n) holomorphic_on A" by (induction n) (auto intro!: holomorphic_intros simp: pochhammer_Suc) lemma holomorphic_on_scaleR [holomorphic_intros]: "f holomorphic_on A \ (\x. c *\<^sub>R f x) holomorphic_on A" by (auto simp: scaleR_conv_of_real intro!: holomorphic_intros) lemma holomorphic_on_Un [holomorphic_intros]: assumes "f holomorphic_on A" "f holomorphic_on B" "open A" "open B" shows "f holomorphic_on (A \ B)" using assms by (auto simp: holomorphic_on_def at_within_open[of _ A] at_within_open[of _ B] at_within_open[of _ "A \ B"] open_Un) lemma holomorphic_on_If_Un [holomorphic_intros]: assumes "f holomorphic_on A" "g holomorphic_on B" "open A" "open B" assumes "\z. z \ A \ z \ B \ f z = g z" shows "(\z. if z \ A then f z else g z) holomorphic_on (A \ B)" (is "?h holomorphic_on _") proof (intro holomorphic_on_Un) note \f holomorphic_on A\ also have "f holomorphic_on A \ ?h holomorphic_on A" by (intro holomorphic_cong) auto finally show \ . next note \g holomorphic_on B\ also have "g holomorphic_on B \ ?h holomorphic_on B" using assms by (intro holomorphic_cong) auto finally show \ . qed (insert assms, auto) -lemma DERIV_deriv_iff_field_differentiable: - "DERIV f x :> deriv f x \ f field_differentiable at x" - unfolding field_differentiable_def by (metis DERIV_imp_deriv) - lemma holomorphic_derivI: "\f holomorphic_on S; open S; x \ S\ \ (f has_field_derivative deriv f x) (at x within T)" by (metis DERIV_deriv_iff_field_differentiable at_within_open holomorphic_on_def has_field_derivative_at_within) -lemma complex_derivative_chain: - "f field_differentiable at x \ g field_differentiable at (f x) - \ deriv (g o f) x = deriv g (f x) * deriv f x" - by (metis DERIV_deriv_iff_field_differentiable DERIV_chain DERIV_imp_deriv) - -lemma deriv_linear [simp]: "deriv (\w. c * w) = (\z. c)" - by (metis DERIV_imp_deriv DERIV_cmult_Id) - -lemma deriv_ident [simp]: "deriv (\w. w) = (\z. 1)" - by (metis DERIV_imp_deriv DERIV_ident) - -lemma deriv_id [simp]: "deriv id = (\z. 1)" - by (simp add: id_def) - -lemma deriv_const [simp]: "deriv (\w. c) = (\z. 0)" - by (metis DERIV_imp_deriv DERIV_const) - -lemma deriv_add [simp]: - "\f field_differentiable at z; g field_differentiable at z\ - \ deriv (\w. f w + g w) z = deriv f z + deriv g z" - unfolding DERIV_deriv_iff_field_differentiable[symmetric] - by (auto intro!: DERIV_imp_deriv derivative_intros) - -lemma deriv_diff [simp]: - "\f field_differentiable at z; g field_differentiable at z\ - \ deriv (\w. f w - g w) z = deriv f z - deriv g z" - unfolding DERIV_deriv_iff_field_differentiable[symmetric] - by (auto intro!: DERIV_imp_deriv derivative_intros) - -lemma deriv_mult [simp]: - "\f field_differentiable at z; g field_differentiable at z\ - \ deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" - unfolding DERIV_deriv_iff_field_differentiable[symmetric] - by (auto intro!: DERIV_imp_deriv derivative_eq_intros) - -lemma deriv_cmult: - "f field_differentiable at z \ deriv (\w. c * f w) z = c * deriv f z" - by simp - -lemma deriv_cmult_right: - "f field_differentiable at z \ deriv (\w. f w * c) z = deriv f z * c" - by simp - -lemma deriv_inverse [simp]: - "\f field_differentiable at z; f z \ 0\ - \ deriv (\w. inverse (f w)) z = - deriv f z / f z ^ 2" - unfolding DERIV_deriv_iff_field_differentiable[symmetric] - by (safe intro!: DERIV_imp_deriv derivative_eq_intros) (auto simp: field_split_simps power2_eq_square) - -lemma deriv_divide [simp]: - "\f field_differentiable at z; g field_differentiable at z; g z \ 0\ - \ deriv (\w. f w / g w) z = (deriv f z * g z - f z * deriv g z) / g z ^ 2" - by (simp add: field_class.field_divide_inverse field_differentiable_inverse) - (simp add: field_split_simps power2_eq_square) - -lemma deriv_cdivide_right: - "f field_differentiable at z \ deriv (\w. f w / c) z = deriv f z / c" - by (simp add: field_class.field_divide_inverse) - lemma complex_derivative_transform_within_open: "\f holomorphic_on s; g holomorphic_on s; open s; z \ s; \w. w \ s \ f w = g w\ \ deriv f z = deriv g z" unfolding holomorphic_on_def by (rule DERIV_imp_deriv) (metis DERIV_deriv_iff_field_differentiable has_field_derivative_transform_within_open at_within_open) -lemma deriv_compose_linear: - "f field_differentiable at (c * z) \ deriv (\w. f (c * w)) z = c * deriv f (c * z)" -apply (rule DERIV_imp_deriv) - unfolding DERIV_deriv_iff_field_differentiable [symmetric] - by (metis (full_types) DERIV_chain2 DERIV_cmult_Id mult.commute) - - -lemma nonzero_deriv_nonconstant: - assumes df: "DERIV f \ :> df" and S: "open S" "\ \ S" and "df \ 0" - shows "\ f constant_on S" -unfolding constant_on_def -by (metis \df \ 0\ has_field_derivative_transform_within_open [OF df S] DERIV_const DERIV_unique) - lemma holomorphic_nonconstant: assumes holf: "f holomorphic_on S" and "open S" "\ \ S" "deriv f \ \ 0" shows "\ f constant_on S" by (rule nonzero_deriv_nonconstant [of f "deriv f \" \ S]) (use assms in \auto simp: holomorphic_derivI\) subsection\Analyticity on a set\ definition\<^marker>\tag important\ analytic_on (infixl "(analytic'_on)" 50) where "f analytic_on S \ \x \ S. \e. 0 < e \ f holomorphic_on (ball x e)" named_theorems\<^marker>\tag important\ analytic_intros "introduction rules for proving analyticity" lemma analytic_imp_holomorphic: "f analytic_on S \ f holomorphic_on S" by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def) (metis centre_in_ball field_differentiable_at_within) lemma analytic_on_open: "open S \ f analytic_on S \ f holomorphic_on S" apply (auto simp: analytic_imp_holomorphic) apply (auto simp: analytic_on_def holomorphic_on_def) by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball) lemma analytic_on_imp_differentiable_at: "f analytic_on S \ x \ S \ f field_differentiable (at x)" apply (auto simp: analytic_on_def holomorphic_on_def) by (metis open_ball centre_in_ball field_differentiable_within_open) lemma analytic_on_subset: "f analytic_on S \ T \ S \ f analytic_on T" by (auto simp: analytic_on_def) lemma analytic_on_Un: "f analytic_on (S \ T) \ f analytic_on S \ f analytic_on T" by (auto simp: analytic_on_def) lemma analytic_on_Union: "f analytic_on (\\) \ (\T \ \. f analytic_on T)" by (auto simp: analytic_on_def) lemma analytic_on_UN: "f analytic_on (\i\I. S i) \ (\i\I. f analytic_on (S i))" by (auto simp: analytic_on_def) lemma analytic_on_holomorphic: "f analytic_on S \ (\T. open T \ S \ T \ f holomorphic_on T)" (is "?lhs = ?rhs") proof - have "?lhs \ (\T. open T \ S \ T \ f analytic_on T)" proof safe assume "f analytic_on S" then show "\T. open T \ S \ T \ f analytic_on T" apply (simp add: analytic_on_def) apply (rule exI [where x="\{U. open U \ f analytic_on U}"], auto) apply (metis open_ball analytic_on_open centre_in_ball) by (metis analytic_on_def) next fix T assume "open T" "S \ T" "f analytic_on T" then show "f analytic_on S" by (metis analytic_on_subset) qed also have "... \ ?rhs" by (auto simp: analytic_on_open) finally show ?thesis . qed lemma analytic_on_linear [analytic_intros,simp]: "((*) c) analytic_on S" by (auto simp add: analytic_on_holomorphic) lemma analytic_on_const [analytic_intros,simp]: "(\z. c) analytic_on S" by (metis analytic_on_def holomorphic_on_const zero_less_one) lemma analytic_on_ident [analytic_intros,simp]: "(\x. x) analytic_on S" by (simp add: analytic_on_def gt_ex) lemma analytic_on_id [analytic_intros]: "id analytic_on S" unfolding id_def by (rule analytic_on_ident) lemma analytic_on_compose: assumes f: "f analytic_on S" and g: "g analytic_on (f ` S)" shows "(g o f) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix x assume x: "x \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g by (metis analytic_on_def g image_eqI x) have "isCont f x" by (metis analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at f x) with e' obtain d where d: "0 < d" and fd: "f ` ball x d \ ball (f x) e'" by (auto simp: continuous_at_ball) have "g \ f holomorphic_on ball x (min d e)" apply (rule holomorphic_on_compose) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball) then show "\e>0. g \ f holomorphic_on ball x e" by (metis d e min_less_iff_conj) qed lemma analytic_on_compose_gen: "f analytic_on S \ g analytic_on T \ (\z. z \ S \ f z \ T) \ g o f analytic_on S" by (metis analytic_on_compose analytic_on_subset image_subset_iff) lemma analytic_on_neg [analytic_intros]: "f analytic_on S \ (\z. -(f z)) analytic_on S" by (metis analytic_on_holomorphic holomorphic_on_minus) lemma analytic_on_add [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" shows "(\z. f z + g z) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g by (metis analytic_on_def g z) have "(\z. f z + g z) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_add) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z + g z) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_diff [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" shows "(\z. f z - g z) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g by (metis analytic_on_def g z) have "(\z. f z - g z) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_diff) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z - g z) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_mult [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" shows "(\z. f z * g z) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g by (metis analytic_on_def g z) have "(\z. f z * g z) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_mult) apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) then show "\e>0. (\z. f z * g z) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_inverse [analytic_intros]: assumes f: "f analytic_on S" and nz: "(\z. z \ S \ f z \ 0)" shows "(\z. inverse (f z)) analytic_on S" unfolding analytic_on_def proof (intro ballI) fix z assume z: "z \ S" then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f by (metis analytic_on_def) have "continuous_on (ball z e) f" by (metis fh holomorphic_on_imp_continuous_on) then obtain e' where e': "0 < e'" and nz': "\y. dist z y < e' \ f y \ 0" by (metis open_ball centre_in_ball continuous_on_open_avoid e z nz) have "(\z. inverse (f z)) holomorphic_on ball z (min e e')" apply (rule holomorphic_on_inverse) apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball) by (metis nz' mem_ball min_less_iff_conj) then show "\e>0. (\z. inverse (f z)) holomorphic_on ball z e" by (metis e e' min_less_iff_conj) qed lemma analytic_on_divide [analytic_intros]: assumes f: "f analytic_on S" and g: "g analytic_on S" and nz: "(\z. z \ S \ g z \ 0)" shows "(\z. f z / g z) analytic_on S" unfolding divide_inverse by (metis analytic_on_inverse analytic_on_mult f g nz) lemma analytic_on_power [analytic_intros]: "f analytic_on S \ (\z. (f z) ^ n) analytic_on S" by (induct n) (auto simp: analytic_on_mult) lemma analytic_on_sum [analytic_intros]: "(\i. i \ I \ (f i) analytic_on S) \ (\x. sum (\i. f i x) I) analytic_on S" by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add) lemma deriv_left_inverse: assumes "f holomorphic_on S" and "g holomorphic_on T" and "open S" and "open T" and "f ` S \ T" and [simp]: "\z. z \ S \ g (f z) = z" and "w \ S" shows "deriv f w * deriv g (f w) = 1" proof - have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w" by (simp add: algebra_simps) also have "... = deriv (g o f) w" using assms - by (metis analytic_on_imp_differentiable_at analytic_on_open complex_derivative_chain image_subset_iff) + by (metis analytic_on_imp_differentiable_at analytic_on_open deriv_chain image_subset_iff) also have "... = deriv id w" proof (rule complex_derivative_transform_within_open [where s=S]) show "g \ f holomorphic_on S" by (rule assms holomorphic_on_compose_gen holomorphic_intros)+ qed (use assms in auto) also have "... = 1" by simp finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Analyticity at a point\ lemma analytic_at_ball: "f analytic_on {z} \ (\e. 0 f holomorphic_on ball z e)" by (metis analytic_on_def singleton_iff) lemma analytic_at: "f analytic_on {z} \ (\s. open s \ z \ s \ f holomorphic_on s)" by (metis analytic_on_holomorphic empty_subsetI insert_subset) lemma analytic_on_analytic_at: "f analytic_on s \ (\z \ s. f analytic_on {z})" by (metis analytic_at_ball analytic_on_def) lemma analytic_at_two: "f analytic_on {z} \ g analytic_on {z} \ (\s. open s \ z \ s \ f holomorphic_on s \ g holomorphic_on s)" (is "?lhs = ?rhs") proof assume ?lhs then obtain s t where st: "open s" "z \ s" "f holomorphic_on s" "open t" "z \ t" "g holomorphic_on t" by (auto simp: analytic_at) show ?rhs apply (rule_tac x="s \ t" in exI) using st apply (auto simp: holomorphic_on_subset) done next assume ?rhs then show ?lhs by (force simp add: analytic_at) qed subsection\<^marker>\tag unimportant\\Combining theorems for derivative with ``analytic at'' hypotheses\ lemma assumes "f analytic_on {z}" "g analytic_on {z}" shows complex_derivative_add_at: "deriv (\w. f w + g w) z = deriv f z + deriv g z" and complex_derivative_diff_at: "deriv (\w. f w - g w) z = deriv f z - deriv g z" and complex_derivative_mult_at: "deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" proof - obtain s where s: "open s" "z \ s" "f holomorphic_on s" "g holomorphic_on s" using assms by (metis analytic_at_two) show "deriv (\w. f w + g w) z = deriv f z + deriv g z" apply (rule DERIV_imp_deriv [OF DERIV_add]) using s apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) done show "deriv (\w. f w - g w) z = deriv f z - deriv g z" apply (rule DERIV_imp_deriv [OF DERIV_diff]) using s apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) done show "deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" apply (rule DERIV_imp_deriv [OF DERIV_mult']) using s apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) done qed lemma deriv_cmult_at: "f analytic_on {z} \ deriv (\w. c * f w) z = c * deriv f z" by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) lemma deriv_cmult_right_at: "f analytic_on {z} \ deriv (\w. f w * c) z = deriv f z * c" by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) subsection\<^marker>\tag unimportant\\Complex differentiation of sequences and series\ (* TODO: Could probably be simplified using Uniform_Limit *) lemma has_complex_derivative_sequence: fixes S :: "complex set" assumes cvs: "convex S" and df: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" and conv: "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ norm (f' n x - g' x) \ e" and "\x l. x \ S \ ((\n. f n x) \ l) sequentially" shows "\g. \x \ S. ((\n. f n x) \ g x) sequentially \ (g has_field_derivative (g' x)) (at x within S)" proof - from assms obtain x l where x: "x \ S" and tf: "((\n. f n x) \ l) sequentially" by blast { fix e::real assume e: "e > 0" then obtain N where N: "\n\N. \x. x \ S \ cmod (f' n x - g' x) \ e" by (metis conv) have "\N. \n\N. \x\S. \h. cmod (f' n x * h - g' x * h) \ e * cmod h" proof (rule exI [of _ N], clarify) fix n y h assume "N \ n" "y \ S" then have "cmod (f' n y - g' y) \ e" by (metis N) then have "cmod h * cmod (f' n y - g' y) \ cmod h * e" by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) then show "cmod (f' n y * h - g' y * h) \ e * cmod h" by (simp add: norm_mult [symmetric] field_simps) qed } note ** = this show ?thesis unfolding has_field_derivative_def proof (rule has_derivative_sequence [OF cvs _ _ x]) show "(\n. f n x) \ l" by (rule tf) next show "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. cmod (f' n x * h - g' x * h) \ e * cmod h" unfolding eventually_sequentially by (blast intro: **) qed (metis has_field_derivative_def df) qed lemma has_complex_derivative_series: fixes S :: "complex set" assumes cvs: "convex S" and df: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" and conv: "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ cmod ((\i e" and "\x l. x \ S \ ((\n. f n x) sums l)" shows "\g. \x \ S. ((\n. f n x) sums g x) \ ((g has_field_derivative g' x) (at x within S))" proof - from assms obtain x l where x: "x \ S" and sf: "((\n. f n x) sums l)" by blast { fix e::real assume e: "e > 0" then obtain N where N: "\n x. n \ N \ x \ S \ cmod ((\i e" by (metis conv) have "\N. \n\N. \x\S. \h. cmod ((\i e * cmod h" proof (rule exI [of _ N], clarify) fix n y h assume "N \ n" "y \ S" then have "cmod ((\i e" by (metis N) then have "cmod h * cmod ((\i cmod h * e" by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) then show "cmod ((\i e * cmod h" by (simp add: norm_mult [symmetric] field_simps sum_distrib_left) qed } note ** = this show ?thesis unfolding has_field_derivative_def proof (rule has_derivative_series [OF cvs _ _ x]) fix n x assume "x \ S" then show "((f n) has_derivative (\z. z * f' n x)) (at x within S)" by (metis df has_field_derivative_def mult_commute_abs) next show " ((\n. f n x) sums l)" by (rule sf) next show "\e. e>0 \ \\<^sub>F n in sequentially. \x\S. \h. cmod ((\i e * cmod h" unfolding eventually_sequentially by (blast intro: **) qed qed subsection\<^marker>\tag unimportant\ \Taylor on Complex Numbers\ lemma sum_Suc_reindex: fixes f :: "nat \ 'a::ab_group_add" shows "sum f {0..n} = f 0 - f (Suc n) + sum (\i. f (Suc i)) {0..n}" by (induct n) auto lemma field_Taylor: assumes S: "convex S" and f: "\i x. x \ S \ i \ n \ (f i has_field_derivative f (Suc i) x) (at x within S)" and B: "\x. x \ S \ norm (f (Suc n) x) \ B" and w: "w \ S" and z: "z \ S" shows "norm(f 0 z - (\i\n. f i w * (z-w) ^ i / (fact i))) \ B * norm(z - w)^(Suc n) / fact n" proof - have wzs: "closed_segment w z \ S" using assms by (metis convex_contains_segment) { fix u assume "u \ closed_segment w z" then have "u \ S" by (metis wzs subsetD) have "(\i\n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) + f (Suc i) u * (z-u)^i / (fact i)) = f (Suc n) u * (z-u) ^ n / (fact n)" proof (induction n) case 0 show ?case by simp next case (Suc n) have "(\i\Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) + f (Suc i) u * (z-u) ^ i / (fact i)) = f (Suc n) u * (z-u) ^ n / (fact n) + f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) - f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))" using Suc by simp also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))" proof - have "(fact(Suc n)) * (f(Suc n) u *(z-u) ^ n / (fact n) + f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / (fact(Suc n)) - f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / (fact(Suc n))) = ((fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / (fact n) + ((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) - ((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))" by (simp add: algebra_simps del: fact_Suc) also have "... = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) + (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" by (simp del: fact_Suc) also have "... = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) + (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" by (simp only: fact_Suc of_nat_mult ac_simps) simp also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)" by (simp add: algebra_simps) finally show ?thesis by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact_Suc) qed finally show ?case . qed then have "((\v. (\i\n. f i v * (z - v)^i / (fact i))) has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n)) (at u within S)" apply (intro derivative_eq_intros) apply (blast intro: assms \u \ S\) apply (rule refl)+ apply (auto simp: field_simps) done } note sum_deriv = this { fix u assume u: "u \ closed_segment w z" then have us: "u \ S" by (metis wzs subsetD) have "norm (f (Suc n) u) * norm (z - u) ^ n \ norm (f (Suc n) u) * norm (u - z) ^ n" by (metis norm_minus_commute order_refl) also have "... \ norm (f (Suc n) u) * norm (z - w) ^ n" by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u]) also have "... \ B * norm (z - w) ^ n" by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us]) finally have "norm (f (Suc n) u) * norm (z - u) ^ n \ B * norm (z - w) ^ n" . } note cmod_bound = this have "(\i\n. f i z * (z - z) ^ i / (fact i)) = (\i\n. (f i z / (fact i)) * 0 ^ i)" by simp also have "\ = f 0 z / (fact 0)" by (subst sum_zero_power) simp finally have "norm (f 0 z - (\i\n. f i w * (z - w) ^ i / (fact i))) \ norm ((\i\n. f i w * (z - w) ^ i / (fact i)) - (\i\n. f i z * (z - z) ^ i / (fact i)))" by (simp add: norm_minus_commute) also have "... \ B * norm (z - w) ^ n / (fact n) * norm (w - z)" apply (rule field_differentiable_bound [where f' = "\w. f (Suc n) w * (z - w)^n / (fact n)" and S = "closed_segment w z", OF convex_closed_segment]) apply (auto simp: ends_in_segment DERIV_subset [OF sum_deriv wzs] norm_divide norm_mult norm_power divide_le_cancel cmod_bound) done also have "... \ B * norm (z - w) ^ Suc n / (fact n)" by (simp add: algebra_simps norm_minus_commute) finally show ?thesis . qed lemma complex_Taylor: assumes S: "convex S" and f: "\i x. x \ S \ i \ n \ (f i has_field_derivative f (Suc i) x) (at x within S)" and B: "\x. x \ S \ cmod (f (Suc n) x) \ B" and w: "w \ S" and z: "z \ S" shows "cmod(f 0 z - (\i\n. f i w * (z-w) ^ i / (fact i))) \ B * cmod(z - w)^(Suc n) / fact n" using assms by (rule field_Taylor) text\Something more like the traditional MVT for real components\ lemma complex_mvt_line: assumes "\u. u \ closed_segment w z \ (f has_field_derivative f'(u)) (at u)" shows "\u. u \ closed_segment w z \ Re(f z) - Re(f w) = Re(f'(u) * (z - w))" proof - have twz: "\t. (1 - t) *\<^sub>R w + t *\<^sub>R z = w + t *\<^sub>R (z - w)" by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) note assms[unfolded has_field_derivative_def, derivative_intros] show ?thesis apply (cut_tac mvt_simple [of 0 1 "Re o f o (\t. (1 - t) *\<^sub>R w + t *\<^sub>R z)" "\u. Re o (\h. f'((1 - u) *\<^sub>R w + u *\<^sub>R z) * h) o (\t. t *\<^sub>R (z - w))"]) apply auto apply (rule_tac x="(1 - x) *\<^sub>R w + x *\<^sub>R z" in exI) apply (auto simp: closed_segment_def twz) [] apply (intro derivative_eq_intros has_derivative_at_withinI, simp_all) apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib) apply (force simp: twz closed_segment_def) done qed lemma complex_Taylor_mvt: assumes "\i x. \x \ closed_segment w z; i \ n\ \ ((f i) has_field_derivative f (Suc i) x) (at x)" shows "\u. u \ closed_segment w z \ Re (f 0 z) = Re ((\i = 0..n. f i w * (z - w) ^ i / (fact i)) + (f (Suc n) u * (z-u)^n / (fact n)) * (z - w))" proof - { fix u assume u: "u \ closed_segment w z" have "(\i = 0..n. (f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = f (Suc 0) u - (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / (fact (Suc n)) + (\i = 0..n. (f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) / (fact (Suc i)))" by (subst sum_Suc_reindex) simp also have "... = f (Suc 0) u - (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / (fact (Suc n)) + (\i = 0..n. f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) - f (Suc i) u * (z-u) ^ i / (fact i))" by (simp only: diff_divide_distrib fact_cancel ac_simps) also have "... = f (Suc 0) u - (f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) / (fact (Suc n)) + f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u" by (subst sum_Suc_diff) auto also have "... = f (Suc n) u * (z-u) ^ n / (fact n)" by (simp only: algebra_simps diff_divide_distrib fact_cancel) finally have "(\i = 0..n. (f (Suc i) u * (z - u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = f (Suc n) u * (z - u) ^ n / (fact n)" . then have "((\u. \i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative f (Suc n) u * (z - u) ^ n / (fact n)) (at u)" apply (intro derivative_eq_intros)+ apply (force intro: u assms) apply (rule refl)+ apply (auto simp: ac_simps) done } then show ?thesis apply (cut_tac complex_mvt_line [of w z "\u. \i = 0..n. f i u * (z-u) ^ i / (fact i)" "\u. (f (Suc n) u * (z-u)^n / (fact n))"]) apply (auto simp add: intro: open_closed_segment) done qed end diff --git a/src/HOL/Analysis/Conformal_Mappings.thy b/src/HOL/Analysis/Conformal_Mappings.thy deleted file mode 100644 --- a/src/HOL/Analysis/Conformal_Mappings.thy +++ /dev/null @@ -1,5085 +0,0 @@ -section \Conformal Mappings and Consequences of Cauchy's Integral Theorem\ - -text\By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\ - -text\Also Cauchy's residue theorem by Wenda Li (2016)\ - -theory Conformal_Mappings -imports Cauchy_Integral_Theorem - -begin - -(* FIXME mv to Cauchy_Integral_Theorem.thy *) -subsection\Cauchy's inequality and more versions of Liouville\ - -lemma Cauchy_higher_deriv_bound: - assumes holf: "f holomorphic_on (ball z r)" - and contf: "continuous_on (cball z r) f" - and fin : "\w. w \ ball z r \ f w \ ball y B0" - and "0 < r" and "0 < n" - shows "norm ((deriv ^^ n) f z) \ (fact n) * B0 / r^n" -proof - - have "0 < B0" using \0 < r\ fin [of z] - by (metis ball_eq_empty ex_in_conv fin not_less) - have le_B0: "\w. cmod (w - z) \ r \ cmod (f w - y) \ B0" - apply (rule continuous_on_closure_norm_le [of "ball z r" "\w. f w - y"]) - apply (auto simp: \0 < r\ dist_norm norm_minus_commute) - apply (rule continuous_intros contf)+ - using fin apply (simp add: dist_commute dist_norm less_eq_real_def) - done - have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w) z - (deriv ^^ n) (\w. y) z" - using \0 < n\ by simp - also have "... = (deriv ^^ n) (\w. f w - y) z" - by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \0 < r\) - finally have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w - y) z" . - have contf': "continuous_on (cball z r) (\u. f u - y)" - by (rule contf continuous_intros)+ - have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" - by (simp add: holf holomorphic_on_diff) - define a where "a = (2 * pi)/(fact n)" - have "0 < a" by (simp add: a_def) - have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" - using \0 < r\ by (simp add: a_def field_split_simps) - have der_dif: "(deriv ^^ n) (\w. f w - y) z = (deriv ^^ n) f z" - using \0 < r\ \0 < n\ - by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) - have "norm ((2 * of_real pi * \)/(fact n) * (deriv ^^ n) (\w. f w - y) z) - \ (B0/r^(Suc n)) * (2 * pi * r)" - apply (rule has_contour_integral_bound_circlepath [of "(\u. (f u - y)/(u - z)^(Suc n))" _ z]) - using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] - using \0 < B0\ \0 < r\ - apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) - done - then show ?thesis - using \0 < r\ - by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) -qed - -lemma Cauchy_inequality: - assumes holf: "f holomorphic_on (ball \ r)" - and contf: "continuous_on (cball \ r) f" - and "0 < r" - and nof: "\x. norm(\-x) = r \ norm(f x) \ B" - shows "norm ((deriv ^^ n) f \) \ (fact n) * B / r^n" -proof - - obtain x where "norm (\-x) = r" - by (metis abs_of_nonneg add_diff_cancel_left' \0 < r\ diff_add_cancel - dual_order.strict_implies_order norm_of_real) - then have "0 \ B" - by (metis nof norm_not_less_zero not_le order_trans) - have "((\u. f u / (u - \) ^ Suc n) has_contour_integral (2 * pi) * \ / fact n * (deriv ^^ n) f \) - (circlepath \ r)" - apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) - using \0 < r\ by simp - then have "norm ((2 * pi * \)/(fact n) * (deriv ^^ n) f \) \ (B / r^(Suc n)) * (2 * pi * r)" - apply (rule has_contour_integral_bound_circlepath) - using \0 \ B\ \0 < r\ - apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) - done - then show ?thesis using \0 < r\ - by (simp add: norm_divide norm_mult field_simps) -qed - -lemma Liouville_polynomial: - assumes holf: "f holomorphic_on UNIV" - and nof: "\z. A \ norm z \ norm(f z) \ B * norm z ^ n" - shows "f \ = (\k\n. (deriv^^k) f 0 / fact k * \ ^ k)" -proof (cases rule: le_less_linear [THEN disjE]) - assume "B \ 0" - then have "\z. A \ norm z \ norm(f z) = 0" - by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) - then have f0: "(f \ 0) at_infinity" - using Lim_at_infinity by force - then have [simp]: "f = (\w. 0)" - using Liouville_weak [OF holf, of 0] - by (simp add: eventually_at_infinity f0) meson - show ?thesis by simp -next - assume "0 < B" - have "((\k. (deriv ^^ k) f 0 / (fact k) * (\ - 0)^k) sums f \)" - apply (rule holomorphic_power_series [where r = "norm \ + 1"]) - using holf holomorphic_on_subset apply auto - done - then have sumsf: "((\k. (deriv ^^ k) f 0 / (fact k) * \^k) sums f \)" by simp - have "(deriv ^^ k) f 0 / fact k * \ ^ k = 0" if "k>n" for k - proof (cases "(deriv ^^ k) f 0 = 0") - case True then show ?thesis by simp - next - case False - define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" - have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" - using \0 < B\ by simp - then have wge1: "1 \ norm w" - by (metis norm_of_real w_def) - then have "w \ 0" by auto - have kB: "0 < fact k * B" - using \0 < B\ by simp - then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" - by simp - then have wgeA: "A \ cmod w" - by (simp only: w_def norm_of_real) - have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" - using \0 < B\ by simp - then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" - by (metis norm_of_real w_def) - then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" - using False by (simp add: field_split_simps mult.commute split: if_split_asm) - also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" - apply (rule Cauchy_inequality) - using holf holomorphic_on_subset apply force - using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast - using \w \ 0\ apply simp - by (metis nof wgeA dist_0_norm dist_norm) - also have "... = fact k * (B * 1 / cmod w ^ (k-n))" - apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) - using \k>n\ \w \ 0\ \0 < B\ apply (simp add: field_split_simps semiring_normalization_rules) - done - also have "... = fact k * B / cmod w ^ (k-n)" - by simp - finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . - then have "1 / cmod w < 1 / cmod w ^ (k - n)" - by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) - then have "cmod w ^ (k - n) < cmod w" - by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) - with self_le_power [OF wge1] have False - by (meson diff_is_0_eq not_gr0 not_le that) - then show ?thesis by blast - qed - then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \ ^ (k + Suc n) = 0" for k - using not_less_eq by blast - then have "(\i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \ ^ (i + Suc n)) sums 0" - by (rule sums_0) - with sums_split_initial_segment [OF sumsf, where n = "Suc n"] - show ?thesis - using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce -qed - -text\Every bounded entire function is a constant function.\ -theorem Liouville_theorem: - assumes holf: "f holomorphic_on UNIV" - and bf: "bounded (range f)" - obtains c where "\z. f z = c" -proof - - obtain B where "\z. cmod (f z) \ B" - by (meson bf bounded_pos rangeI) - then show ?thesis - using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast -qed - -text\A holomorphic function f has only isolated zeros unless f is 0.\ - -lemma powser_0_nonzero: - fixes a :: "nat \ 'a::{real_normed_field,banach}" - assumes r: "0 < r" - and sm: "\x. norm (x - \) < r \ (\n. a n * (x - \) ^ n) sums (f x)" - and [simp]: "f \ = 0" - and m0: "a m \ 0" and "m>0" - obtains s where "0 < s" and "\z. z \ cball \ s - {\} \ f z \ 0" -proof - - have "r \ conv_radius a" - using sm sums_summable by (auto simp: le_conv_radius_iff [where \=\]) - obtain m where am: "a m \ 0" and az [simp]: "(\n. n a n = 0)" - apply (rule_tac m = "LEAST n. a n \ 0" in that) - using m0 - apply (rule LeastI2) - apply (fastforce intro: dest!: not_less_Least)+ - done - define b where "b i = a (i+m) / a m" for i - define g where "g x = suminf (\i. b i * (x - \) ^ i)" for x - have [simp]: "b 0 = 1" - by (simp add: am b_def) - { fix x::'a - assume "norm (x - \) < r" - then have "(\n. (a m * (x - \)^m) * (b n * (x - \)^n)) sums (f x)" - using am az sm sums_zero_iff_shift [of m "(\n. a n * (x - \) ^ n)" "f x"] - by (simp add: b_def monoid_mult_class.power_add algebra_simps) - then have "x \ \ \ (\n. b n * (x - \)^n) sums (f x / (a m * (x - \)^m))" - using am by (simp add: sums_mult_D) - } note bsums = this - then have "norm (x - \) < r \ summable (\n. b n * (x - \)^n)" for x - using sums_summable by (cases "x=\") auto - then have "r \ conv_radius b" - by (simp add: le_conv_radius_iff [where \=\]) - then have "r/2 < conv_radius b" - using not_le order_trans r by fastforce - then have "continuous_on (cball \ (r/2)) g" - using powser_continuous_suminf [of "r/2" b \] by (simp add: g_def) - then obtain s where "s>0" "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ dist (g x) (g \) < 1/2" - apply (rule continuous_onE [where x=\ and e = "1/2"]) - using r apply (auto simp: norm_minus_commute dist_norm) - done - moreover have "g \ = 1" - by (simp add: g_def) - ultimately have gnz: "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ (g x) \ 0" - by fastforce - have "f x \ 0" if "x \ \" "norm (x - \) \ s" "norm (x - \) \ r/2" for x - using bsums [of x] that gnz [of x] - apply (auto simp: g_def) - using r sums_iff by fastforce - then show ?thesis - apply (rule_tac s="min s (r/2)" in that) - using \0 < r\ \0 < s\ by (auto simp: dist_commute dist_norm) -qed - -subsection \Analytic continuation\ - -proposition isolated_zeros: - assumes holf: "f holomorphic_on S" - and "open S" "connected S" "\ \ S" "f \ = 0" "\ \ S" "f \ \ 0" - obtains r where "0 < r" and "ball \ r \ S" and - "\z. z \ ball \ r - {\} \ f z \ 0" -proof - - obtain r where "0 < r" and r: "ball \ r \ S" - using \open S\ \\ \ S\ open_contains_ball_eq by blast - have powf: "((\n. (deriv ^^ n) f \ / (fact n) * (z - \)^n) sums f z)" if "z \ ball \ r" for z - apply (rule holomorphic_power_series [OF _ that]) - apply (rule holomorphic_on_subset [OF holf r]) - done - obtain m where m: "(deriv ^^ m) f \ / (fact m) \ 0" - using holomorphic_fun_eq_0_on_connected [OF holf \open S\ \connected S\ _ \\ \ S\ \\ \ S\] \f \ \ 0\ - by auto - then have "m \ 0" using assms(5) funpow_0 by fastforce - obtain s where "0 < s" and s: "\z. z \ cball \ s - {\} \ f z \ 0" - apply (rule powser_0_nonzero [OF \0 < r\ powf \f \ = 0\ m]) - using \m \ 0\ by (auto simp: dist_commute dist_norm) - have "0 < min r s" by (simp add: \0 < r\ \0 < s\) - then show ?thesis - apply (rule that) - using r s by auto -qed - -proposition analytic_continuation: - assumes holf: "f holomorphic_on S" - and "open S" and "connected S" - and "U \ S" and "\ \ S" - and "\ islimpt U" - and fU0 [simp]: "\z. z \ U \ f z = 0" - and "w \ S" - shows "f w = 0" -proof - - obtain e where "0 < e" and e: "cball \ e \ S" - using \open S\ \\ \ S\ open_contains_cball_eq by blast - define T where "T = cball \ e \ U" - have contf: "continuous_on (closure T) f" - by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on - holomorphic_on_subset inf.cobounded1) - have fT0 [simp]: "\x. x \ T \ f x = 0" - by (simp add: T_def) - have "\r. \\e>0. \x'\U. x' \ \ \ dist x' \ < e; 0 < r\ \ \x'\cball \ e \ U. x' \ \ \ dist x' \ < r" - by (metis \0 < e\ IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj) - then have "\ islimpt T" using \\ islimpt U\ - by (auto simp: T_def islimpt_approachable) - then have "\ \ closure T" - by (simp add: closure_def) - then have "f \ = 0" - by (auto simp: continuous_constant_on_closure [OF contf]) - show ?thesis - apply (rule ccontr) - apply (rule isolated_zeros [OF holf \open S\ \connected S\ \\ \ S\ \f \ = 0\ \w \ S\], assumption) - by (metis open_ball \\ islimpt T\ centre_in_ball fT0 insertE insert_Diff islimptE) -qed - -corollary analytic_continuation_open: - assumes "open s" and "open s'" and "s \ {}" and "connected s'" - and "s \ s'" - assumes "f holomorphic_on s'" and "g holomorphic_on s'" - and "\z. z \ s \ f z = g z" - assumes "z \ s'" - shows "f z = g z" -proof - - from \s \ {}\ obtain \ where "\ \ s" by auto - with \open s\ have \: "\ islimpt s" - by (intro interior_limit_point) (auto simp: interior_open) - have "f z - g z = 0" - by (rule analytic_continuation[of "\z. f z - g z" s' s \]) - (insert assms \\ \ s\ \, auto intro: holomorphic_intros) - thus ?thesis by simp -qed - -subsection\Open mapping theorem\ - -lemma holomorphic_contract_to_zero: - assumes contf: "continuous_on (cball \ r) f" - and holf: "f holomorphic_on ball \ r" - and "0 < r" - and norm_less: "\z. norm(\ - z) = r \ norm(f \) < norm(f z)" - obtains z where "z \ ball \ r" "f z = 0" -proof - - { assume fnz: "\w. w \ ball \ r \ f w \ 0" - then have "0 < norm (f \)" - by (simp add: \0 < r\) - have fnz': "\w. w \ cball \ r \ f w \ 0" - by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) - have "frontier(cball \ r) \ {}" - using \0 < r\ by simp - define g where [abs_def]: "g z = inverse (f z)" for z - have contg: "continuous_on (cball \ r) g" - unfolding g_def using contf continuous_on_inverse fnz' by blast - have holg: "g holomorphic_on ball \ r" - unfolding g_def using fnz holf holomorphic_on_inverse by blast - have "frontier (cball \ r) \ cball \ r" - by (simp add: subset_iff) - then have contf': "continuous_on (frontier (cball \ r)) f" - and contg': "continuous_on (frontier (cball \ r)) g" - by (blast intro: contf contg continuous_on_subset)+ - have froc: "frontier(cball \ r) \ {}" - using \0 < r\ by simp - moreover have "continuous_on (frontier (cball \ r)) (norm o f)" - using contf' continuous_on_compose continuous_on_norm_id by blast - ultimately obtain w where w: "w \ frontier(cball \ r)" - and now: "\x. x \ frontier(cball \ r) \ norm (f w) \ norm (f x)" - apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]]) - apply simp - done - then have fw: "0 < norm (f w)" - by (simp add: fnz') - have "continuous_on (frontier (cball \ r)) (norm o g)" - using contg' continuous_on_compose continuous_on_norm_id by blast - then obtain v where v: "v \ frontier(cball \ r)" - and nov: "\x. x \ frontier(cball \ r) \ norm (g v) \ norm (g x)" - apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]]) - apply simp - done - then have fv: "0 < norm (f v)" - by (simp add: fnz') - have "norm ((deriv ^^ 0) g \) \ fact 0 * norm (g v) / r ^ 0" - by (rule Cauchy_inequality [OF holg contg \0 < r\]) (simp add: dist_norm nov) - then have "cmod (g \) \ norm (g v)" - by simp - with w have wr: "norm (\ - w) = r" and nfw: "norm (f w) \ norm (f \)" - apply (simp_all add: dist_norm) - by (metis \0 < cmod (f \)\ g_def less_imp_inverse_less norm_inverse not_le now order_trans v) - with fw have False - using norm_less by force - } - with that show ?thesis by blast -qed - -theorem open_mapping_thm: - assumes holf: "f holomorphic_on S" - and S: "open S" and "connected S" - and "open U" and "U \ S" - and fne: "\ f constant_on S" - shows "open (f ` U)" -proof - - have *: "open (f ` U)" - if "U \ {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\x. \y \ U. f y \ x" - for U - proof (clarsimp simp: open_contains_ball) - fix \ assume \: "\ \ U" - show "\e>0. ball (f \) e \ f ` U" - proof - - have hol: "(\z. f z - f \) holomorphic_on U" - by (rule holomorphic_intros that)+ - obtain s where "0 < s" and sbU: "ball \ s \ U" - and sne: "\z. z \ ball \ s - {\} \ (\z. f z - f \) z \ 0" - using isolated_zeros [OF hol U \] by (metis fneU right_minus_eq) - obtain r where "0 < r" and r: "cball \ r \ ball \ s" - apply (rule_tac r="s/2" in that) - using \0 < s\ by auto - have "cball \ r \ U" - using sbU r by blast - then have frsbU: "frontier (cball \ r) \ U" - using Diff_subset frontier_def order_trans by fastforce - then have cof: "compact (frontier(cball \ r))" - by blast - have frne: "frontier (cball \ r) \ {}" - using \0 < r\ by auto - have contfr: "continuous_on (frontier (cball \ r)) (\z. norm (f z - f \))" - by (metis continuous_on_norm continuous_on_subset frsbU hol holomorphic_on_imp_continuous_on) - obtain w where "norm (\ - w) = r" - and w: "(\z. norm (\ - z) = r \ norm (f w - f \) \ norm(f z - f \))" - apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) - apply (simp add: dist_norm) - done - moreover define \ where "\ \ norm (f w - f \) / 3" - ultimately have "0 < \" - using \0 < r\ dist_complex_def r sne by auto - have "ball (f \) \ \ f ` U" - proof - fix \ - assume \: "\ \ ball (f \) \" - have *: "cmod (\ - f \) < cmod (\ - f z)" if "cmod (\ - z) = r" for z - proof - - have lt: "cmod (f w - f \) / 3 < cmod (\ - f z)" - using w [OF that] \ - using dist_triangle2 [of "f \" "\" "f z"] dist_triangle2 [of "f \" "f z" \] - by (simp add: \_def dist_norm norm_minus_commute) - show ?thesis - by (metis \_def dist_commute dist_norm less_trans lt mem_ball \) - qed - have "continuous_on (cball \ r) (\z. \ - f z)" - apply (rule continuous_intros)+ - using \cball \ r \ U\ \f holomorphic_on U\ - apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) - done - moreover have "(\z. \ - f z) holomorphic_on ball \ r" - apply (rule holomorphic_intros)+ - apply (metis \cball \ r \ U\ \f holomorphic_on U\ holomorphic_on_subset interior_cball interior_subset) - done - ultimately obtain z where "z \ ball \ r" "\ - f z = 0" - apply (rule holomorphic_contract_to_zero) - apply (blast intro!: \0 < r\ *)+ - done - then show "\ \ f ` U" - using \cball \ r \ U\ by fastforce - qed - then show ?thesis using \0 < \\ by blast - qed - qed - have "open (f ` X)" if "X \ components U" for X - proof - - have holfU: "f holomorphic_on U" - using \U \ S\ holf holomorphic_on_subset by blast - have "X \ {}" - using that by (simp add: in_components_nonempty) - moreover have "open X" - using that \open U\ open_components by auto - moreover have "connected X" - using that in_components_maximal by blast - moreover have "f holomorphic_on X" - by (meson that holfU holomorphic_on_subset in_components_maximal) - moreover have "\y\X. f y \ x" for x - proof (rule ccontr) - assume not: "\ (\y\X. f y \ x)" - have "X \ S" - using \U \ S\ in_components_subset that by blast - obtain w where w: "w \ X" using \X \ {}\ by blast - have wis: "w islimpt X" - using w \open X\ interior_eq by auto - have hol: "(\z. f z - x) holomorphic_on S" - by (simp add: holf holomorphic_on_diff) - with fne [unfolded constant_on_def] - analytic_continuation[OF hol S \connected S\ \X \ S\ _ wis] not \X \ S\ w - show False by auto - qed - ultimately show ?thesis - by (rule *) - qed - then have "open (f ` \(components U))" - by (metis (no_types, lifting) imageE image_Union open_Union) - then show ?thesis - by force -qed - -text\No need for \<^term>\S\ to be connected. But the nonconstant condition is stronger.\ -corollary\<^marker>\tag unimportant\ open_mapping_thm2: - assumes holf: "f holomorphic_on S" - and S: "open S" - and "open U" "U \ S" - and fnc: "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" - shows "open (f ` U)" -proof - - have "S = \(components S)" by simp - with \U \ S\ have "U = (\C \ components S. C \ U)" by auto - then have "f ` U = (\C \ components S. f ` (C \ U))" - using image_UN by fastforce - moreover - { fix C assume "C \ components S" - with S \C \ components S\ open_components in_components_connected - have C: "open C" "connected C" by auto - have "C \ S" - by (metis \C \ components S\ in_components_maximal) - have nf: "\ f constant_on C" - apply (rule fnc) - using C \C \ S\ \C \ components S\ in_components_nonempty by auto - have "f holomorphic_on C" - by (metis holf holomorphic_on_subset \C \ S\) - then have "open (f ` (C \ U))" - apply (rule open_mapping_thm [OF _ C _ _ nf]) - apply (simp add: C \open U\ open_Int, blast) - done - } ultimately show ?thesis - by force -qed - -corollary\<^marker>\tag unimportant\ open_mapping_thm3: - assumes holf: "f holomorphic_on S" - and "open S" and injf: "inj_on f S" - shows "open (f ` S)" -apply (rule open_mapping_thm2 [OF holf]) -using assms -apply (simp_all add:) -using injective_not_constant subset_inj_on by blast - -subsection\Maximum modulus principle\ - -text\If \<^term>\f\ is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is - properly within the domain of \<^term>\f\.\ - -proposition maximum_modulus_principle: - assumes holf: "f holomorphic_on S" - and S: "open S" and "connected S" - and "open U" and "U \ S" and "\ \ U" - and no: "\z. z \ U \ norm(f z) \ norm(f \)" - shows "f constant_on S" -proof (rule ccontr) - assume "\ f constant_on S" - then have "open (f ` U)" - using open_mapping_thm assms by blast - moreover have "\ open (f ` U)" - proof - - have "\t. cmod (f \ - t) < e \ t \ f ` U" if "0 < e" for e - apply (rule_tac x="if 0 < Re(f \) then f \ + (e/2) else f \ - (e/2)" in exI) - using that - apply (simp add: dist_norm) - apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym) - done - then show ?thesis - unfolding open_contains_ball by (metis \\ \ U\ contra_subsetD dist_norm imageI mem_ball) - qed - ultimately show False - by blast -qed - -proposition maximum_modulus_frontier: - assumes holf: "f holomorphic_on (interior S)" - and contf: "continuous_on (closure S) f" - and bos: "bounded S" - and leB: "\z. z \ frontier S \ norm(f z) \ B" - and "\ \ S" - shows "norm(f \) \ B" -proof - - have "compact (closure S)" using bos - by (simp add: bounded_closure compact_eq_bounded_closed) - moreover have "continuous_on (closure S) (cmod \ f)" - using contf continuous_on_compose continuous_on_norm_id by blast - ultimately obtain z where zin: "z \ closure S" and z: "\y. y \ closure S \ (cmod \ f) y \ (cmod \ f) z" - using continuous_attains_sup [of "closure S" "norm o f"] \\ \ S\ by auto - then consider "z \ frontier S" | "z \ interior S" using frontier_def by auto - then have "norm(f z) \ B" - proof cases - case 1 then show ?thesis using leB by blast - next - case 2 - have zin: "z \ connected_component_set (interior S) z" - by (simp add: 2) - have "f constant_on (connected_component_set (interior S) z)" - apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin]) - apply (metis connected_component_subset holf holomorphic_on_subset) - apply (simp_all add: open_connected_component) - by (metis closure_subset comp_eq_dest_lhs interior_subset subsetCE z connected_component_in) - then obtain c where c: "\w. w \ connected_component_set (interior S) z \ f w = c" - by (auto simp: constant_on_def) - have "f ` closure(connected_component_set (interior S) z) \ {c}" - apply (rule image_closure_subset) - apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) - using c - apply auto - done - then have cc: "\w. w \ closure(connected_component_set (interior S) z) \ f w = c" by blast - have "frontier(connected_component_set (interior S) z) \ {}" - apply (simp add: frontier_eq_empty) - by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV) - then obtain w where w: "w \ frontier(connected_component_set (interior S) z)" - by auto - then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def) - also have "... \ B" - apply (rule leB) - using w -using frontier_interior_subset frontier_of_connected_component_subset by blast - finally show ?thesis . - qed - then show ?thesis - using z \\ \ S\ closure_subset by fastforce -qed - -corollary\<^marker>\tag unimportant\ maximum_real_frontier: - assumes holf: "f holomorphic_on (interior S)" - and contf: "continuous_on (closure S) f" - and bos: "bounded S" - and leB: "\z. z \ frontier S \ Re(f z) \ B" - and "\ \ S" - shows "Re(f \) \ B" -using maximum_modulus_frontier [of "exp o f" S "exp B"] - Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms -by auto - -subsection\<^marker>\tag unimportant\ \Factoring out a zero according to its order\ - -lemma holomorphic_factor_order_of_zero: - assumes holf: "f holomorphic_on S" - and os: "open S" - and "\ \ S" "0 < n" - and dnz: "(deriv ^^ n) f \ \ 0" - and dfz: "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" - obtains g r where "0 < r" - "g holomorphic_on ball \ r" - "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" - "\w. w \ ball \ r \ g w \ 0" -proof - - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) - then have holfb: "f holomorphic_on ball \ r" - using holf holomorphic_on_subset by blast - define g where "g w = suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" for w - have sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" - and feq: "f w - f \ = (w - \)^n * g w" - if w: "w \ ball \ r" for w - proof - - define powf where "powf = (\i. (deriv ^^ i) f \/(fact i) * (w - \)^i)" - have sing: "{.. = 0 then {} else {0})" - unfolding powf_def using \0 < n\ dfz by (auto simp: dfz; metis funpow_0 not_gr0) - have "powf sums f w" - unfolding powf_def by (rule holomorphic_power_series [OF holfb w]) - moreover have "(\i" - apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric]) - apply simp - apply (simp only: dfz sing) - apply (simp add: powf_def) - done - ultimately have fsums: "(\i. powf (i+n)) sums (f w - f \)" - using w sums_iff_shift' by metis - then have *: "summable (\i. (w - \) ^ n * ((deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n)))" - unfolding powf_def using sums_summable - by (auto simp: power_add mult_ac) - have "summable (\i. (deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n))" - proof (cases "w=\") - case False then show ?thesis - using summable_mult [OF *, of "1 / (w - \) ^ n"] by simp - next - case True then show ?thesis - by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"] - split: if_split_asm) - qed - then show sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" - by (simp add: summable_sums_iff g_def) - show "f w - f \ = (w - \)^n * g w" - apply (rule sums_unique2) - apply (rule fsums [unfolded powf_def]) - using sums_mult [OF sumsg, of "(w - \) ^ n"] - by (auto simp: power_add mult_ac) - qed - then have holg: "g holomorphic_on ball \ r" - by (meson sumsg power_series_holomorphic) - then have contg: "continuous_on (ball \ r) g" - by (blast intro: holomorphic_on_imp_continuous_on) - have "g \ \ 0" - using dnz unfolding g_def - by (subst suminf_finite [of "{0}"]) auto - obtain d where "0 < d" and d: "\w. w \ ball \ d \ g w \ 0" - apply (rule exE [OF continuous_on_avoid [OF contg _ \g \ \ 0\]]) - using \0 < r\ - apply force - by (metis \0 < r\ less_trans mem_ball not_less_iff_gr_or_eq) - show ?thesis - apply (rule that [where g=g and r ="min r d"]) - using \0 < r\ \0 < d\ holg - apply (auto simp: feq holomorphic_on_subset subset_ball d) - done -qed - - -lemma holomorphic_factor_order_of_zero_strong: - assumes holf: "f holomorphic_on S" "open S" "\ \ S" "0 < n" - and "(deriv ^^ n) f \ \ 0" - and "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" - obtains g r where "0 < r" - "g holomorphic_on ball \ r" - "\w. w \ ball \ r \ f w - f \ = ((w - \) * g w) ^ n" - "\w. w \ ball \ r \ g w \ 0" -proof - - obtain g r where "0 < r" - and holg: "g holomorphic_on ball \ r" - and feq: "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" - and gne: "\w. w \ ball \ r \ g w \ 0" - by (auto intro: holomorphic_factor_order_of_zero [OF assms]) - have con: "continuous_on (ball \ r) (\z. deriv g z / g z)" - by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on) - have cd: "\x. dist \ x < r \ (\z. deriv g z / g z) field_differentiable at x" - apply (rule derivative_intros)+ - using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) - apply (metis open_ball at_within_open holg holomorphic_on_def mem_ball) - using gne mem_ball by blast - obtain h where h: "\x. x \ ball \ r \ (h has_field_derivative deriv g x / g x) (at x)" - apply (rule exE [OF holomorphic_convex_primitive [of "ball \ r" "{}" "\z. deriv g z / g z"]]) - apply (auto simp: con cd) - apply (metis open_ball at_within_open mem_ball) - done - then have "continuous_on (ball \ r) h" - by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open) - then have con: "continuous_on (ball \ r) (\x. exp (h x) / g x)" - by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne) - have 0: "dist \ x < r \ ((\x. exp (h x) / g x) has_field_derivative 0) (at x)" for x - apply (rule h derivative_eq_intros | simp)+ - apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2]) - using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h) - done - obtain c where c: "\x. x \ ball \ r \ exp (h x) / g x = c" - by (rule DERIV_zero_connected_constant [of "ball \ r" "{}" "\x. exp(h x) / g x"]) (auto simp: con 0) - have hol: "(\z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \ r" - apply (rule holomorphic_on_compose [unfolded o_def, where g = exp]) - apply (rule holomorphic_intros)+ - using h holomorphic_on_open apply blast - apply (rule holomorphic_intros)+ - using \0 < n\ apply simp - apply (rule holomorphic_intros)+ - done - show ?thesis - apply (rule that [where g="\z. exp((Ln(inverse c) + h z)/n)" and r =r]) - using \0 < r\ \0 < n\ - apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric]) - apply (rule hol) - apply (simp add: Transcendental.exp_add gne) - done -qed - - -lemma - fixes k :: "'a::wellorder" - assumes a_def: "a == LEAST x. P x" and P: "P k" - shows def_LeastI: "P a" and def_Least_le: "a \ k" -unfolding a_def -by (rule LeastI Least_le; rule P)+ - -lemma holomorphic_factor_zero_nonconstant: - assumes holf: "f holomorphic_on S" and S: "open S" "connected S" - and "\ \ S" "f \ = 0" - and nonconst: "\ f constant_on S" - obtains g r n - where "0 < n" "0 < r" "ball \ r \ S" - "g holomorphic_on ball \ r" - "\w. w \ ball \ r \ f w = (w - \)^n * g w" - "\w. w \ ball \ r \ g w \ 0" -proof (cases "\n>0. (deriv ^^ n) f \ = 0") - case True then show ?thesis - using holomorphic_fun_eq_const_on_connected [OF holf S _ \\ \ S\] nonconst by (simp add: constant_on_def) -next - case False - then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \ \ 0" by blast - obtain r0 where "r0 > 0" "ball \ r0 \ S" using S openE \\ \ S\ by auto - define n where "n \ LEAST n. (deriv ^^ n) f \ \ 0" - have n_ne: "(deriv ^^ n) f \ \ 0" - by (rule def_LeastI [OF n_def]) (rule n0) - then have "0 < n" using \f \ = 0\ - using funpow_0 by fastforce - have n_min: "\k. k < n \ (deriv ^^ k) f \ = 0" - using def_Least_le [OF n_def] not_le by blast - then obtain g r1 - where "0 < r1" "g holomorphic_on ball \ r1" - "\w. w \ ball \ r1 \ f w = (w - \) ^ n * g w" - "\w. w \ ball \ r1 \ g w \ 0" - by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne] simp: \f \ = 0\) - then show ?thesis - apply (rule_tac g=g and r="min r0 r1" and n=n in that) - using \0 < n\ \0 < r0\ \0 < r1\ \ball \ r0 \ S\ - apply (auto simp: subset_ball intro: holomorphic_on_subset) - done -qed - - -lemma holomorphic_lower_bound_difference: - assumes holf: "f holomorphic_on S" and S: "open S" "connected S" - and "\ \ S" and "\ \ S" - and fne: "f \ \ f \" - obtains k n r - where "0 < k" "0 < r" - "ball \ r \ S" - "\w. w \ ball \ r \ k * norm(w - \)^n \ norm(f w - f \)" -proof - - define n where "n = (LEAST n. 0 < n \ (deriv ^^ n) f \ \ 0)" - obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \ \ 0" - using fne holomorphic_fun_eq_const_on_connected [OF holf S] \\ \ S\ \\ \ S\ by blast - then have "0 < n" and n_ne: "(deriv ^^ n) f \ \ 0" - unfolding n_def by (metis (mono_tags, lifting) LeastI)+ - have n_min: "\k. \0 < k; k < n\ \ (deriv ^^ k) f \ = 0" - unfolding n_def by (blast dest: not_less_Least) - then obtain g r - where "0 < r" and holg: "g holomorphic_on ball \ r" - and fne: "\w. w \ ball \ r \ f w - f \ = (w - \) ^ n * g w" - and gnz: "\w. w \ ball \ r \ g w \ 0" - by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne]) - obtain e where "e>0" and e: "ball \ e \ S" using assms by (blast elim!: openE) - then have holfb: "f holomorphic_on ball \ e" - using holf holomorphic_on_subset by blast - define d where "d = (min e r) / 2" - have "0 < d" using \0 < r\ \0 < e\ by (simp add: d_def) - have "d < r" - using \0 < r\ by (auto simp: d_def) - then have cbb: "cball \ d \ ball \ r" - by (auto simp: cball_subset_ball_iff) - then have "g holomorphic_on cball \ d" - by (rule holomorphic_on_subset [OF holg]) - then have "closed (g ` cball \ d)" - by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on) - moreover have "g ` cball \ d \ {}" - using \0 < d\ by auto - ultimately obtain x where x: "x \ g ` cball \ d" and "\y. y \ g ` cball \ d \ dist 0 x \ dist 0 y" - by (rule distance_attains_inf) blast - then have leg: "\w. w \ cball \ d \ norm x \ norm (g w)" - by auto - have "ball \ d \ cball \ d" by auto - also have "... \ ball \ e" using \0 < d\ d_def by auto - also have "... \ S" by (rule e) - finally have dS: "ball \ d \ S" . - moreover have "x \ 0" using gnz x \d < r\ by auto - ultimately show ?thesis - apply (rule_tac k="norm x" and n=n and r=d in that) - using \d < r\ leg - apply (auto simp: \0 < d\ fne norm_mult norm_power algebra_simps mult_right_mono) - done -qed - -lemma - assumes holf: "f holomorphic_on (S - {\})" and \: "\ \ interior S" - shows holomorphic_on_extend_lim: - "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ - ((\z. (z - \) * f z) \ 0) (at \)" - (is "?P = ?Q") - and holomorphic_on_extend_bounded: - "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ - (\B. eventually (\z. norm(f z) \ B) (at \))" - (is "?P = ?R") -proof - - obtain \ where "0 < \" and \: "ball \ \ \ S" - using \ mem_interior by blast - have "?R" if holg: "g holomorphic_on S" and gf: "\z. z \ S - {\} \ g z = f z" for g - proof - - have *: "\\<^sub>F z in at \. dist (g z) (g \) < 1 \ cmod (f z) \ cmod (g \) + 1" - apply (simp add: eventually_at) - apply (rule_tac x="\" in exI) - using \ \0 < \\ - apply (clarsimp simp:) - apply (drule_tac c=x in subsetD) - apply (simp add: dist_commute) - by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD) - have "continuous_on (interior S) g" - by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset) - then have "\x. x \ interior S \ (g \ g x) (at x)" - using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast - then have "(g \ g \) (at \)" - by (simp add: \) - then show ?thesis - apply (rule_tac x="norm(g \) + 1" in exI) - apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) - done - qed - moreover have "?Q" if "\\<^sub>F z in at \. cmod (f z) \ B" for B - by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) - moreover have "?P" if "(\z. (z - \) * f z) \\\ 0" - proof - - define h where [abs_def]: "h z = (z - \)^2 * f z" for z - have h0: "(h has_field_derivative 0) (at \)" - apply (simp add: h_def has_field_derivative_iff) - apply (rule Lim_transform_within [OF that, of 1]) - apply (auto simp: field_split_simps power2_eq_square) - done - have holh: "h holomorphic_on S" - proof (simp add: holomorphic_on_def, clarify) - fix z assume "z \ S" - show "h field_differentiable at z within S" - proof (cases "z = \") - case True then show ?thesis - using field_differentiable_at_within field_differentiable_def h0 by blast - next - case False - then have "f field_differentiable at z within S" - using holomorphic_onD [OF holf, of z] \z \ S\ - unfolding field_differentiable_def has_field_derivative_iff - by (force intro: exI [where x="dist \ z"] elim: Lim_transform_within_set [unfolded eventually_at]) - then show ?thesis - by (simp add: h_def power2_eq_square derivative_intros) - qed - qed - define g where [abs_def]: "g z = (if z = \ then deriv h \ else (h z - h \) / (z - \))" for z - have holg: "g holomorphic_on S" - unfolding g_def by (rule pole_lemma [OF holh \]) - show ?thesis - apply (rule_tac x="\z. if z = \ then deriv g \ else (g z - g \)/(z - \)" in exI) - apply (rule conjI) - apply (rule pole_lemma [OF holg \]) - apply (auto simp: g_def power2_eq_square divide_simps) - using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square) - done - qed - ultimately show "?P = ?Q" and "?P = ?R" - by meson+ -qed - -lemma pole_at_infinity: - assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \ l) at_infinity" - obtains a n where "\z. f z = (\i\n. a i * z^i)" -proof (cases "l = 0") - case False - with tendsto_inverse [OF lim] show ?thesis - apply (rule_tac a="(\n. inverse l)" and n=0 in that) - apply (simp add: Liouville_weak [OF holf, of "inverse l"]) - done -next - case True - then have [simp]: "l = 0" . - show ?thesis - proof (cases "\r. 0 < r \ (\z \ ball 0 r - {0}. f(inverse z) \ 0)") - case True - then obtain r where "0 < r" and r: "\z. z \ ball 0 r - {0} \ f(inverse z) \ 0" - by auto - have 1: "inverse \ f \ inverse holomorphic_on ball 0 r - {0}" - by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+ - have 2: "0 \ interior (ball 0 r)" - using \0 < r\ by simp - have "\B. 0 eventually (\z. cmod ((inverse \ f \ inverse) z) \ B) (at 0)" - apply (rule exI [where x=1]) - apply simp - using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] - apply (rule eventually_mono) - apply (simp add: dist_norm) - done - with holomorphic_on_extend_bounded [OF 1 2] - obtain g where holg: "g holomorphic_on ball 0 r" - and geq: "\z. z \ ball 0 r - {0} \ g z = (inverse \ f \ inverse) z" - by meson - have ifi0: "(inverse \ f \ inverse) \0\ 0" - using \l = 0\ lim lim_at_infinity_0 by blast - have g2g0: "g \0\ g 0" - using \0 < r\ centre_in_ball continuous_at continuous_on_eq_continuous_at holg - by (blast intro: holomorphic_on_imp_continuous_on) - have g2g1: "g \0\ 0" - apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) - using \0 < r\ by (auto simp: geq) - have [simp]: "g 0 = 0" - by (rule tendsto_unique [OF _ g2g0 g2g1]) simp - have "ball 0 r - {0::complex} \ {}" - using \0 < r\ - apply (clarsimp simp: ball_def dist_norm) - apply (drule_tac c="of_real r/2" in subsetD, auto) - done - then obtain w::complex where "w \ 0" and w: "norm w < r" by force - then have "g w \ 0" by (simp add: geq r) - obtain B n e where "0 < B" "0 < e" "e \ r" - and leg: "\w. norm w < e \ B * cmod w ^ n \ cmod (g w)" - apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w]) - using \0 < r\ w \g w \ 0\ by (auto simp: ball_subset_ball_iff) - have "cmod (f z) \ cmod z ^ n / B" if "2/e \ cmod z" for z - proof - - have ize: "inverse z \ ball 0 e - {0}" using that \0 < e\ - by (auto simp: norm_divide field_split_simps algebra_simps) - then have [simp]: "z \ 0" and izr: "inverse z \ ball 0 r - {0}" using \e \ r\ - by auto - then have [simp]: "f z \ 0" - using r [of "inverse z"] by simp - have [simp]: "f z = inverse (g (inverse z))" - using izr geq [of "inverse z"] by simp - show ?thesis using ize leg [of "inverse z"] \0 < B\ \0 < e\ - by (simp add: field_split_simps norm_divide algebra_simps) - qed - then show ?thesis - apply (rule_tac a = "\k. (deriv ^^ k) f 0 / (fact k)" and n=n in that) - apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp) - done - next - case False - then have fi0: "\r. r > 0 \ \z\ball 0 r - {0}. f (inverse z) = 0" - by simp - have fz0: "f z = 0" if "0 < r" and lt1: "\x. x \ 0 \ cmod x < r \ inverse (cmod (f (inverse x))) < 1" - for z r - proof - - have f0: "(f \ 0) at_infinity" - proof - - have DIM_complex[intro]: "2 \ DIM(complex)" \ \should not be necessary!\ - by simp - from lt1 have "f (inverse x) \ 0 \ x \ 0 \ cmod x < r \ 1 < cmod (f (inverse x))" for x - using one_less_inverse by force - then have **: "cmod (f (inverse x)) \ 1 \ x \ 0 \ cmod x < r \ f (inverse x) = 0" for x - by force - then have *: "(f \ inverse) ` (ball 0 r - {0}) \ {0} \ - ball 0 1" - by force - have "continuous_on (inverse ` (ball 0 r - {0})) f" - using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast - then have "connected ((f \ inverse) ` (ball 0 r - {0}))" - apply (intro connected_continuous_image continuous_intros) - apply (force intro: connected_punctured_ball)+ - done - then have "{0} \ (f \ inverse) ` (ball 0 r - {0}) = {} \ - ball 0 1 \ (f \ inverse) ` (ball 0 r - {0}) = {}" - by (rule connected_closedD) (use * in auto) - then have "w \ 0 \ cmod w < r \ f (inverse w) = 0" for w - using fi0 **[of w] \0 < r\ - apply (auto simp add: inf.commute [of "- ball 0 1"] Diff_eq [symmetric] image_subset_iff dest: less_imp_le) - apply fastforce - apply (drule bspec [of _ _ w]) - apply (auto dest: less_imp_le) - done - then show ?thesis - apply (simp add: lim_at_infinity_0) - apply (rule tendsto_eventually) - apply (simp add: eventually_at) - apply (rule_tac x=r in exI) - apply (simp add: \0 < r\ dist_norm) - done - qed - obtain w where "w \ ball 0 r - {0}" and "f (inverse w) = 0" - using False \0 < r\ by blast - then show ?thesis - by (auto simp: f0 Liouville_weak [OF holf, of 0]) - qed - show ?thesis - apply (rule that [of "\n. 0" 0]) - using lim [unfolded lim_at_infinity_0] - apply (simp add: Lim_at dist_norm norm_inverse) - apply (drule_tac x=1 in spec) - using fz0 apply auto - done - qed -qed - -subsection\<^marker>\tag unimportant\ \Entire proper functions are precisely the non-trivial polynomials\ - -lemma proper_map_polyfun: - fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" - assumes "closed S" and "compact K" and c: "c i \ 0" "1 \ i" "i \ n" - shows "compact (S \ {z. (\i\n. c i * z^i) \ K})" -proof - - obtain B where "B > 0" and B: "\x. x \ K \ norm x \ B" - by (metis compact_imp_bounded \compact K\ bounded_pos) - have *: "norm x \ b" - if "\x. b \ norm x \ B + 1 \ norm (\i\n. c i * x ^ i)" - "(\i\n. c i * x ^ i) \ K" for b x - proof - - have "norm (\i\n. c i * x ^ i) \ B" - using B that by blast - moreover have "\ B + 1 \ B" - by simp - ultimately show "norm x \ b" - using that by (metis (no_types) less_eq_real_def not_less order_trans) - qed - have "bounded {z. (\i\n. c i * z ^ i) \ K}" - using polyfun_extremal [where c=c and B="B+1", OF c] - by (auto simp: bounded_pos eventually_at_infinity_pos *) - moreover have "closed ((\z. (\i\n. c i * z ^ i)) -` K)" - apply (intro allI continuous_closed_vimage continuous_intros) - using \compact K\ compact_eq_bounded_closed by blast - ultimately show ?thesis - using closed_Int_compact [OF \closed S\] compact_eq_bounded_closed - by (auto simp add: vimage_def) -qed - -lemma proper_map_polyfun_univ: - fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" - assumes "compact K" "c i \ 0" "1 \ i" "i \ n" - shows "compact ({z. (\i\n. c i * z^i) \ K})" - using proper_map_polyfun [of UNIV K c i n] assms by simp - -lemma proper_map_polyfun_eq: - assumes "f holomorphic_on UNIV" - shows "(\k. compact k \ compact {z. f z \ k}) \ - (\c n. 0 < n \ (c n \ 0) \ f = (\z. \i\n. c i * z^i))" - (is "?lhs = ?rhs") -proof - assume compf [rule_format]: ?lhs - have 2: "\k. 0 < k \ a k \ 0 \ f = (\z. \i \ k. a i * z ^ i)" - if "\z. f z = (\i\n. a i * z ^ i)" for a n - proof (cases "\i\n. 0 a i = 0") - case True - then have [simp]: "\z. f z = a 0" - by (simp add: that sum.atMost_shift) - have False using compf [of "{a 0}"] by simp - then show ?thesis .. - next - case False - then obtain k where k: "0 < k" "k\n" "a k \ 0" by force - define m where "m = (GREATEST k. k\n \ a k \ 0)" - have m: "m\n \ a m \ 0" - unfolding m_def - apply (rule GreatestI_nat [where b = n]) - using k apply auto - done - have [simp]: "a i = 0" if "m < i" "i \ n" for i - using Greatest_le_nat [where b = "n" and P = "\k. k\n \ a k \ 0"] - using m_def not_le that by auto - have "k \ m" - unfolding m_def - apply (rule Greatest_le_nat [where b = "n"]) - using k apply auto - done - with k m show ?thesis - by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right) - qed - have "((inverse \ f) \ 0) at_infinity" - proof (rule Lim_at_infinityI) - fix e::real assume "0 < e" - with compf [of "cball 0 (inverse e)"] - show "\B. \x. B \ cmod x \ dist ((inverse \ f) x) 0 \ e" - apply simp - apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) - apply (rule_tac x="b+1" in exI) - apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one) - done - qed - then show ?rhs - apply (rule pole_at_infinity [OF assms]) - using 2 apply blast - done -next - assume ?rhs - then obtain c n where "0 < n" "c n \ 0" "f = (\z. \i\n. c i * z ^ i)" by blast - then have "compact {z. f z \ k}" if "compact k" for k - by (auto intro: proper_map_polyfun_univ [OF that]) - then show ?lhs by blast -qed - -subsection \Relating invertibility and nonvanishing of derivative\ - -lemma has_complex_derivative_locally_injective: - assumes holf: "f holomorphic_on S" - and S: "\ \ S" "open S" - and dnz: "deriv f \ \ 0" - obtains r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" -proof - - have *: "\d>0. \x. dist \ x < d \ onorm (\v. deriv f x * v - deriv f \ * v) < e" if "e > 0" for e - proof - - have contdf: "continuous_on S (deriv f)" - by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open S\) - obtain \ where "\>0" and \: "\x. \x \ S; dist x \ \ \\ \ cmod (deriv f x - deriv f \) \ e/2" - using continuous_onE [OF contdf \\ \ S\, of "e/2"] \0 < e\ - by (metis dist_complex_def half_gt_zero less_imp_le) - obtain \ where "\>0" "ball \ \ \ S" - by (metis openE [OF \open S\ \\ \ S\]) - with \\>0\ have "\\>0. \x. dist \ x < \ \ onorm (\v. deriv f x * v - deriv f \ * v) \ e/2" - apply (rule_tac x="min \ \" in exI) - apply (intro conjI allI impI Operator_Norm.onorm_le) - apply simp - apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult) - apply (rule mult_right_mono [OF \]) - apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \) - done - with \e>0\ show ?thesis by force - qed - have "inj ((*) (deriv f \))" - using dnz by simp - then obtain g' where g': "linear g'" "g' \ (*) (deriv f \) = id" - using linear_injective_left_inverse [of "(*) (deriv f \)"] - by (auto simp: linear_times) - show ?thesis - apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\z h. deriv f z * h" and g' = g']) - using g' * - apply (simp_all add: linear_conv_bounded_linear that) - using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf - holomorphic_on_imp_differentiable_at \open S\ apply blast - done -qed - -lemma has_complex_derivative_locally_invertible: - assumes holf: "f holomorphic_on S" - and S: "\ \ S" "open S" - and dnz: "deriv f \ \ 0" - obtains r where "r > 0" "ball \ r \ S" "open (f ` (ball \ r))" "inj_on f (ball \ r)" -proof - - obtain r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" - by (blast intro: that has_complex_derivative_locally_injective [OF assms]) - then have \: "\ \ ball \ r" by simp - then have nc: "\ f constant_on ball \ r" - using \inj_on f (ball \ r)\ injective_not_constant by fastforce - have holf': "f holomorphic_on ball \ r" - using \ball \ r \ S\ holf holomorphic_on_subset by blast - have "open (f ` ball \ r)" - apply (rule open_mapping_thm [OF holf']) - using nc apply auto - done - then show ?thesis - using \0 < r\ \ball \ r \ S\ \inj_on f (ball \ r)\ that by blast -qed - -lemma holomorphic_injective_imp_regular: - assumes holf: "f holomorphic_on S" - and "open S" and injf: "inj_on f S" - and "\ \ S" - shows "deriv f \ \ 0" -proof - - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) - have holf': "f holomorphic_on ball \ r" - using \ball \ r \ S\ holf holomorphic_on_subset by blast - show ?thesis - proof (cases "\n>0. (deriv ^^ n) f \ = 0") - case True - have fcon: "f w = f \" if "w \ ball \ r" for w - apply (rule holomorphic_fun_eq_const_on_connected [OF holf']) - using True \0 < r\ that by auto - have False - using fcon [of "\ + r/2"] \0 < r\ r injf unfolding inj_on_def - by (metis \\ \ S\ contra_subsetD dist_commute fcon mem_ball perfect_choose_dist) - then show ?thesis .. - next - case False - then obtain n0 where n0: "n0 > 0 \ (deriv ^^ n0) f \ \ 0" by blast - define n where [abs_def]: "n = (LEAST n. n > 0 \ (deriv ^^ n) f \ \ 0)" - have n_ne: "n > 0" "(deriv ^^ n) f \ \ 0" - using def_LeastI [OF n_def n0] by auto - have n_min: "\k. 0 < k \ k < n \ (deriv ^^ k) f \ = 0" - using def_Least_le [OF n_def] not_le by auto - obtain g \ where "0 < \" - and holg: "g holomorphic_on ball \ \" - and fd: "\w. w \ ball \ \ \ f w - f \ = ((w - \) * g w) ^ n" - and gnz: "\w. w \ ball \ \ \ g w \ 0" - apply (rule holomorphic_factor_order_of_zero_strong [OF holf \open S\ \\ \ S\ n_ne]) - apply (blast intro: n_min)+ - done - show ?thesis - proof (cases "n=1") - case True - with n_ne show ?thesis by auto - next - case False - have holgw: "(\w. (w - \) * g w) holomorphic_on ball \ (min r \)" - apply (rule holomorphic_intros)+ - using holg by (simp add: holomorphic_on_subset subset_ball) - have gd: "\w. dist \ w < \ \ (g has_field_derivative deriv g w) (at w)" - using holg - by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH) - have *: "\w. w \ ball \ (min r \) - \ ((\w. (w - \) * g w) has_field_derivative ((w - \) * deriv g w + g w)) - (at w)" - by (rule gd derivative_eq_intros | simp)+ - have [simp]: "deriv (\w. (w - \) * g w) \ \ 0" - using * [of \] \0 < \\ \0 < r\ by (simp add: DERIV_imp_deriv gnz) - obtain T where "\ \ T" "open T" and Tsb: "T \ ball \ (min r \)" and oimT: "open ((\w. (w - \) * g w) ` T)" - apply (rule has_complex_derivative_locally_invertible [OF holgw, of \]) - using \0 < r\ \0 < \\ - apply (simp_all add:) - by (meson open_ball centre_in_ball) - define U where "U = (\w. (w - \) * g w) ` T" - have "open U" by (metis oimT U_def) - have "0 \ U" - apply (auto simp: U_def) - apply (rule image_eqI [where x = \]) - apply (auto simp: \\ \ T\) - done - then obtain \ where "\>0" and \: "cball 0 \ \ U" - using \open U\ open_contains_cball by blast - then have "\ * exp(2 * of_real pi * \ * (0/n)) \ cball 0 \" - "\ * exp(2 * of_real pi * \ * (1/n)) \ cball 0 \" - by (auto simp: norm_mult) - with \ have "\ * exp(2 * of_real pi * \ * (0/n)) \ U" - "\ * exp(2 * of_real pi * \ * (1/n)) \ U" by blast+ - then obtain y0 y1 where "y0 \ T" and y0: "(y0 - \) * g y0 = \ * exp(2 * of_real pi * \ * (0/n))" - and "y1 \ T" and y1: "(y1 - \) * g y1 = \ * exp(2 * of_real pi * \ * (1/n))" - by (auto simp: U_def) - then have "y0 \ ball \ \" "y1 \ ball \ \" using Tsb by auto - moreover have "y0 \ y1" - using y0 y1 \\ > 0\ complex_root_unity_eq_1 [of n 1] \n > 0\ False by auto - moreover have "T \ S" - by (meson Tsb min.cobounded1 order_trans r subset_ball) - ultimately have False - using inj_onD [OF injf, of y0 y1] \y0 \ T\ \y1 \ T\ - using fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne - apply (simp add: y0 y1 power_mult_distrib) - apply (force simp: algebra_simps) - done - then show ?thesis .. - qed - qed -qed - -text\Hence a nice clean inverse function theorem\ - -lemma has_field_derivative_inverse_strong: - fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" - shows "\DERIV f x :> f'; f' \ 0; open S; x \ S; continuous_on S f; - \z. z \ S \ g (f z) = z\ - \ DERIV g (f x) :> inverse (f')" - unfolding has_field_derivative_def - by (rule has_derivative_inverse_strong [of S x f g]) auto - -lemma has_field_derivative_inverse_strong_x: - fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" - shows "\DERIV f (g y) :> f'; f' \ 0; open S; continuous_on S f; g y \ S; f(g y) = y; - \z. z \ S \ g (f z) = z\ - \ DERIV g y :> inverse (f')" - unfolding has_field_derivative_def - by (rule has_derivative_inverse_strong_x [of S g y f]) auto - -proposition holomorphic_has_inverse: - assumes holf: "f holomorphic_on S" - and "open S" and injf: "inj_on f S" - obtains g where "g holomorphic_on (f ` S)" - "\z. z \ S \ deriv f z * deriv g (f z) = 1" - "\z. z \ S \ g(f z) = z" -proof - - have ofs: "open (f ` S)" - by (rule open_mapping_thm3 [OF assms]) - have contf: "continuous_on S f" - by (simp add: holf holomorphic_on_imp_continuous_on) - have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \ S" for z - proof - - have 1: "(f has_field_derivative deriv f z) (at z)" - using DERIV_deriv_iff_field_differentiable \z \ S\ \open S\ holf holomorphic_on_imp_differentiable_at - by blast - have 2: "deriv f z \ 0" - using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast - show ?thesis - apply (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) - apply (simp add: holf holomorphic_on_imp_continuous_on) - by (simp add: injf the_inv_into_f_f) - qed - show ?thesis - proof - show "the_inv_into S f holomorphic_on f ` S" - by (simp add: holomorphic_on_open ofs) (blast intro: *) - next - fix z assume "z \ S" - have "deriv f z \ 0" - using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast - then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" - using * [OF \z \ S\] by (simp add: DERIV_imp_deriv) - next - fix z assume "z \ S" - show "the_inv_into S f (f z) = z" - by (simp add: \z \ S\ injf the_inv_into_f_f) - qed -qed - -subsection\The Schwarz Lemma\ - -lemma Schwarz1: - assumes holf: "f holomorphic_on S" - and contf: "continuous_on (closure S) f" - and S: "open S" "connected S" - and boS: "bounded S" - and "S \ {}" - obtains w where "w \ frontier S" - "\z. z \ closure S \ norm (f z) \ norm (f w)" -proof - - have connf: "continuous_on (closure S) (norm o f)" - using contf continuous_on_compose continuous_on_norm_id by blast - have coc: "compact (closure S)" - by (simp add: \bounded S\ bounded_closure compact_eq_bounded_closed) - then obtain x where x: "x \ closure S" and xmax: "\z. z \ closure S \ norm(f z) \ norm(f x)" - apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]]) - using \S \ {}\ apply auto - done - then show ?thesis - proof (cases "x \ frontier S") - case True - then show ?thesis using that xmax by blast - next - case False - then have "x \ S" - using \open S\ frontier_def interior_eq x by auto - then have "f constant_on S" - apply (rule maximum_modulus_principle [OF holf S \open S\ order_refl]) - using closure_subset apply (blast intro: xmax) - done - then have "f constant_on (closure S)" - by (rule constant_on_closureI [OF _ contf]) - then obtain c where c: "\x. x \ closure S \ f x = c" - by (meson constant_on_def) - obtain w where "w \ frontier S" - by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) - then show ?thesis - by (simp add: c frontier_def that) - qed -qed - -lemma Schwarz2: - "\f holomorphic_on ball 0 r; - 0 < s; ball w s \ ball 0 r; - \z. norm (w-z) < s \ norm(f z) \ norm(f w)\ - \ f constant_on ball 0 r" -by (rule maximum_modulus_principle [where U = "ball w s" and \ = w]) (simp_all add: dist_norm) - -lemma Schwarz3: - assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" - obtains h where "h holomorphic_on (ball 0 r)" and "\z. norm z < r \ f z = z * (h z)" and "deriv f 0 = h 0" -proof - - define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z - have d0: "deriv f 0 = h 0" - by (simp add: h_def) - moreover have "h holomorphic_on (ball 0 r)" - by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) - moreover have "norm z < r \ f z = z * h z" for z - by (simp add: h_def) - ultimately show ?thesis - using that by blast -qed - -proposition Schwarz_Lemma: - assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" - and no: "\z. norm z < 1 \ norm (f z) < 1" - and \: "norm \ < 1" - shows "norm (f \) \ norm \" and "norm(deriv f 0) \ 1" - and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) - \ norm(deriv f 0) = 1) - \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" - (is "?P \ ?Q") -proof - - obtain h where holh: "h holomorphic_on (ball 0 1)" - and fz_eq: "\z. norm z < 1 \ f z = z * (h z)" and df0: "deriv f 0 = h 0" - by (rule Schwarz3 [OF holf]) auto - have noh_le: "norm (h z) \ 1" if z: "norm z < 1" for z - proof - - have "norm (h z) < a" if a: "1 < a" for a - proof - - have "max (inverse a) (norm z) < 1" - using z a by (simp_all add: inverse_less_1_iff) - then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" - using Rats_dense_in_real by blast - then have nzr: "norm z < r" and ira: "inverse r < a" - using z a less_imp_inverse_less by force+ - then have "0 < r" - by (meson norm_not_less_zero not_le order.strict_trans2) - have holh': "h holomorphic_on ball 0 r" - by (meson holh \r < 1\ holomorphic_on_subset less_eq_real_def subset_ball) - have conth': "continuous_on (cball 0 r) h" - by (meson \r < 1\ dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) - obtain w where w: "norm w = r" and lenw: "\z. norm z < r \ norm(h z) \ norm(h w)" - apply (rule Schwarz1 [OF holh']) using conth' \0 < r\ by auto - have "h w = f w / w" using fz_eq \r < 1\ nzr w by auto - then have "cmod (h z) < inverse r" - by (metis \0 < r\ \r < 1\ divide_strict_right_mono inverse_eq_divide - le_less_trans lenw no norm_divide nzr w) - then show ?thesis using ira by linarith - qed - then show "norm (h z) \ 1" - using not_le by blast - qed - show "cmod (f \) \ cmod \" - proof (cases "\ = 0") - case True then show ?thesis by auto - next - case False - then show ?thesis - by (simp add: noh_le fz_eq \ mult_left_le norm_mult) - qed - show no_df0: "norm(deriv f 0) \ 1" - by (simp add: \\z. cmod z < 1 \ cmod (h z) \ 1\ df0) - show "?Q" if "?P" - using that - proof - assume "\z. cmod z < 1 \ z \ 0 \ cmod (f z) = cmod z" - then obtain \ where \: "cmod \ < 1" "\ \ 0" "cmod (f \) = cmod \" by blast - then have [simp]: "norm (h \) = 1" - by (simp add: fz_eq norm_mult) - have "ball \ (1 - cmod \) \ ball 0 1" - by (simp add: ball_subset_ball_iff) - moreover have "\z. cmod (\ - z) < 1 - cmod \ \ cmod (h z) \ cmod (h \)" - apply (simp add: algebra_simps) - by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) - ultimately obtain c where c: "\z. norm z < 1 \ h z = c" - using Schwarz2 [OF holh, of "1 - norm \" \, unfolded constant_on_def] \ by auto - then have "norm c = 1" - using \ by force - with c show ?thesis - using fz_eq by auto - next - assume [simp]: "cmod (deriv f 0) = 1" - then obtain c where c: "\z. norm z < 1 \ h z = c" - using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le - by auto - moreover have "norm c = 1" using df0 c by auto - ultimately show ?thesis - using fz_eq by auto - qed -qed - -corollary Schwarz_Lemma': - assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" - and no: "\z. norm z < 1 \ norm (f z) < 1" - shows "((\\. norm \ < 1 \ norm (f \) \ norm \) - \ norm(deriv f 0) \ 1) - \ (((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) - \ norm(deriv f 0) = 1) - \ (\\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1))" - using Schwarz_Lemma [OF assms] - by (metis (no_types) norm_eq_zero zero_less_one) - -subsection\The Schwarz reflection principle\ - -lemma hol_pal_lem0: - assumes "d \ a \ k" "k \ d \ b" - obtains c where - "c \ closed_segment a b" "d \ c = k" - "\z. z \ closed_segment a c \ d \ z \ k" - "\z. z \ closed_segment c b \ k \ d \ z" -proof - - obtain c where cin: "c \ closed_segment a b" and keq: "k = d \ c" - using connected_ivt_hyperplane [of "closed_segment a b" a b d k] - by (auto simp: assms) - have "closed_segment a c \ {z. d \ z \ k}" "closed_segment c b \ {z. k \ d \ z}" - unfolding segment_convex_hull using assms keq - by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) - then show ?thesis using cin that by fastforce -qed - -lemma hol_pal_lem1: - assumes "convex S" "open S" - and abc: "a \ S" "b \ S" "c \ S" - "d \ 0" and lek: "d \ a \ k" "d \ b \ k" "d \ c \ k" - and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" - and contf: "continuous_on S f" - shows "contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof - - have "interior (convex hull {a, b, c}) \ interior(S \ {x. d \ x \ k})" - apply (rule interior_mono) - apply (rule hull_minimal) - apply (simp add: abc lek) - apply (rule convex_Int [OF \convex S\ convex_halfspace_le]) - done - also have "... \ {z \ S. d \ z < k}" - by (force simp: interior_open [OF \open S\] \d \ 0\) - finally have *: "interior (convex hull {a, b, c}) \ {z \ S. d \ z < k}" . - have "continuous_on (convex hull {a,b,c}) f" - using \convex S\ contf abc continuous_on_subset subset_hull - by fastforce - moreover have "f holomorphic_on interior (convex hull {a,b,c})" - by (rule holomorphic_on_subset [OF holf1 *]) - ultimately show ?thesis - using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 - by blast -qed - -lemma hol_pal_lem2: - assumes S: "convex S" "open S" - and abc: "a \ S" "b \ S" "c \ S" - and "d \ 0" and lek: "d \ a \ k" "d \ b \ k" - and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" - and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" - and contf: "continuous_on S f" - shows "contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof (cases "d \ c \ k") - case True show ?thesis - by (rule hol_pal_lem1 [OF S abc \d \ 0\ lek True holf1 contf]) -next - case False - then have "d \ c > k" by force - obtain a' where a': "a' \ closed_segment b c" and "d \ a' = k" - and ba': "\z. z \ closed_segment b a' \ d \ z \ k" - and a'c: "\z. z \ closed_segment a' c \ k \ d \ z" - apply (rule hol_pal_lem0 [of d b k c, OF \d \ b \ k\]) - using False by auto - obtain b' where b': "b' \ closed_segment a c" and "d \ b' = k" - and ab': "\z. z \ closed_segment a b' \ d \ z \ k" - and b'c: "\z. z \ closed_segment b' c \ k \ d \ z" - apply (rule hol_pal_lem0 [of d a k c, OF \d \ a \ k\]) - using False by auto - have a'b': "a' \ S \ b' \ S" - using a' abc b' convex_contains_segment \convex S\ by auto - have "continuous_on (closed_segment c a) f" - by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) - then have 1: "contour_integral (linepath c a) f = - contour_integral (linepath c b') f + contour_integral (linepath b' a) f" - apply (rule contour_integral_split_linepath) - using b' by (simp add: closed_segment_commute) - have "continuous_on (closed_segment b c) f" - by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) - then have 2: "contour_integral (linepath b c) f = - contour_integral (linepath b a') f + contour_integral (linepath a' c) f" - by (rule contour_integral_split_linepath [OF _ a']) - have 3: "contour_integral (reversepath (linepath b' a')) f = - - contour_integral (linepath b' a') f" - by (rule contour_integral_reversepath [OF valid_path_linepath]) - have fcd_le: "f field_differentiable at x" - if "x \ interior S \ x \ interior {x. d \ x \ k}" for x - proof - - have "f holomorphic_on S \ {c. d \ c < k}" - by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) - then have "\C D. x \ interior C \ interior D \ f holomorphic_on interior C \ interior D" - using that - by (metis Collect_mem_eq Int_Collect \d \ 0\ interior_halfspace_le interior_open \open S\) - then show "f field_differentiable at x" - by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) - qed - have ab_le: "\x. x \ closed_segment a b \ d \ x \ k" - proof - - fix x :: complex - assume "x \ closed_segment a b" - then have "\C. x \ C \ b \ C \ a \ C \ \ convex C" - by (meson contra_subsetD convex_contains_segment) - then show "d \ x \ k" - by (metis lek convex_halfspace_le mem_Collect_eq) - qed - have "continuous_on (S \ {x. d \ x \ k}) f" using contf - by (simp add: continuous_on_subset) - then have "(f has_contour_integral 0) - (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" - apply (rule Cauchy_theorem_convex [where K = "{}"]) - apply (simp_all add: path_image_join convex_Int convex_halfspace_le \convex S\ fcd_le ab_le - closed_segment_subset abc a'b' ba') - by (metis \d \ a' = k\ \d \ b' = k\ convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl) - then have 4: "contour_integral (linepath a b) f + - contour_integral (linepath b a') f + - contour_integral (linepath a' b') f + - contour_integral (linepath b' a) f = 0" - by (rule has_chain_integral_chain_integral4) - have fcd_ge: "f field_differentiable at x" - if "x \ interior S \ x \ interior {x. k \ d \ x}" for x - proof - - have f2: "f holomorphic_on S \ {c. k < d \ c}" - by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) - have f3: "interior S = S" - by (simp add: interior_open \open S\) - then have "x \ S \ interior {c. k \ d \ c}" - using that by simp - then show "f field_differentiable at x" - using f3 f2 unfolding holomorphic_on_def - by (metis (no_types) \d \ 0\ at_within_interior interior_Int interior_halfspace_ge interior_interior) - qed - have "continuous_on (S \ {x. k \ d \ x}) f" using contf - by (simp add: continuous_on_subset) - then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" - apply (rule Cauchy_theorem_convex [where K = "{}"]) - apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \convex S\ - fcd_ge closed_segment_subset abc a'b' a'c) - by (metis \d \ a' = k\ b'c closed_segment_commute convex_contains_segment - convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl) - then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" - by (rule has_chain_integral_chain_integral3) - show ?thesis - using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) -qed - -lemma hol_pal_lem3: - assumes S: "convex S" "open S" - and abc: "a \ S" "b \ S" "c \ S" - and "d \ 0" and lek: "d \ a \ k" - and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" - and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" - and contf: "continuous_on S f" - shows "contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof (cases "d \ b \ k") - case True show ?thesis - by (rule hol_pal_lem2 [OF S abc \d \ 0\ lek True holf1 holf2 contf]) -next - case False - show ?thesis - proof (cases "d \ c \ k") - case True - have "contour_integral (linepath c a) f + - contour_integral (linepath a b) f + - contour_integral (linepath b c) f = 0" - by (rule hol_pal_lem2 [OF S \c \ S\ \a \ S\ \b \ S\ \d \ 0\ \d \ c \ k\ lek holf1 holf2 contf]) - then show ?thesis - by (simp add: algebra_simps) - next - case False - have "contour_integral (linepath b c) f + - contour_integral (linepath c a) f + - contour_integral (linepath a b) f = 0" - apply (rule hol_pal_lem2 [OF S \b \ S\ \c \ S\ \a \ S\, of "-d" "-k"]) - using \d \ 0\ \\ d \ b \ k\ False by (simp_all add: holf1 holf2 contf) - then show ?thesis - by (simp add: algebra_simps) - qed -qed - -lemma hol_pal_lem4: - assumes S: "convex S" "open S" - and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" - and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" - and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" - and contf: "continuous_on S f" - shows "contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0" -proof (cases "d \ a \ k") - case True show ?thesis - by (rule hol_pal_lem3 [OF S abc \d \ 0\ True holf1 holf2 contf]) -next - case False - show ?thesis - apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"]) - using \d \ 0\ False by (simp_all add: holf1 holf2 contf) -qed - -lemma holomorphic_on_paste_across_line: - assumes S: "open S" and "d \ 0" - and holf1: "f holomorphic_on (S \ {z. d \ z < k})" - and holf2: "f holomorphic_on (S \ {z. k < d \ z})" - and contf: "continuous_on S f" - shows "f holomorphic_on S" -proof - - have *: "\t. open t \ p \ t \ continuous_on t f \ - (\a b c. convex hull {a, b, c} \ t \ - contour_integral (linepath a b) f + - contour_integral (linepath b c) f + - contour_integral (linepath c a) f = 0)" - if "p \ S" for p - proof - - obtain e where "e>0" and e: "ball p e \ S" - using \p \ S\ openE S by blast - then have "continuous_on (ball p e) f" - using contf continuous_on_subset by blast - moreover have "f holomorphic_on {z. dist p z < e \ d \ z < k}" - apply (rule holomorphic_on_subset [OF holf1]) - using e by auto - moreover have "f holomorphic_on {z. dist p z < e \ k < d \ z}" - apply (rule holomorphic_on_subset [OF holf2]) - using e by auto - ultimately show ?thesis - apply (rule_tac x="ball p e" in exI) - using \e > 0\ e \d \ 0\ - apply (simp add:, clarify) - apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k]) - apply (auto simp: subset_hull) - done - qed - show ?thesis - by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) -qed - -proposition Schwarz_reflection: - assumes "open S" and cnjs: "cnj ` S \ S" - and holf: "f holomorphic_on (S \ {z. 0 < Im z})" - and contf: "continuous_on (S \ {z. 0 \ Im z}) f" - and f: "\z. \z \ S; z \ \\ \ (f z) \ \" - shows "(\z. if 0 \ Im z then f z else cnj(f(cnj z))) holomorphic_on S" -proof - - have 1: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. 0 < Im z})" - by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) - have cont_cfc: "continuous_on (S \ {z. Im z \ 0}) (cnj o f o cnj)" - apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) - using cnjs apply auto - done - have "cnj \ f \ cnj field_differentiable at x within S \ {z. Im z < 0}" - if "x \ S" "Im x < 0" "f field_differentiable at (cnj x) within S \ {z. 0 < Im z}" for x - using that - apply (simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm, clarify) - apply (rule_tac x="cnj f'" in exI) - apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) - apply (drule_tac x="cnj xa" in bspec) - using cnjs apply force - apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) - done - then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \ {z. Im z < 0})" - using holf cnjs - by (force simp: holomorphic_on_def) - have 2: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. Im z < 0})" - apply (rule iffD1 [OF holomorphic_cong [OF refl]]) - using hol_cfc by auto - have [simp]: "(S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0}) = S" - by force - have "continuous_on ((S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0})) - (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" - apply (rule continuous_on_cases_local) - using cont_cfc contf - apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) - using f Reals_cnj_iff complex_is_Real_iff apply auto - done - then have 3: "continuous_on S (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" - by force - show ?thesis - apply (rule holomorphic_on_paste_across_line [OF \open S\, of "- \" _ 0]) - using 1 2 3 - apply auto - done -qed - -subsection\Bloch's theorem\ - -lemma Bloch_lemma_0: - assumes holf: "f holomorphic_on cball 0 r" and "0 < r" - and [simp]: "f 0 = 0" - and le: "\z. norm z < r \ norm(deriv f z) \ 2 * norm(deriv f 0)" - shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \ f ` ball 0 r" -proof - - have "sqrt 2 < 3/2" - by (rule real_less_lsqrt) (auto simp: power2_eq_square) - then have sq3: "0 < 3 - 2 * sqrt 2" by simp - show ?thesis - proof (cases "deriv f 0 = 0") - case True then show ?thesis by simp - next - case False - define C where "C = 2 * norm(deriv f 0)" - have "0 < C" using False by (simp add: C_def) - have holf': "f holomorphic_on ball 0 r" using holf - using ball_subset_cball holomorphic_on_subset by blast - then have holdf': "deriv f holomorphic_on ball 0 r" - by (rule holomorphic_deriv [OF _ open_ball]) - have "Le1": "norm(deriv f z - deriv f 0) \ norm z / (r - norm z) * C" - if "norm z < r" for z - proof - - have T1: "norm(deriv f z - deriv f 0) \ norm z / (R - norm z) * C" - if R: "norm z < R" "R < r" for R - proof - - have "0 < R" using R - by (metis less_trans norm_zero zero_less_norm_iff) - have df_le: "\x. norm x < r \ norm (deriv f x) \ C" - using le by (simp add: C_def) - have hol_df: "deriv f holomorphic_on cball 0 R" - apply (rule holomorphic_on_subset) using R holdf' by auto - have *: "((\w. deriv f w / (w - z)) has_contour_integral 2 * pi * \ * deriv f z) (circlepath 0 R)" - if "norm z < R" for z - using \0 < R\ that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] - by (force simp: winding_number_circlepath) - have **: "((\x. deriv f x / (x - z) - deriv f x / x) has_contour_integral - of_real (2 * pi) * \ * (deriv f z - deriv f 0)) - (circlepath 0 R)" - using has_contour_integral_diff [OF * [of z] * [of 0]] \0 < R\ that - by (simp add: algebra_simps) - have [simp]: "\x. norm x = R \ x \ z" using that(1) by blast - have "norm (deriv f x / (x - z) - deriv f x / x) - \ C * norm z / (R * (R - norm z))" - if "norm x = R" for x - proof - - have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = - norm (deriv f x) * norm z" - by (simp add: norm_mult right_diff_distrib') - show ?thesis - using \0 < R\ \0 < C\ R that - apply (simp add: norm_mult norm_divide divide_simps) - using df_le norm_triangle_ineq2 \0 < C\ apply (auto intro!: mult_mono) - done - qed - then show ?thesis - using has_contour_integral_bound_circlepath - [OF **, of "C * norm z/(R*(R - norm z))"] - \0 < R\ \0 < C\ R - apply (simp add: norm_mult norm_divide) - apply (simp add: divide_simps mult.commute) - done - qed - obtain r' where r': "norm z < r'" "r' < r" - using Rats_dense_in_real [of "norm z" r] \norm z < r\ by blast - then have [simp]: "closure {r'<.. norm(f z)" - if r: "norm z < r" for z - proof - - have 1: "\x. x \ ball 0 r \ - ((\z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) - (at x within ball 0 r)" - by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ - have 2: "closed_segment 0 z \ ball 0 r" - by (metis \0 < r\ convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) - have 3: "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" - apply (rule integrable_on_cmult_right [where 'b=real, simplified]) - apply (rule integrable_on_cdivide [where 'b=real, simplified]) - apply (rule integrable_on_cmult_left [where 'b=real, simplified]) - apply (rule ident_integrable_on) - done - have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm z * norm z * x * C / (r - norm z)" - if x: "0 \ x" "x \ 1" for x - proof - - have [simp]: "x * norm z < r" - using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) - have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \ norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" - apply (rule Le1) using r x \0 < r\ by simp - also have "... \ norm (x *\<^sub>R z) / (r - norm z) * C" - using r x \0 < r\ - apply (simp add: field_split_simps) - by (simp add: \0 < C\ mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) - finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm (x *\<^sub>R z) / (r - norm z) * C * norm z" - by (rule mult_right_mono) simp - with x show ?thesis by (simp add: algebra_simps) - qed - have le_norm: "abc \ norm d - e \ norm(f - d) \ e \ abc \ norm f" for abc d e and f::complex - by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) - have "norm (integral {0..1} (\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) - \ integral {0..1} (\t. (norm z)\<^sup>2 * t / (r - norm z) * C)" - apply (rule integral_norm_bound_integral) - using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 - apply (simp add: has_contour_integral_linepath has_integral_integrable_integral) - apply (rule 3) - apply (simp add: norm_mult power2_eq_square 4) - done - then have int_le: "norm (f z - deriv f 0 * z) \ (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" - using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 - apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) - done - show ?thesis - apply (rule le_norm [OF _ int_le]) - using \norm z < r\ - apply (simp add: power2_eq_square divide_simps C_def norm_mult) - proof - - have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" - by (simp add: algebra_simps) - then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" - by (simp add: algebra_simps) - qed - qed - have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" - by (auto simp: sqrt2_less_2) - have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" - apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) - apply (subst closure_ball) - using \0 < r\ mult_pos_pos sq201 - apply (auto simp: cball_subset_cball_iff) - done - have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" - apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force) - using \0 < r\ mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff) - using False \0 < r\ centre_in_ball holf' holomorphic_nonconstant by blast - have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = - ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" - by simp - also have "... \ f ` ball 0 ((1 - sqrt 2 / 2) * r)" - proof - - have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \ norm (f z)" - if "norm z = (1 - sqrt 2 / 2) * r" for z - apply (rule order_trans [OF _ *]) - using \0 < r\ - apply (simp_all add: field_simps power2_eq_square that) - apply (simp add: mult.assoc [symmetric]) - done - show ?thesis - apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) - using \0 < r\ sq201 3 apply simp_all - using C_def \0 < C\ sq3 apply force - done - qed - also have "... \ f ` ball 0 r" - apply (rule image_subsetI [OF imageI], simp) - apply (erule less_le_trans) - using \0 < r\ apply (auto simp: field_simps) - done - finally show ?thesis . - qed -qed - -lemma Bloch_lemma: - assumes holf: "f holomorphic_on cball a r" and "0 < r" - and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" - shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" -proof - - have fz: "(\z. f (a + z)) = f o (\z. (a + z))" - by (simp add: o_def) - have hol0: "(\z. f (a + z)) holomorphic_on cball 0 r" - unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ - then have [simp]: "\x. norm x < r \ (\z. f (a + z)) field_differentiable at x" - by (metis open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) - have [simp]: "\z. norm z < r \ f field_differentiable at (a + z)" - by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute) - then have [simp]: "f field_differentiable at a" - by (metis add.comm_neutral \0 < r\ norm_eq_zero) - have hol1: "(\z. f (a + z) - f a) holomorphic_on cball 0 r" - by (intro holomorphic_intros hol0) - then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\z. f (a + z) - f a) 0)) - \ (\z. f (a + z) - f a) ` ball 0 r" - apply (rule Bloch_lemma_0) - apply (simp_all add: \0 < r\) - apply (simp add: fz complex_derivative_chain) - apply (simp add: dist_norm le) - done - then show ?thesis - apply clarify - apply (drule_tac c="x - f a" in subsetD) - apply (force simp: fz \0 < r\ dist_norm complex_derivative_chain field_differentiable_compose)+ - done -qed - -proposition Bloch_unit: - assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" - obtains b r where "1/12 < r" and "ball b r \ f ` (ball a 1)" -proof - - define r :: real where "r = 249/256" - have "0 < r" "r < 1" by (auto simp: r_def) - define g where "g z = deriv f z * of_real(r - norm(z - a))" for z - have "deriv f holomorphic_on ball a 1" - by (rule holomorphic_deriv [OF holf open_ball]) - then have "continuous_on (ball a 1) (deriv f)" - using holomorphic_on_imp_continuous_on by blast - then have "continuous_on (cball a r) (deriv f)" - by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \r < 1\) - then have "continuous_on (cball a r) g" - by (simp add: g_def continuous_intros) - then have 1: "compact (g ` cball a r)" - by (rule compact_continuous_image [OF _ compact_cball]) - have 2: "g ` cball a r \ {}" - using \r > 0\ by auto - obtain p where pr: "p \ cball a r" - and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" - using distance_attains_sup [OF 1 2, of 0] by force - define t where "t = (r - norm(p - a)) / 2" - have "norm (p - a) \ r" - using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) - then have "norm (p - a) < r" using pr - by (simp add: norm_minus_commute dist_norm) - then have "0 < t" - by (simp add: t_def) - have cpt: "cball p t \ ball a r" - using \0 < t\ by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) - have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" - if "y \ cball a r" for y - proof - - have [simp]: "norm (y - a) \ r" - using that by (simp add: dist_norm norm_minus_commute) - have "norm (g y) \ norm (g p)" - using pge [OF that] by simp - then have "norm (deriv f y) * abs (r - norm (y - a)) \ norm (deriv f p) * abs (r - norm (p - a))" - by (simp only: dist_norm g_def norm_mult norm_of_real) - with that \norm (p - a) < r\ show ?thesis - by (simp add: dist_norm field_split_simps) - qed - have le_norm_dfp: "r / (r - norm (p - a)) \ norm (deriv f p)" - using gen_le_dfp [of a] \r > 0\ by auto - have 1: "f holomorphic_on cball p t" - apply (rule holomorphic_on_subset [OF holf]) - using cpt \r < 1\ order_subst1 subset_ball by auto - have 2: "norm (deriv f z) \ 2 * norm (deriv f p)" if "z \ ball p t" for z - proof - - have z: "z \ cball a r" - by (meson ball_subset_cball subsetD cpt that) - then have "norm(z - a) < r" - by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) - have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" - using gen_le_dfp [OF z] by simp - with \norm (z - a) < r\ \norm (p - a) < r\ - have "norm (deriv f z) \ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" - by (simp add: field_simps) - also have "... \ 2 * norm (deriv f p)" - apply (rule mult_right_mono) - using that \norm (p - a) < r\ \norm(z - a) < r\ - apply (simp_all add: field_simps t_def dist_norm [symmetric]) - using dist_triangle3 [of z a p] by linarith - finally show ?thesis . - qed - have sqrt2: "sqrt 2 < 2113/1494" - by (rule real_less_lsqrt) (auto simp: power2_eq_square) - then have sq3: "0 < 3 - 2 * sqrt 2" by simp - have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" - using sq3 sqrt2 by (auto simp: field_simps r_def) - also have "... \ cmod (deriv f p) * (r - cmod (p - a))" - using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) - finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" - using pos_divide_less_eq half_gt_zero_iff sq3 by blast - then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" - using sq3 by (simp add: mult.commute t_def) - have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball p t" - by (rule Bloch_lemma [OF 1 \0 < t\ 2]) - also have "... \ f ` ball a 1" - apply (rule image_mono) - apply (rule order_trans [OF ball_subset_cball]) - apply (rule order_trans [OF cpt]) - using \0 < t\ \r < 1\ apply (simp add: ball_subset_ball_iff dist_norm) - done - finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball a 1" . - with ** show ?thesis - by (rule that) -qed - -theorem Bloch: - assumes holf: "f holomorphic_on ball a r" and "0 < r" - and r': "r' \ r * norm (deriv f a) / 12" - obtains b where "ball b r' \ f ` (ball a r)" -proof (cases "deriv f a = 0") - case True with r' show ?thesis - using ball_eq_empty that by fastforce -next - case False - define C where "C = deriv f a" - have "0 < norm C" using False by (simp add: C_def) - have dfa: "f field_differentiable at a" - apply (rule holomorphic_on_imp_differentiable_at [OF holf]) - using \0 < r\ by auto - have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" - by (simp add: o_def) - have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" - apply (rule holomorphic_on_subset [OF holf]) - using \0 < r\ apply (force simp: dist_norm norm_mult) - done - have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" - apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ - using \0 < r\ by (simp add: C_def False) - have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative - (deriv f (a + of_real r * z) / C)) (at z)" - if "norm z < 1" for z - proof - - have *: "((\x. f (a + of_real r * x)) has_field_derivative - (deriv f (a + of_real r * z) * of_real r)) (at z)" - apply (simp add: fo) - apply (rule DERIV_chain [OF field_differentiable_derivI]) - apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) - using \0 < r\ apply (simp add: dist_norm norm_mult that) - apply (rule derivative_eq_intros | simp)+ - done - show ?thesis - apply (rule derivative_eq_intros * | simp)+ - using \0 < r\ by (auto simp: C_def False) - qed - have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" - apply (subst deriv_cdivide_right) - apply (simp add: field_differentiable_def fo) - apply (rule exI) - apply (rule DERIV_chain [OF field_differentiable_derivI]) - apply (simp add: dfa) - apply (rule derivative_eq_intros | simp add: C_def False fo)+ - using \0 < r\ - apply (simp add: C_def False fo) - apply (simp add: derivative_intros dfa complex_derivative_chain) - done - have sb1: "(*) (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 - \ f ` ball a r" - using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) - have sb2: "ball (C * r * b) r' \ (*) (C * r) ` ball b t" - if "1 / 12 < t" for b t - proof - - have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" - using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right - by auto - show ?thesis - apply clarify - apply (rule_tac x="x / (C * r)" in image_eqI) - using \0 < r\ - apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) - apply (erule less_le_trans) - apply (rule order_trans [OF r' *]) - done - qed - show ?thesis - apply (rule Bloch_unit [OF 1 2]) - apply (rename_tac t) - apply (rule_tac b="(C * of_real r) * b" in that) - apply (drule image_mono [where f = "\z. (C * of_real r) * z"]) - using sb1 sb2 - apply force - done -qed - -corollary Bloch_general: - assumes holf: "f holomorphic_on s" and "a \ s" - and tle: "\z. z \ frontier s \ t \ dist a z" - and rle: "r \ t * norm(deriv f a) / 12" - obtains b where "ball b r \ f ` s" -proof - - consider "r \ 0" | "0 < t * norm(deriv f a) / 12" using rle by force - then show ?thesis - proof cases - case 1 then show ?thesis - by (simp add: ball_empty that) - next - case 2 - show ?thesis - proof (cases "deriv f a = 0") - case True then show ?thesis - using rle by (simp add: ball_empty that) - next - case False - then have "t > 0" - using 2 by (force simp: zero_less_mult_iff) - have "\ ball a t \ s \ ball a t \ frontier s \ {}" - apply (rule connected_Int_frontier [of "ball a t" s], simp_all) - using \0 < t\ \a \ s\ centre_in_ball apply blast - done - with tle have *: "ball a t \ s" by fastforce - then have 1: "f holomorphic_on ball a t" - using holf using holomorphic_on_subset by blast - show ?thesis - apply (rule Bloch [OF 1 \t > 0\ rle]) - apply (rule_tac b=b in that) - using * apply force - done - qed - qed -qed - -subsection \Cauchy's residue theorem\ - -text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. - Interactive Theorem Proving\ - -definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where - "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" - -lemma Eps_cong: - assumes "\x. P x = Q x" - shows "Eps P = Eps Q" - using ext[of P Q, OF assms] by simp - -lemma residue_cong: - assumes eq: "eventually (\z. f z = g z) (at z)" and "z = z'" - shows "residue f z = residue g z'" -proof - - from assms have eq': "eventually (\z. g z = f z) (at z)" - by (simp add: eq_commute) - let ?P = "\f c e. (\\>0. \ < e \ - (f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \))" - have "residue f z = residue g z" unfolding residue_def - proof (rule Eps_cong) - fix c :: complex - have "\e>0. ?P g c e" - if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g - proof - - from that(1) obtain e where e: "e > 0" "?P f c e" - by blast - from that(2) obtain e' where e': "e' > 0" "\z'. z' \ z \ dist z' z < e' \ f z' = g z'" - unfolding eventually_at by blast - have "?P g c (min e e')" - proof (intro allI exI impI, goal_cases) - case (1 \) - hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" - using e(2) by auto - thus ?case - proof (rule has_contour_integral_eq) - fix z' assume "z' \ path_image (circlepath z \)" - hence "dist z' z < e'" and "z' \ z" - using 1 by (auto simp: dist_commute) - with e'(2)[of z'] show "f z' = g z'" by simp - qed - qed - moreover from e and e' have "min e e' > 0" by auto - ultimately show ?thesis by blast - qed - from this[OF _ eq] and this[OF _ eq'] - show "(\e>0. ?P f c e) \ (\e>0. ?P g c e)" - by blast - qed - with assms show ?thesis by simp -qed - -lemma contour_integral_circlepath_eq: - assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" - and e2_cball:"cball z e2 \ s" - shows - "f contour_integrable_on circlepath z e1" - "f contour_integrable_on circlepath z e2" - "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" -proof - - define l where "l \ linepath (z+e2) (z+e1)" - have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto - have "e2>0" using \e1>0\ \e1\e2\ by auto - have zl_img:"z\path_image l" - proof - assume "z \ path_image l" - then have "e2 \ cmod (e2 - e1)" - using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def - by (auto simp add:closed_segment_commute) - thus False using \e2>0\ \e1>0\ \e1\e2\ - apply (subst (asm) norm_of_real) - by auto - qed - define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" - show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" - proof - - show "f contour_integrable_on circlepath z e2" - apply (intro contour_integrable_continuous_circlepath[OF - continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) - using \e2>0\ e2_cball by auto - show "f contour_integrable_on (circlepath z e1)" - apply (intro contour_integrable_continuous_circlepath[OF - continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) - using \e1>0\ \e1\e2\ e2_cball by auto - qed - have [simp]:"f contour_integrable_on l" - proof - - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ - by (intro closed_segment_subset,auto simp add:dist_norm) - hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def - by auto - then show "f contour_integrable_on l" unfolding l_def - apply (intro contour_integrable_continuous_linepath[OF - continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) - by auto - qed - let ?ig="\g. contour_integral g f" - have "(f has_contour_integral 0) g" - proof (rule Cauchy_theorem_global[OF _ f_holo]) - show "open (s - {z})" using \open s\ by auto - show "valid_path g" unfolding g_def l_def by auto - show "pathfinish g = pathstart g" unfolding g_def l_def by auto - next - have path_img:"path_image g \ cball z e2" - proof - - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ - by (intro closed_segment_subset,auto simp add:dist_norm) - moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto - ultimately show ?thesis unfolding g_def l_def using \e2>0\ - by (simp add: path_image_join closed_segment_commute) - qed - show "path_image g \ s - {z}" - proof - - have "z\path_image g" using zl_img - unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) - moreover note \cball z e2 \ s\ and path_img - ultimately show ?thesis by auto - qed - show "winding_number g w = 0" when"w \ s - {z}" for w - proof - - have "winding_number g w = 0" when "w\s" using that e2_cball - apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) - by (auto simp add:g_def l_def) - moreover have "winding_number g z=0" - proof - - let ?Wz="\g. winding_number g z" - have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) - + ?Wz (reversepath l)" - using \e2>0\ \e1>0\ zl_img unfolding g_def l_def - by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ - also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" - using zl_img - apply (subst (2) winding_number_reversepath) - by (auto simp add:l_def closed_segment_commute) - also have "... = 0" - proof - - have "?Wz (circlepath z e2) = 1" using \e2>0\ - by (auto intro: winding_number_circlepath_centre) - moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \e1>0\ - apply (subst winding_number_reversepath) - by (auto intro: winding_number_circlepath_centre) - ultimately show ?thesis by auto - qed - finally show ?thesis . - qed - ultimately show ?thesis using that by auto - qed - qed - then have "0 = ?ig g" using contour_integral_unique by simp - also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) - + ?ig (reversepath l)" - unfolding g_def - by (auto simp add:contour_integrable_reversepath_eq) - also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" - by (auto simp add:contour_integral_reversepath) - finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" - by simp -qed - -lemma base_residue: - assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" - and r_cball:"cball z r \ s" - shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" -proof - - obtain e where "e>0" and e_cball:"cball z e \ s" - using open_contains_cball[of s] \open s\ \z\s\ by auto - define c where "c \ 2 * pi * \" - define i where "i \ contour_integral (circlepath z e) f / c" - have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ - proof - - have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" - "f contour_integrable_on circlepath z \" - "f contour_integrable_on circlepath z e" - using \\ - by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ - then show ?thesis unfolding i_def c_def - by (auto intro:has_contour_integral_integral) - qed - then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" - unfolding residue_def c_def - apply (rule_tac someI[of _ i],intro exI[where x=e]) - by (auto simp add:\e>0\ c_def) - then obtain e' where "e'>0" - and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" - by auto - let ?int="\e. contour_integral (circlepath z e) f" - define \ where "\ \ Min {r,e'} / 2" - have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto - have "(f has_contour_integral c * (residue f z)) (circlepath z \)" - using e'_def[rule_format,OF \\>0\ \\] . - then show ?thesis unfolding c_def - using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] - by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) -qed - -lemma residue_holo: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s" - shows "residue f z = 0" -proof - - define c where "c \ 2 * pi * \" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ - using open_contains_cball_eq by blast - have "(f has_contour_integral c*residue f z) (circlepath z e)" - using f_holo - by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) - moreover have "(f has_contour_integral 0) (circlepath z e)" - using f_holo e_cball \e>0\ - by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) - ultimately have "c*residue f z =0" - using has_contour_integral_unique by blast - thus ?thesis unfolding c_def by auto -qed - -lemma residue_const:"residue (\_. c) z = 0" - by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) - -lemma residue_add: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" - and g_holo:"g holomorphic_on s - {z}" - shows "residue (\z. f z + g z) z= residue f z + residue g z" -proof - - define c where "c \ 2 * pi * \" - define fg where "fg \ (\z. f z+g z)" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ - using open_contains_cball_eq by blast - have "(fg has_contour_integral c * residue fg z) (circlepath z e)" - unfolding fg_def using f_holo g_holo - apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) - by (auto intro:holomorphic_intros) - moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" - unfolding fg_def using f_holo g_holo - by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) - ultimately have "c*(residue f z + residue g z) = c * residue fg z" - using has_contour_integral_unique by (auto simp add:distrib_left) - thus ?thesis unfolding fg_def - by (auto simp add:c_def) -qed - -lemma residue_lmul: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" - shows "residue (\z. c * (f z)) z= c * residue f z" -proof (cases "c=0") - case True - thus ?thesis using residue_const by auto -next - case False - define c' where "c' \ 2 * pi * \" - define f' where "f' \ (\z. c * (f z))" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ - using open_contains_cball_eq by blast - have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" - unfolding f'_def using f_holo - apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) - by (auto intro:holomorphic_intros) - moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" - unfolding f'_def using f_holo - by (auto intro: has_contour_integral_lmul - base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) - ultimately have "c' * residue f' z = c * (c' * residue f z)" - using has_contour_integral_unique by auto - thus ?thesis unfolding f'_def c'_def using False - by (auto simp add:field_simps) -qed - -lemma residue_rmul: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" - shows "residue (\z. (f z) * c) z= residue f z * c" -using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) - -lemma residue_div: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" - shows "residue (\z. (f z) / c) z= residue f z / c " -using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) - -lemma residue_neg: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" - shows "residue (\z. - (f z)) z= - residue f z" -using residue_lmul[OF assms,of "-1"] by auto - -lemma residue_diff: - assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" - and g_holo:"g holomorphic_on s - {z}" - shows "residue (\z. f z - g z) z= residue f z - residue g z" -using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] -by (auto intro:holomorphic_intros g_holo) - -lemma residue_simple: - assumes "open s" "z\s" and f_holo:"f holomorphic_on s" - shows "residue (\w. f w / (w - z)) z = f z" -proof - - define c where "c \ 2 * pi * \" - define f' where "f' \ \w. f w / (w - z)" - obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ - using open_contains_cball_eq by blast - have "(f' has_contour_integral c * f z) (circlepath z e)" - unfolding f'_def c_def using \e>0\ f_holo e_cball - by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) - moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" - unfolding f'_def using f_holo - apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) - by (auto intro!:holomorphic_intros) - ultimately have "c * f z = c * residue f' z" - using has_contour_integral_unique by blast - thus ?thesis unfolding c_def f'_def by auto -qed - -lemma residue_simple': - assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" - and lim: "((\w. f w * (w - z)) \ c) (at z)" - shows "residue f z = c" -proof - - define g where "g = (\w. if w = z then c else f w * (w - z))" - from holo have "(\w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P") - by (force intro: holomorphic_intros) - also have "?P \ g holomorphic_on (s - {z})" - by (intro holomorphic_cong refl) (simp_all add: g_def) - finally have *: "g holomorphic_on (s - {z})" . - - note lim - also have "(\w. f w * (w - z)) \z\ c \ g \z\ g z" - by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter) - finally have **: "g \z\ g z" . - - have g_holo: "g holomorphic_on s" - by (rule no_isolated_singularity'[where K = "{z}"]) - (insert assms * **, simp_all add: at_within_open_NO_MATCH) - from s and this have "residue (\w. g w / (w - z)) z = g z" - by (rule residue_simple) - also have "\\<^sub>F za in at z. g za / (za - z) = f za" - unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def) - hence "residue (\w. g w / (w - z)) z = residue f z" - by (intro residue_cong refl) - finally show ?thesis - by (simp add: g_def) -qed - -lemma residue_holomorphic_over_power: - assumes "open A" "z0 \ A" "f holomorphic_on A" - shows "residue (\z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n" -proof - - let ?f = "\z. f z / (z - z0) ^ Suc n" - from assms(1,2) obtain r where r: "r > 0" "cball z0 r \ A" - by (auto simp: open_contains_cball) - have "(?f has_contour_integral 2 * pi * \ * residue ?f z0) (circlepath z0 r)" - using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros) - moreover have "(?f has_contour_integral 2 * pi * \ / fact n * (deriv ^^ n) f z0) (circlepath z0 r)" - using assms r - by (intro Cauchy_has_contour_integral_higher_derivative_circlepath) - (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on) - ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" - by (rule has_contour_integral_unique) - thus ?thesis by (simp add: field_simps) -qed - -lemma residue_holomorphic_over_power': - assumes "open A" "0 \ A" "f holomorphic_on A" - shows "residue (\z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" - using residue_holomorphic_over_power[OF assms] by simp - -lemma get_integrable_path: - assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\s-pts" "b\s-pts" - obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" - "path_image g \ s-pts" "f contour_integrable_on g" using assms -proof (induct arbitrary:s thesis a rule:finite_induct[OF \finite pts\]) - case 1 - obtain g where "valid_path g" "path_image g \ s" "pathstart g = a" "pathfinish g = b" - using connected_open_polynomial_connected[OF \open s\,of a b ] \connected (s - {})\ - valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto - moreover have "f contour_integrable_on g" - using contour_integrable_holomorphic_simple[OF _ \open s\ \valid_path g\ \path_image g \ s\,of f] - \f holomorphic_on s - {}\ - by auto - ultimately show ?case using "1"(1)[of g] by auto -next - case idt:(2 p pts) - obtain e where "e>0" and e:"\w\ball a e. w \ s \ (w \ a \ w \ insert p pts)" - using finite_ball_avoid[OF \open s\ \finite (insert p pts)\, of a] - \a \ s - insert p pts\ - by auto - define a' where "a' \ a+e/2" - have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ - by (auto simp add:dist_complex_def a'_def) - then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" - "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" - using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) - by (metis Diff_insert2 open_delete) - define g where "g \ linepath a a' +++ g'" - have "valid_path g" unfolding g_def by (auto intro: valid_path_join) - moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto - moreover have "path_image g \ s - insert p pts" unfolding g_def - proof (rule subset_path_image_join) - have "closed_segment a a' \ ball a e" using \e>0\ - by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) - then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) - by auto - next - show "path_image g' \ s - insert p pts" using g'(4) by blast - qed - moreover have "f contour_integrable_on g" - proof - - have "closed_segment a a' \ ball a e" using \e>0\ - by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) - then have "continuous_on (closed_segment a a') f" - using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] - apply (elim continuous_on_subset) - by auto - then have "f contour_integrable_on linepath a a'" - using contour_integrable_continuous_linepath by auto - then show ?thesis unfolding g_def - apply (rule contour_integrable_joinI) - by (auto simp add: \e>0\) - qed - ultimately show ?case using idt.prems(1)[of g] by auto -qed - -lemma Cauchy_theorem_aux: - assumes "open s" "connected (s-pts)" "finite pts" "pts \ s" "f holomorphic_on s-pts" - "valid_path g" "pathfinish g = pathstart g" "path_image g \ s-pts" - "\z. (z \ s) \ winding_number g z = 0" - "\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" - shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" - using assms -proof (induct arbitrary:s g rule:finite_induct[OF \finite pts\]) - case 1 - then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) -next - case (2 p pts) - note fin[simp] = \finite (insert p pts)\ - and connected = \connected (s - insert p pts)\ - and valid[simp] = \valid_path g\ - and g_loop[simp] = \pathfinish g = pathstart g\ - and holo[simp]= \f holomorphic_on s - insert p pts\ - and path_img = \path_image g \ s - insert p pts\ - and winding = \\z. z \ s \ winding_number g z = 0\ - and h = \\pa\s. 0 < h pa \ (\w\cball pa (h pa). w \ s \ (w \ pa \ w \ insert p pts))\ - have "h p>0" and "p\s" - and h_p: "\w\cball p (h p). w \ s \ (w \ p \ w \ insert p pts)" - using h \insert p pts \ s\ by auto - obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" - "path_image pg \ s-insert p pts" "f contour_integrable_on pg" - proof - - have "p + h p\cball p (h p)" using h[rule_format,of p] - by (simp add: \p \ s\ dist_norm) - then have "p + h p \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ - by fastforce - moreover have "pathstart g \ s - insert p pts " using path_img by auto - ultimately show ?thesis - using get_integrable_path[OF \open s\ connected fin holo,of "pathstart g" "p+h p"] that - by blast - qed - obtain n::int where "n=winding_number g p" - using integer_winding_number[OF _ g_loop,of p] valid path_img - by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) - define p_circ where "p_circ \ circlepath p (h p)" - define p_circ_pt where "p_circ_pt \ linepath (p+h p) (p+h p)" - define n_circ where "n_circ \ \n. ((+++) p_circ ^^ n) p_circ_pt" - define cp where "cp \ if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" - have n_circ:"valid_path (n_circ k)" - "winding_number (n_circ k) p = k" - "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" - "path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" - "p \ path_image (n_circ k)" - "\p'. p'\s - pts \ winding_number (n_circ k) p'=0 \ p'\path_image (n_circ k)" - "f contour_integrable_on (n_circ k)" - "contour_integral (n_circ k) f = k * contour_integral p_circ f" - for k - proof (induct k) - case 0 - show "valid_path (n_circ 0)" - and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" - and "winding_number (n_circ 0) p = of_nat 0" - and "pathstart (n_circ 0) = p + h p" - and "pathfinish (n_circ 0) = p + h p" - and "p \ path_image (n_circ 0)" - unfolding n_circ_def p_circ_pt_def using \h p > 0\ - by (auto simp add: dist_norm) - show "winding_number (n_circ 0) p'=0 \ p'\path_image (n_circ 0)" when "p'\s- pts" for p' - unfolding n_circ_def p_circ_pt_def - apply (auto intro!:winding_number_trivial) - by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ - show "f contour_integrable_on (n_circ 0)" - unfolding n_circ_def p_circ_pt_def - by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing) - show "contour_integral (n_circ 0) f = of_nat 0 * contour_integral p_circ f" - unfolding n_circ_def p_circ_pt_def by auto - next - case (Suc k) - have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto - have pcirc:"p \ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" - using Suc(3) unfolding p_circ_def using \h p > 0\ by (auto simp add: p_circ_def) - have pcirc_image:"path_image p_circ \ s - insert p pts" - proof - - have "path_image p_circ \ cball p (h p)" using \0 < h p\ p_circ_def by auto - then show ?thesis using h_p pcirc(1) by auto - qed - have pcirc_integrable:"f contour_integrable_on p_circ" - by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] - contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on - holomorphic_on_subset[OF holo]) - show "valid_path (n_circ (Suc k))" - using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto - show "path_image (n_circ (Suc k)) - = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" - proof - - have "path_image p_circ = sphere p (h p)" - unfolding p_circ_def using \0 < h p\ by auto - then show ?thesis unfolding n_Suc using Suc.hyps(5) \h p>0\ - by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) - qed - then show "p \ path_image (n_circ (Suc k))" using \h p>0\ by auto - show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" - proof - - have "winding_number p_circ p = 1" - by (simp add: \h p > 0\ p_circ_def winding_number_circlepath_centre) - moreover have "p \ path_image (n_circ k)" using Suc(5) \h p>0\ by auto - then have "winding_number (p_circ +++ n_circ k) p - = winding_number p_circ p + winding_number (n_circ k) p" - using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc - apply (intro winding_number_join) - by auto - ultimately show ?thesis using Suc(2) unfolding n_circ_def - by auto - qed - show "pathstart (n_circ (Suc k)) = p + h p" - by (simp add: n_circ_def p_circ_def) - show "pathfinish (n_circ (Suc k)) = p + h p" - using Suc(4) unfolding n_circ_def by auto - show "winding_number (n_circ (Suc k)) p'=0 \ p'\path_image (n_circ (Suc k))" when "p'\s-pts" for p' - proof - - have " p' \ path_image p_circ" using \p \ s\ h p_circ_def that using pcirc_image by blast - moreover have "p' \ path_image (n_circ k)" - using Suc.hyps(7) that by blast - moreover have "winding_number p_circ p' = 0" - proof - - have "path_image p_circ \ cball p (h p)" - using h unfolding p_circ_def using \p \ s\ by fastforce - moreover have "p'\cball p (h p)" using \p \ s\ h that "2.hyps"(2) by fastforce - ultimately show ?thesis unfolding p_circ_def - apply (intro winding_number_zero_outside) - by auto - qed - ultimately show ?thesis - unfolding n_Suc - apply (subst winding_number_join) - by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that]) - qed - show "f contour_integrable_on (n_circ (Suc k))" - unfolding n_Suc - by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) - show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" - unfolding n_Suc - by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] - Suc(9) algebra_simps) - qed - have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" - "valid_path cp" "path_image cp \ s - insert p pts" - "winding_number cp p = - n" - "\p'. p'\s - pts \ winding_number cp p'=0 \ p' \ path_image cp" - "f contour_integrable_on cp" - "contour_integral cp f = - n * contour_integral p_circ f" - proof - - show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" - using n_circ unfolding cp_def by auto - next - have "sphere p (h p) \ s - insert p pts" - using h[rule_format,of p] \insert p pts \ s\ by force - moreover have "p + complex_of_real (h p) \ s - insert p pts" - using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) - ultimately show "path_image cp \ s - insert p pts" unfolding cp_def - using n_circ(5) by auto - next - show "winding_number cp p = - n" - unfolding cp_def using winding_number_reversepath n_circ \h p>0\ - by (auto simp: valid_path_imp_path) - next - show "winding_number cp p'=0 \ p' \ path_image cp" when "p'\s - pts" for p' - unfolding cp_def - apply (auto) - apply (subst winding_number_reversepath) - by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1)) - next - show "f contour_integrable_on cp" unfolding cp_def - using contour_integrable_reversepath_eq n_circ(1,8) by auto - next - show "contour_integral cp f = - n * contour_integral p_circ f" - unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) - by auto - qed - define g' where "g' \ g +++ pg +++ cp +++ (reversepath pg)" - have "contour_integral g' f = (\p\pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" - proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) - show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) - show "open (s - {p})" using \open s\ by auto - show " pts \ s - {p}" using \insert p pts \ s\ \ p \ pts\ by blast - show "f holomorphic_on s - {p} - pts" using holo \p \ pts\ by (metis Diff_insert2) - show "valid_path g'" - unfolding g'_def cp_def using n_circ valid pg g_loop - by (auto intro!:valid_path_join ) - show "pathfinish g' = pathstart g'" - unfolding g'_def cp_def using pg(2) by simp - show "path_image g' \ s - {p} - pts" - proof - - define s' where "s' \ s - {p} - pts" - have s':"s' = s-insert p pts " unfolding s'_def by auto - then show ?thesis using path_img pg(4) cp(4) - unfolding g'_def - apply (fold s'_def s') - apply (intro subset_path_image_join) - by auto - qed - note path_join_imp[simp] - show "\z. z \ s - {p} \ winding_number g' z = 0" - proof clarify - fix z assume z:"z\s - {p}" - have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z - + winding_number (pg +++ cp +++ (reversepath pg)) z" - proof (rule winding_number_join) - show "path g" using \valid_path g\ by (simp add: valid_path_imp_path) - show "z \ path_image g" using z path_img by auto - show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp - by (simp add: valid_path_imp_path) - next - have "path_image (pg +++ cp +++ reversepath pg) \ s - insert p pts" - using pg(4) cp(4) by (auto simp:subset_path_image_join) - then show "z \ path_image (pg +++ cp +++ reversepath pg)" using z by auto - next - show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto - qed - also have "... = winding_number g z + (winding_number pg z - + winding_number (cp +++ (reversepath pg)) z)" - proof (subst add_left_cancel,rule winding_number_join) - show "path pg" and "path (cp +++ reversepath pg)" - and "pathfinish pg = pathstart (cp +++ reversepath pg)" - by (auto simp add: valid_path_imp_path) - show "z \ path_image pg" using pg(4) z by blast - show "z \ path_image (cp +++ reversepath pg)" using z - by (metis Diff_iff \z \ path_image pg\ contra_subsetD cp(4) insertI1 - not_in_path_image_join path_image_reversepath singletonD) - qed - also have "... = winding_number g z + (winding_number pg z - + (winding_number cp z + winding_number (reversepath pg) z))" - apply (auto intro!:winding_number_join simp: valid_path_imp_path) - apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) - by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) - also have "... = winding_number g z + winding_number cp z" - apply (subst winding_number_reversepath) - apply (auto simp: valid_path_imp_path) - by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) - finally have "winding_number g' z = winding_number g z + winding_number cp z" - unfolding g'_def . - moreover have "winding_number g z + winding_number cp z = 0" - using winding z \n=winding_number g p\ by auto - ultimately show "winding_number g' z = 0" unfolding g'_def by auto - qed - show "\pa\s - {p}. 0 < h pa \ (\w\cball pa (h pa). w \ s - {p} \ (w \ pa \ w \ pts))" - using h by fastforce - qed - moreover have "contour_integral g' f = contour_integral g f - - winding_number g p * contour_integral p_circ f" - proof - - have "contour_integral g' f = contour_integral g f - + contour_integral (pg +++ cp +++ reversepath pg) f" - unfolding g'_def - apply (subst contour_integral_join) - by (auto simp add:open_Diff[OF \open s\,OF finite_imp_closed[OF fin]] - intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] - contour_integrable_reversepath) - also have "... = contour_integral g f + contour_integral pg f - + contour_integral (cp +++ reversepath pg) f" - apply (subst contour_integral_join) - by (auto simp add:contour_integrable_reversepath) - also have "... = contour_integral g f + contour_integral pg f - + contour_integral cp f + contour_integral (reversepath pg) f" - apply (subst contour_integral_join) - by (auto simp add:contour_integrable_reversepath) - also have "... = contour_integral g f + contour_integral cp f" - using contour_integral_reversepath - by (auto simp add:contour_integrable_reversepath) - also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f" - using \n=winding_number g p\ by auto - finally show ?thesis . - qed - moreover have "winding_number g' p' = winding_number g p'" when "p'\pts" for p' - proof - - have [simp]: "p' \ path_image g" "p' \ path_image pg" "p'\path_image cp" - using "2.prems"(8) that - apply blast - apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) - by (meson DiffD2 cp(4) rev_subsetD subset_insertI that) - have "winding_number g' p' = winding_number g p' - + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def - apply (subst winding_number_join) - apply (simp_all add: valid_path_imp_path) - apply (intro not_in_path_image_join) - by auto - also have "... = winding_number g p' + winding_number pg p' - + winding_number (cp +++ reversepath pg) p'" - apply (subst winding_number_join) - apply (simp_all add: valid_path_imp_path) - apply (intro not_in_path_image_join) - by auto - also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p' - + winding_number (reversepath pg) p'" - apply (subst winding_number_join) - by (simp_all add: valid_path_imp_path) - also have "... = winding_number g p' + winding_number cp p'" - apply (subst winding_number_reversepath) - by (simp_all add: valid_path_imp_path) - also have "... = winding_number g p'" using that by auto - finally show ?thesis . - qed - ultimately show ?case unfolding p_circ_def - apply (subst (asm) sum.cong[OF refl, - of pts _ "\p. winding_number g p * contour_integral (circlepath p (h p)) f"]) - by (auto simp add:sum.insert[OF \finite pts\ \p\pts\] algebra_simps) -qed - -lemma Cauchy_theorem_singularities: - assumes "open s" "connected s" "finite pts" and - holo:"f holomorphic_on s-pts" and - "valid_path g" and - loop:"pathfinish g = pathstart g" and - "path_image g \ s-pts" and - homo:"\z. (z \ s) \ winding_number g z = 0" and - avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" - shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" - (is "?L=?R") -proof - - define circ where "circ \ \p. winding_number g p * contour_integral (circlepath p (h p)) f" - define pts1 where "pts1 \ pts \ s" - define pts2 where "pts2 \ pts - pts1" - have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" - unfolding pts1_def pts2_def by auto - have "contour_integral g f = (\p\pts1. circ p)" unfolding circ_def - proof (rule Cauchy_theorem_aux[OF \open s\ _ _ \pts1\s\ _ \valid_path g\ loop _ homo]) - have "finite pts1" unfolding pts1_def using \finite pts\ by auto - then show "connected (s - pts1)" - using \open s\ \connected s\ connected_open_delete_finite[of s] by auto - next - show "finite pts1" using \pts = pts1 \ pts2\ assms(3) by auto - show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) - show "path_image g \ s - pts1" using assms(7) pts1_def by auto - show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ pts1))" - by (simp add: avoid pts1_def) - qed - moreover have "sum circ pts2=0" - proof - - have "winding_number g p=0" when "p\pts2" for p - using \pts2 \ s={}\ that homo[rule_format,of p] by auto - thus ?thesis unfolding circ_def - apply (intro sum.neutral) - by auto - qed - moreover have "?R=sum circ pts1 + sum circ pts2" - unfolding circ_def - using sum.union_disjoint[OF _ _ \pts1 \ pts2 = {}\] \finite pts\ \pts=pts1 \ pts2\ - by blast - ultimately show ?thesis - apply (fold circ_def) - by auto -qed - -theorem Residue_theorem: - fixes s pts::"complex set" and f::"complex \ complex" - and g::"real \ complex" - assumes "open s" "connected s" "finite pts" and - holo:"f holomorphic_on s-pts" and - "valid_path g" and - loop:"pathfinish g = pathstart g" and - "path_image g \ s-pts" and - homo:"\z. (z \ s) \ winding_number g z = 0" - shows "contour_integral g f = 2 * pi * \ *(\p\pts. winding_number g p * residue f p)" -proof - - define c where "c \ 2 * pi * \" - obtain h where avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" - using finite_cball_avoid[OF \open s\ \finite pts\] by metis - have "contour_integral g f - = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" - using Cauchy_theorem_singularities[OF assms avoid] . - also have "... = (\p\pts. c * winding_number g p * residue f p)" - proof (intro sum.cong) - show "pts = pts" by simp - next - fix x assume "x \ pts" - show "winding_number g x * contour_integral (circlepath x (h x)) f - = c * winding_number g x * residue f x" - proof (cases "x\s") - case False - then have "winding_number g x=0" using homo by auto - thus ?thesis by auto - next - case True - have "contour_integral (circlepath x (h x)) f = c* residue f x" - using \x\pts\ \finite pts\ avoid[rule_format,OF True] - apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) - by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \open s\ finite_imp_closed]) - then show ?thesis by auto - qed - qed - also have "... = c * (\p\pts. winding_number g p * residue f p)" - by (simp add: sum_distrib_left algebra_simps) - finally show ?thesis unfolding c_def . -qed - -subsection \Non-essential singular points\ - -definition\<^marker>\tag important\ is_pole :: - "('a::topological_space \ 'b::real_normed_vector) \ 'a \ bool" where - "is_pole f a = (LIM x (at a). f x :> at_infinity)" - -lemma is_pole_cong: - assumes "eventually (\x. f x = g x) (at a)" "a=b" - shows "is_pole f a \ is_pole g b" - unfolding is_pole_def using assms by (intro filterlim_cong,auto) - -lemma is_pole_transform: - assumes "is_pole f a" "eventually (\x. f x = g x) (at a)" "a=b" - shows "is_pole g b" - using is_pole_cong assms by auto - -lemma is_pole_tendsto: - fixes f::"('a::topological_space \ 'b::real_normed_div_algebra)" - shows "is_pole f x \ ((inverse o f) \ 0) (at x)" -unfolding is_pole_def -by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) - -lemma is_pole_inverse_holomorphic: - assumes "open s" - and f_holo:"f holomorphic_on (s-{z})" - and pole:"is_pole f z" - and non_z:"\x\s-{z}. f x\0" - shows "(\x. if x=z then 0 else inverse (f x)) holomorphic_on s" -proof - - define g where "g \ \x. if x=z then 0 else inverse (f x)" - have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] - apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \ f"]) - by (simp_all add:g_def) - moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto - hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def - by (auto elim!:continuous_on_inverse simp add:non_z) - hence "continuous_on (s-{z}) g" unfolding g_def - apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) - by auto - ultimately have "continuous_on s g" using open_delete[OF \open s\] \open s\ - by (auto simp add:continuous_on_eq_continuous_at) - moreover have "(inverse o f) holomorphic_on (s-{z})" - unfolding comp_def using f_holo - by (auto elim!:holomorphic_on_inverse simp add:non_z) - hence "g holomorphic_on (s-{z})" - apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) - by (auto simp add:g_def) - ultimately show ?thesis unfolding g_def using \open s\ - by (auto elim!: no_isolated_singularity) -qed - -lemma not_is_pole_holomorphic: - assumes "open A" "x \ A" "f holomorphic_on A" - shows "\is_pole f x" -proof - - have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact - with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at) - hence "f \x\ f x" by (simp add: isCont_def) - thus "\is_pole f x" unfolding is_pole_def - using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto -qed - -lemma is_pole_inverse_power: "n > 0 \ is_pole (\z::complex. 1 / (z - a) ^ n) a" - unfolding is_pole_def inverse_eq_divide [symmetric] - by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros) - (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros) - -lemma is_pole_inverse: "is_pole (\z::complex. 1 / (z - a)) a" - using is_pole_inverse_power[of 1 a] by simp - -lemma is_pole_divide: - fixes f :: "'a :: t2_space \ 'b :: real_normed_field" - assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \ 0" - shows "is_pole (\z. f z / g z) z" -proof - - have "filterlim (\z. f z * inverse (g z)) at_infinity (at z)" - by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"] - filterlim_compose[OF filterlim_inverse_at_infinity])+ - (insert assms, auto simp: isCont_def) - thus ?thesis by (simp add: field_split_simps is_pole_def) -qed - -lemma is_pole_basic: - assumes "f holomorphic_on A" "open A" "z \ A" "f z \ 0" "n > 0" - shows "is_pole (\w. f w / (w - z) ^ n) z" -proof (rule is_pole_divide) - have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact - with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at) - have "filterlim (\w. (w - z) ^ n) (nhds 0) (at z)" - using assms by (auto intro!: tendsto_eq_intros) - thus "filterlim (\w. (w - z) ^ n) (at 0) (at z)" - by (intro filterlim_atI tendsto_eq_intros) - (insert assms, auto simp: eventually_at_filter) -qed fact+ - -lemma is_pole_basic': - assumes "f holomorphic_on A" "open A" "0 \ A" "f 0 \ 0" "n > 0" - shows "is_pole (\w. f w / w ^ n) 0" - using is_pole_basic[of f A 0] assms by simp - -text \The proposition - \<^term>\\x. ((f::complex\complex) \ x) (at z) \ is_pole f z\ -can be interpreted as the complex function \<^term>\f\ has a non-essential singularity at \<^term>\z\ -(i.e. the singularity is either removable or a pole).\ -definition not_essential::"[complex \ complex, complex] \ bool" where - "not_essential f z = (\x. f\z\x \ is_pole f z)" - -definition isolated_singularity_at::"[complex \ complex, complex] \ bool" where - "isolated_singularity_at f z = (\r>0. f analytic_on ball z r-{z})" - -named_theorems singularity_intros "introduction rules for singularities" - -lemma holomorphic_factor_unique: - fixes f::"complex \ complex" and z::complex and r::real and m n::int - assumes "r>0" "g z\0" "h z\0" - and asm:"\w\ball z r-{z}. f w = g w * (w-z) powr n \ g w\0 \ f w = h w * (w - z) powr m \ h w\0" - and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" - shows "n=m" -proof - - have [simp]:"at z within ball z r \ bot" using \r>0\ - by (auto simp add:at_within_ball_bot_iff) - have False when "n>m" - proof - - have "(h \ 0) (at z within ball z r)" - proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (n - m) * g w"]) - have "\w\ball z r-{z}. h w = (w-z)powr(n-m) * g w" - using \n>m\ asm \r>0\ - apply (auto simp add:field_simps powr_diff) - by force - then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ - \ (x' - z) powr (n - m) * g x' = h x'" for x' by auto - next - define F where "F \ at z within ball z r" - define f' where "f' \ \x. (x - z) powr (n-m)" - have "f' z=0" using \n>m\ unfolding f'_def by auto - moreover have "continuous F f'" unfolding f'_def F_def continuous_def - apply (subst Lim_ident_at) - using \n>m\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) - ultimately have "(f' \ 0) F" unfolding F_def - by (simp add: continuous_within) - moreover have "(g \ g z) F" - using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \r>0\ - unfolding F_def by auto - ultimately show " ((\w. f' w * g w) \ 0) F" using tendsto_mult by fastforce - qed - moreover have "(h \ h z) (at z within ball z r)" - using holomorphic_on_imp_continuous_on[OF h_holo] - by (auto simp add:continuous_on_def \r>0\) - ultimately have "h z=0" by (auto intro!: tendsto_unique) - thus False using \h z\0\ by auto - qed - moreover have False when "m>n" - proof - - have "(g \ 0) (at z within ball z r)" - proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (m - n) * h w"]) - have "\w\ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \m>n\ asm - apply (auto simp add:field_simps powr_diff) - by force - then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ - \ (x' - z) powr (m - n) * h x' = g x'" for x' by auto - next - define F where "F \ at z within ball z r" - define f' where "f' \\x. (x - z) powr (m-n)" - have "f' z=0" using \m>n\ unfolding f'_def by auto - moreover have "continuous F f'" unfolding f'_def F_def continuous_def - apply (subst Lim_ident_at) - using \m>n\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) - ultimately have "(f' \ 0) F" unfolding F_def - by (simp add: continuous_within) - moreover have "(h \ h z) F" - using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \r>0\ - unfolding F_def by auto - ultimately show " ((\w. f' w * h w) \ 0) F" using tendsto_mult by fastforce - qed - moreover have "(g \ g z) (at z within ball z r)" - using holomorphic_on_imp_continuous_on[OF g_holo] - by (auto simp add:continuous_on_def \r>0\) - ultimately have "g z=0" by (auto intro!: tendsto_unique) - thus False using \g z\0\ by auto - qed - ultimately show "n=m" by fastforce -qed - -lemma holomorphic_factor_puncture: - assumes f_iso:"isolated_singularity_at f z" - and "not_essential f z" \ \\<^term>\f\ has either a removable singularity or a pole at \<^term>\z\\ - and non_zero:"\\<^sub>Fw in (at z). f w\0" \ \\<^term>\f\ will not be constantly zero in a neighbour of \<^term>\z\\ - shows "\!n::int. \g r. 0 < r \ g holomorphic_on cball z r \ g z\0 - \ (\w\cball z r-{z}. f w = g w * (w-z) powr n \ g w\0)" -proof - - define P where "P = (\f n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 - \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" - have imp_unique:"\!n::int. \g r. P f n g r" when "\n g r. P f n g r" - proof (rule ex_ex1I[OF that]) - fix n1 n2 :: int - assume g1_asm:"\g1 r1. P f n1 g1 r1" and g2_asm:"\g2 r2. P f n2 g2 r2" - define fac where "fac \ \n g r. \w\cball z r-{z}. f w = g w * (w - z) powr (of_int n) \ g w \ 0" - obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\0" - and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto - obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\0" - and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto - define r where "r \ min r1 r2" - have "r>0" using \r1>0\ \r2>0\ unfolding r_def by auto - moreover have "\w\ball z r-{z}. f w = g1 w * (w-z) powr n1 \ g1 w\0 - \ f w = g2 w * (w - z) powr n2 \ g2 w\0" - using \fac n1 g1 r1\ \fac n2 g2 r2\ unfolding fac_def r_def - by fastforce - ultimately show "n1=n2" using g1_holo g2_holo \g1 z\0\ \g2 z\0\ - apply (elim holomorphic_factor_unique) - by (auto simp add:r_def) - qed - - have P_exist:"\ n g r. P h n g r" when - "\z'. (h \ z') (at z)" "isolated_singularity_at h z" "\\<^sub>Fw in (at z). h w\0" - for h - proof - - from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}" - unfolding isolated_singularity_at_def by auto - obtain z' where "(h \ z') (at z)" using \\z'. (h \ z') (at z)\ by auto - define h' where "h'=(\x. if x=z then z' else h x)" - have "h' holomorphic_on ball z r" - apply (rule no_isolated_singularity'[of "{z}"]) - subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \h \z\ z'\ empty_iff h'_def insert_iff) - subgoal using \h analytic_on ball z r - {z}\ analytic_imp_holomorphic h'_def holomorphic_transform - by fastforce - by auto - have ?thesis when "z'=0" - proof - - have "h' z=0" using that unfolding h'_def by auto - moreover have "\ h' constant_on ball z r" - using \\\<^sub>Fw in (at z). h w\0\ unfolding constant_on_def frequently_def eventually_at h'_def - apply simp - by (metis \0 < r\ centre_in_ball dist_commute mem_ball that) - moreover note \h' holomorphic_on ball z r\ - ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \ ball z r" and - g:"g holomorphic_on ball z r1" - "\w. w \ ball z r1 \ h' w = (w - z) ^ n * g w" - "\w. w \ ball z r1 \ g w \ 0" - using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified, - OF \h' holomorphic_on ball z r\ \r>0\ \h' z=0\ \\ h' constant_on ball z r\] - by (auto simp add:dist_commute) - define rr where "rr=r1/2" - have "P h' n g rr" - unfolding P_def rr_def - using \n>0\ \r1>0\ g by (auto simp add:powr_nat) - then have "P h n g rr" - unfolding h'_def P_def by auto - then show ?thesis unfolding P_def by blast - qed - moreover have ?thesis when "z'\0" - proof - - have "h' z\0" using that unfolding h'_def by auto - obtain r1 where "r1>0" "cball z r1 \ ball z r" "\x\cball z r1. h' x\0" - proof - - have "isCont h' z" "h' z\0" - by (auto simp add: Lim_cong_within \h \z\ z'\ \z'\0\ continuous_at h'_def) - then obtain r2 where r2:"r2>0" "\x\ball z r2. h' x\0" - using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto - define r1 where "r1=min r2 r / 2" - have "0 < r1" "cball z r1 \ ball z r" - using \r2>0\ \r>0\ unfolding r1_def by auto - moreover have "\x\cball z r1. h' x \ 0" - using r2 unfolding r1_def by simp - ultimately show ?thesis using that by auto - qed - then have "P h' 0 h' r1" using \h' holomorphic_on ball z r\ unfolding P_def by auto - then have "P h 0 h' r1" unfolding P_def h'_def by auto - then show ?thesis unfolding P_def by blast - qed - ultimately show ?thesis by auto - qed - - have ?thesis when "\x. (f \ x) (at z)" - apply (rule_tac imp_unique[unfolded P_def]) - using P_exist[OF that(1) f_iso non_zero] unfolding P_def . - moreover have ?thesis when "is_pole f z" - proof (rule imp_unique[unfolded P_def]) - obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\x\ball z e-{z}. f x\0" - proof - - have "\\<^sub>F z in at z. f z \ 0" - using \is_pole f z\ filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def - by auto - then obtain e1 where e1:"e1>0" "\x\ball z e1-{z}. f x\0" - using that eventually_at[of "\x. f x\0" z UNIV,simplified] by (auto simp add:dist_commute) - obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}" - using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto - define e where "e=min e1 e2" - show ?thesis - apply (rule that[of e]) - using e1 e2 unfolding e_def by auto - qed - - define h where "h \ \x. inverse (f x)" - - have "\n g r. P h n g r" - proof - - have "h \z\ 0" - using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce - moreover have "\\<^sub>Fw in (at z). h w\0" - using non_zero - apply (elim frequently_rev_mp) - unfolding h_def eventually_at by (auto intro:exI[where x=1]) - moreover have "isolated_singularity_at h z" - unfolding isolated_singularity_at_def h_def - apply (rule exI[where x=e]) - using e_holo e_nz \e>0\ by (metis open_ball analytic_on_open - holomorphic_on_inverse open_delete) - ultimately show ?thesis - using P_exist[of h] by auto - qed - then obtain n g r - where "0 < r" and - g_holo:"g holomorphic_on cball z r" and "g z\0" and - g_fac:"(\w\cball z r-{z}. h w = g w * (w - z) powr of_int n \ g w \ 0)" - unfolding P_def by auto - have "P f (-n) (inverse o g) r" - proof - - have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\cball z r - {z}" for w - using g_fac[rule_format,of w] that unfolding h_def - apply (auto simp add:powr_minus ) - by (metis inverse_inverse_eq inverse_mult_distrib) - then show ?thesis - unfolding P_def comp_def - using \r>0\ g_holo g_fac \g z\0\ by (auto intro:holomorphic_intros) - qed - then show "\x g r. 0 < r \ g holomorphic_on cball z r \ g z \ 0 - \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int x \ g w \ 0)" - unfolding P_def by blast - qed - ultimately show ?thesis using \not_essential f z\ unfolding not_essential_def by presburger -qed - -lemma not_essential_transform: - assumes "not_essential g z" - assumes "\\<^sub>F w in (at z). g w = f w" - shows "not_essential f z" - using assms unfolding not_essential_def - by (simp add: filterlim_cong is_pole_cong) - -lemma isolated_singularity_at_transform: - assumes "isolated_singularity_at g z" - assumes "\\<^sub>F w in (at z). g w = f w" - shows "isolated_singularity_at f z" -proof - - obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}" - using assms(1) unfolding isolated_singularity_at_def by auto - obtain r2 where "r2>0" and r2:" \x. x \ z \ dist x z < r2 \ g x = f x" - using assms(2) unfolding eventually_at by auto - define r3 where "r3=min r1 r2" - have "r3>0" unfolding r3_def using \r1>0\ \r2>0\ by auto - moreover have "f analytic_on ball z r3 - {z}" - proof - - have "g holomorphic_on ball z r3 - {z}" - using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto) - then have "f holomorphic_on ball z r3 - {z}" - using r2 unfolding r3_def - by (auto simp add:dist_commute elim!:holomorphic_transform) - then show ?thesis by (subst analytic_on_open,auto) - qed - ultimately show ?thesis unfolding isolated_singularity_at_def by auto -qed - -lemma not_essential_powr[singularity_intros]: - assumes "LIM w (at z). f w :> (at x)" - shows "not_essential (\w. (f w) powr (of_int n)) z" -proof - - define fp where "fp=(\w. (f w) powr (of_int n))" - have ?thesis when "n>0" - proof - - have "(\w. (f w) ^ (nat n)) \z\ x ^ nat n" - using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros) - then have "fp \z\ x ^ nat n" unfolding fp_def - apply (elim Lim_transform_within[where d=1],simp) - by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power) - then show ?thesis unfolding not_essential_def fp_def by auto - qed - moreover have ?thesis when "n=0" - proof - - have "fp \z\ 1 " - apply (subst tendsto_cong[where g="\_.1"]) - using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto - then show ?thesis unfolding fp_def not_essential_def by auto - qed - moreover have ?thesis when "n<0" - proof (cases "x=0") - case True - have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity" - apply (subst filterlim_inverse_at_iff[symmetric],simp) - apply (rule filterlim_atI) - subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros) - subgoal using filterlim_at_within_not_equal[OF assms,of 0] - by (eventually_elim,insert that,auto) - done - then have "LIM w (at z). fp w :> at_infinity" - proof (elim filterlim_mono_eventually) - show "\\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x" - using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def - apply eventually_elim - using powr_of_int that by auto - qed auto - then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto - next - case False - let ?xx= "inverse (x ^ (nat (-n)))" - have "(\w. inverse ((f w) ^ (nat (-n)))) \z\?xx" - using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros) - then have "fp \z\?xx" - apply (elim Lim_transform_within[where d=1],simp) - unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less - not_le power_eq_0_iff powr_0 powr_of_int that) - then show ?thesis unfolding fp_def not_essential_def by auto - qed - ultimately show ?thesis by linarith -qed - -lemma isolated_singularity_at_powr[singularity_intros]: - assumes "isolated_singularity_at f z" "\\<^sub>F w in (at z). f w\0" - shows "isolated_singularity_at (\w. (f w) powr (of_int n)) z" -proof - - obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}" - using assms(1) unfolding isolated_singularity_at_def by auto - then have r1:"f holomorphic_on ball z r1 - {z}" - using analytic_on_open[of "ball z r1-{z}" f] by blast - obtain r2 where "r2>0" and r2:"\w. w \ z \ dist w z < r2 \ f w \ 0" - using assms(2) unfolding eventually_at by auto - define r3 where "r3=min r1 r2" - have "(\w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}" - apply (rule holomorphic_on_powr_of_int) - subgoal unfolding r3_def using r1 by auto - subgoal unfolding r3_def using r2 by (auto simp add:dist_commute) - done - moreover have "r3>0" unfolding r3_def using \0 < r1\ \0 < r2\ by linarith - ultimately show ?thesis unfolding isolated_singularity_at_def - apply (subst (asm) analytic_on_open[symmetric]) - by auto -qed - -lemma non_zero_neighbour: - assumes f_iso:"isolated_singularity_at f z" - and f_ness:"not_essential f z" - and f_nconst:"\\<^sub>Fw in (at z). f w\0" - shows "\\<^sub>F w in (at z). f w\0" -proof - - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" - "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" - using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto - have "f w \ 0" when " w \ z" "dist w z < fr" for w - proof - - have "f w = fp w * (w - z) powr of_int fn" "fp w \ 0" - using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute) - moreover have "(w - z) powr of_int fn \0" - unfolding powr_eq_0_iff using \w\z\ by auto - ultimately show ?thesis by auto - qed - then show ?thesis using \fr>0\ unfolding eventually_at by auto -qed - -lemma non_zero_neighbour_pole: - assumes "is_pole f z" - shows "\\<^sub>F w in (at z). f w\0" - using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0] - unfolding is_pole_def by auto - -lemma non_zero_neighbour_alt: - assumes holo: "f holomorphic_on S" - and "open S" "connected S" "z \ S" "\ \ S" "f \ \ 0" - shows "\\<^sub>F w in (at z). f w\0 \ w\S" -proof (cases "f z = 0") - case True - from isolated_zeros[OF holo \open S\ \connected S\ \z \ S\ True \\ \ S\ \f \ \ 0\] - obtain r where "0 < r" "ball z r \ S" "\w \ ball z r - {z}.f w \ 0" by metis - then show ?thesis unfolding eventually_at - apply (rule_tac x=r in exI) - by (auto simp add:dist_commute) -next - case False - obtain r1 where r1:"r1>0" "\y. dist z y < r1 \ f y \ 0" - using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at - holo holomorphic_on_imp_continuous_on by blast - obtain r2 where r2:"r2>0" "ball z r2 \ S" - using assms(2) assms(4) openE by blast - show ?thesis unfolding eventually_at - apply (rule_tac x="min r1 r2" in exI) - using r1 r2 by (auto simp add:dist_commute) -qed - -lemma not_essential_times[singularity_intros]: - assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" - assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" - shows "not_essential (\w. f w * g w) z" -proof - - define fg where "fg = (\w. f w * g w)" - have ?thesis when "\ ((\\<^sub>Fw in (at z). f w\0) \ (\\<^sub>Fw in (at z). g w\0))" - proof - - have "\\<^sub>Fw in (at z). fg w=0" - using that[unfolded frequently_def, simplified] unfolding fg_def - by (auto elim: eventually_rev_mp) - from tendsto_cong[OF this] have "fg \z\0" by auto - then show ?thesis unfolding not_essential_def fg_def by auto - qed - moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" and g_nconst:"\\<^sub>Fw in (at z). g w\0" - proof - - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" - "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" - using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto - obtain gn gp gr where [simp]:"gp z \ 0" and "gr > 0" - and gr: "gp holomorphic_on cball z gr" - "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" - using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto - - define r1 where "r1=(min fr gr)" - have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto - have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" - when "w\ball z r1 - {z}" for w - proof - - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" - using fr(2)[rule_format,of w] that unfolding r1_def by auto - moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" - using gr(2)[rule_format, of w] that unfolding r1_def by auto - ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" - unfolding fg_def by (auto simp add:powr_add) - qed - - have [intro]: "fp \z\fp z" "gp \z\gp z" - using fr(1) \fr>0\ gr(1) \gr>0\ - by (meson open_ball ball_subset_cball centre_in_ball - continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on - holomorphic_on_subset)+ - have ?thesis when "fn+gn>0" - proof - - have "(\w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \z\0" - using that by (auto intro!:tendsto_eq_intros) - then have "fg \z\ 0" - apply (elim Lim_transform_within[OF _ \r1>0\]) - by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self - eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int - that) - then show ?thesis unfolding not_essential_def fg_def by auto - qed - moreover have ?thesis when "fn+gn=0" - proof - - have "(\w. fp w * gp w) \z\fp z*gp z" - using that by (auto intro!:tendsto_eq_intros) - then have "fg \z\ fp z*gp z" - apply (elim Lim_transform_within[OF _ \r1>0\]) - apply (subst fg_times) - by (auto simp add:dist_commute that) - then show ?thesis unfolding not_essential_def fg_def by auto - qed - moreover have ?thesis when "fn+gn<0" - proof - - have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity" - apply (rule filterlim_divide_at_infinity) - apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI) - using eventually_at_topological by blast - then have "is_pole fg z" unfolding is_pole_def - apply (elim filterlim_transform_within[OF _ _ \r1>0\],simp) - apply (subst fg_times,simp add:dist_commute) - apply (subst powr_of_int) - using that by (auto simp add:field_split_simps) - then show ?thesis unfolding not_essential_def fg_def by auto - qed - ultimately show ?thesis unfolding not_essential_def fg_def by fastforce - qed - ultimately show ?thesis by auto -qed - -lemma not_essential_inverse[singularity_intros]: - assumes f_ness:"not_essential f z" - assumes f_iso:"isolated_singularity_at f z" - shows "not_essential (\w. inverse (f w)) z" -proof - - define vf where "vf = (\w. inverse (f w))" - have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" - proof - - have "\\<^sub>Fw in (at z). f w=0" - using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) - then have "\\<^sub>Fw in (at z). vf w=0" - unfolding vf_def by auto - from tendsto_cong[OF this] have "vf \z\0" unfolding vf_def by auto - then show ?thesis unfolding not_essential_def vf_def by auto - qed - moreover have ?thesis when "is_pole f z" - proof - - have "vf \z\0" - using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast - then show ?thesis unfolding not_essential_def vf_def by auto - qed - moreover have ?thesis when "\x. f\z\x " and f_nconst:"\\<^sub>Fw in (at z). f w\0" - proof - - from that obtain fz where fz:"f\z\fz" by auto - have ?thesis when "fz=0" - proof - - have "(\w. inverse (vf w)) \z\0" - using fz that unfolding vf_def by auto - moreover have "\\<^sub>F w in at z. inverse (vf w) \ 0" - using non_zero_neighbour[OF f_iso f_ness f_nconst] - unfolding vf_def by auto - ultimately have "is_pole vf z" - using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto - then show ?thesis unfolding not_essential_def vf_def by auto - qed - moreover have ?thesis when "fz\0" - proof - - have "vf \z\inverse fz" - using fz that unfolding vf_def by (auto intro:tendsto_eq_intros) - then show ?thesis unfolding not_essential_def vf_def by auto - qed - ultimately show ?thesis by auto - qed - ultimately show ?thesis using f_ness unfolding not_essential_def by auto -qed - -lemma isolated_singularity_at_inverse[singularity_intros]: - assumes f_iso:"isolated_singularity_at f z" - and f_ness:"not_essential f z" - shows "isolated_singularity_at (\w. inverse (f w)) z" -proof - - define vf where "vf = (\w. inverse (f w))" - have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" - proof - - have "\\<^sub>Fw in (at z). f w=0" - using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) - then have "\\<^sub>Fw in (at z). vf w=0" - unfolding vf_def by auto - then obtain d1 where "d1>0" and d1:"\x. x \ z \ dist x z < d1 \ vf x = 0" - unfolding eventually_at by auto - then have "vf holomorphic_on ball z d1-{z}" - apply (rule_tac holomorphic_transform[of "\_. 0"]) - by (auto simp add:dist_commute) - then have "vf analytic_on ball z d1 - {z}" - by (simp add: analytic_on_open open_delete) - then show ?thesis using \d1>0\ unfolding isolated_singularity_at_def vf_def by auto - qed - moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" - proof - - have "\\<^sub>F w in at z. f w \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] . - then obtain d1 where d1:"d1>0" "\x. x \ z \ dist x z < d1 \ f x \ 0" - unfolding eventually_at by auto - obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}" - using f_iso unfolding isolated_singularity_at_def by auto - define d3 where "d3=min d1 d2" - have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto - moreover have "vf analytic_on ball z d3 - {z}" - unfolding vf_def - apply (rule analytic_on_inverse) - subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto - subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute) - done - ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto - qed - ultimately show ?thesis by auto -qed - -lemma not_essential_divide[singularity_intros]: - assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" - assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" - shows "not_essential (\w. f w / g w) z" -proof - - have "not_essential (\w. f w * inverse (g w)) z" - apply (rule not_essential_times[where g="\w. inverse (g w)"]) - using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse) - then show ?thesis by (simp add:field_simps) -qed - -lemma - assumes f_iso:"isolated_singularity_at f z" - and g_iso:"isolated_singularity_at g z" - shows isolated_singularity_at_times[singularity_intros]: - "isolated_singularity_at (\w. f w * g w) z" and - isolated_singularity_at_add[singularity_intros]: - "isolated_singularity_at (\w. f w + g w) z" -proof - - obtain d1 d2 where "d1>0" "d2>0" - and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}" - using f_iso g_iso unfolding isolated_singularity_at_def by auto - define d3 where "d3=min d1 d2" - have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto - - have "(\w. f w * g w) analytic_on ball z d3 - {z}" - apply (rule analytic_on_mult) - using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) - then show "isolated_singularity_at (\w. f w * g w) z" - using \d3>0\ unfolding isolated_singularity_at_def by auto - have "(\w. f w + g w) analytic_on ball z d3 - {z}" - apply (rule analytic_on_add) - using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) - then show "isolated_singularity_at (\w. f w + g w) z" - using \d3>0\ unfolding isolated_singularity_at_def by auto -qed - -lemma isolated_singularity_at_uminus[singularity_intros]: - assumes f_iso:"isolated_singularity_at f z" - shows "isolated_singularity_at (\w. - f w) z" - using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast - -lemma isolated_singularity_at_id[singularity_intros]: - "isolated_singularity_at (\w. w) z" - unfolding isolated_singularity_at_def by (simp add: gt_ex) - -lemma isolated_singularity_at_minus[singularity_intros]: - assumes f_iso:"isolated_singularity_at f z" - and g_iso:"isolated_singularity_at g z" - shows "isolated_singularity_at (\w. f w - g w) z" - using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\w. - g w"] - ,OF g_iso] by simp - -lemma isolated_singularity_at_divide[singularity_intros]: - assumes f_iso:"isolated_singularity_at f z" - and g_iso:"isolated_singularity_at g z" - and g_ness:"not_essential g z" - shows "isolated_singularity_at (\w. f w / g w) z" - using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso, - of "\w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps) - -lemma isolated_singularity_at_const[singularity_intros]: - "isolated_singularity_at (\w. c) z" - unfolding isolated_singularity_at_def by (simp add: gt_ex) - -lemma isolated_singularity_at_holomorphic: - assumes "f holomorphic_on s-{z}" "open s" "z\s" - shows "isolated_singularity_at f z" - using assms unfolding isolated_singularity_at_def - by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) - -subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ - - -definition\<^marker>\tag important\ zorder :: "(complex \ complex) \ complex \ int" where - "zorder f z = (THE n. (\h r. r>0 \ h holomorphic_on cball z r \ h z\0 - \ (\w\cball z r - {z}. f w = h w * (w-z) powr (of_int n) - \ h w \0)))" - -definition\<^marker>\tag important\ zor_poly - ::"[complex \ complex, complex] \ complex \ complex" where - "zor_poly f z = (SOME h. \r. r > 0 \ h holomorphic_on cball z r \ h z \ 0 - \ (\w\cball z r - {z}. f w = h w * (w - z) powr (zorder f z) - \ h w \0))" - -lemma zorder_exist: - fixes f::"complex \ complex" and z::complex - defines "n\zorder f z" and "g\zor_poly f z" - assumes f_iso:"isolated_singularity_at f z" - and f_ness:"not_essential f z" - and f_nconst:"\\<^sub>Fw in (at z). f w\0" - shows "g z\0 \ (\r. r>0 \ g holomorphic_on cball z r - \ (\w\cball z r - {z}. f w = g w * (w-z) powr n \ g w \0))" -proof - - define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 - \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" - have "\!n. \g r. P n g r" - using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto - then have "\g r. P n g r" - unfolding n_def P_def zorder_def - by (drule_tac theI',argo) - then have "\r. P n g r" - unfolding P_def zor_poly_def g_def n_def - by (drule_tac someI_ex,argo) - then obtain r1 where "P n g r1" by auto - then show ?thesis unfolding P_def by auto -qed - -lemma - fixes f::"complex \ complex" and z::complex - assumes f_iso:"isolated_singularity_at f z" - and f_ness:"not_essential f z" - and f_nconst:"\\<^sub>Fw in (at z). f w\0" - shows zorder_inverse: "zorder (\w. inverse (f w)) z = - zorder f z" - and zor_poly_inverse: "\\<^sub>Fw in (at z). zor_poly (\w. inverse (f w)) z w - = inverse (zor_poly f z w)" -proof - - define vf where "vf = (\w. inverse (f w))" - define fn vfn where - "fn = zorder f z" and "vfn = zorder vf z" - define fp vfp where - "fp = zor_poly f z" and "vfp = zor_poly vf z" - - obtain fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" - "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" - using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def] - by auto - have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" - and fr_nz: "inverse (fp w)\0" - when "w\ball z fr - {z}" for w - proof - - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" - using fr(2)[rule_format,of w] that by auto - then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\0" - unfolding vf_def by (auto simp add:powr_minus) - qed - obtain vfr where [simp]:"vfp z \ 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" - "(\w\cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0)" - proof - - have "isolated_singularity_at vf z" - using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . - moreover have "not_essential vf z" - using not_essential_inverse[OF f_ness f_iso] unfolding vf_def . - moreover have "\\<^sub>F w in at z. vf w \ 0" - using f_nconst unfolding vf_def by (auto elim:frequently_elim1) - ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto - qed - - - define r1 where "r1 = min fr vfr" - have "r1>0" using \fr>0\ \vfr>0\ unfolding r1_def by simp - show "vfn = - fn" - apply (rule holomorphic_factor_unique[of r1 vfp z "\w. inverse (fp w)" vf]) - subgoal using \r1>0\ by simp - subgoal by simp - subgoal by simp - subgoal - proof (rule ballI) - fix w assume "w \ ball z r1 - {z}" - then have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" unfolding r1_def by auto - from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] - show "vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0 - \ vf w = inverse (fp w) * (w - z) powr of_int (- fn) \ inverse (fp w) \ 0" by auto - qed - subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) - subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros) - done - - have "vfp w = inverse (fp w)" when "w\ball z r1-{z}" for w - proof - - have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" "w\z" using that unfolding r1_def by auto - from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \vfn = - fn\ \w\z\ - show ?thesis by auto - qed - then show "\\<^sub>Fw in (at z). vfp w = inverse (fp w)" - unfolding eventually_at using \r1>0\ - apply (rule_tac x=r1 in exI) - by (auto simp add:dist_commute) -qed - -lemma - fixes f g::"complex \ complex" and z::complex - assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" - and f_ness:"not_essential f z" and g_ness:"not_essential g z" - and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" - shows zorder_times:"zorder (\w. f w * g w) z = zorder f z + zorder g z" and - zor_poly_times:"\\<^sub>Fw in (at z). zor_poly (\w. f w * g w) z w - = zor_poly f z w *zor_poly g z w" -proof - - define fg where "fg = (\w. f w * g w)" - define fn gn fgn where - "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z" - define fp gp fgp where - "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z" - have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" - using fg_nconst by (auto elim!:frequently_elim1) - obtain fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" - "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" - using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto - obtain gr where [simp]:"gp z \ 0" and "gr > 0" - and gr: "gp holomorphic_on cball z gr" - "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" - using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto - define r1 where "r1=min fr gr" - have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto - have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" - when "w\ball z r1 - {z}" for w - proof - - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" - using fr(2)[rule_format,of w] that unfolding r1_def by auto - moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" - using gr(2)[rule_format, of w] that unfolding r1_def by auto - ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" - unfolding fg_def by (auto simp add:powr_add) - qed - - obtain fgr where [simp]:"fgp z \ 0" and "fgr > 0" - and fgr: "fgp holomorphic_on cball z fgr" - "\w\cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0" - proof - - have "fgp z \ 0 \ (\r>0. fgp holomorphic_on cball z r - \ (\w\cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0))" - apply (rule zorder_exist[of fg z, folded fgn_def fgp_def]) - subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] . - subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] . - subgoal unfolding fg_def using fg_nconst . - done - then show ?thesis using that by blast - qed - define r2 where "r2 = min fgr r1" - have "r2>0" using \r1>0\ \fgr>0\ unfolding r2_def by simp - show "fgn = fn + gn " - apply (rule holomorphic_factor_unique[of r2 fgp z "\w. fp w * gp w" fg]) - subgoal using \r2>0\ by simp - subgoal by simp - subgoal by simp - subgoal - proof (rule ballI) - fix w assume "w \ ball z r2 - {z}" - then have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" unfolding r2_def by auto - from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] - show "fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0 - \ fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \ fp w * gp w \ 0" by auto - qed - subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) - subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) - done - - have "fgp w = fp w *gp w" when "w\ball z r2-{z}" for w - proof - - have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" "w\z" using that unfolding r2_def by auto - from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \fgn = fn + gn\ \w\z\ - show ?thesis by auto - qed - then show "\\<^sub>Fw in (at z). fgp w = fp w * gp w" - using \r2>0\ unfolding eventually_at by (auto simp add:dist_commute) -qed - -lemma - fixes f g::"complex \ complex" and z::complex - assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" - and f_ness:"not_essential f z" and g_ness:"not_essential g z" - and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" - shows zorder_divide:"zorder (\w. f w / g w) z = zorder f z - zorder g z" and - zor_poly_divide:"\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w - = zor_poly f z w / zor_poly g z w" -proof - - have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" - using fg_nconst by (auto elim!:frequently_elim1) - define vg where "vg=(\w. inverse (g w))" - have "zorder (\w. f w * vg w) z = zorder f z + zorder vg z" - apply (rule zorder_times[OF f_iso _ f_ness,of vg]) - subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . - subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . - subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) - done - then show "zorder (\w. f w / g w) z = zorder f z - zorder g z" - using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def - by (auto simp add:field_simps) - - have "\\<^sub>F w in at z. zor_poly (\w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w" - apply (rule zor_poly_times[OF f_iso _ f_ness,of vg]) - subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . - subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . - subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) - done - then show "\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" - using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def - apply eventually_elim - by (auto simp add:field_simps) -qed - -lemma zorder_exist_zero: - fixes f::"complex \ complex" and z::complex - defines "n\zorder f z" and "g\zor_poly f z" - assumes holo: "f holomorphic_on s" and - "open s" "connected s" "z\s" - and non_const: "\w\s. f w \ 0" - shows "(if f z=0 then n > 0 else n=0) \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r - \ (\w\cball z r. f w = g w * (w-z) ^ nat n \ g w \0))" -proof - - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" - "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" - proof - - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r - \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" - proof (rule zorder_exist[of f z,folded g_def n_def]) - show "isolated_singularity_at f z" unfolding isolated_singularity_at_def - using holo assms(4,6) - by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE) - show "not_essential f z" unfolding not_essential_def - using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on - by fastforce - have "\\<^sub>F w in at z. f w \ 0 \ w\s" - proof - - obtain w where "w\s" "f w\0" using non_const by auto - then show ?thesis - by (rule non_zero_neighbour_alt[OF holo \open s\ \connected s\ \z\s\]) - qed - then show "\\<^sub>F w in at z. f w \ 0" - apply (elim eventually_frequentlyE) - by auto - qed - then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" - "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" - by auto - obtain r2 where r2: "r2>0" "cball z r2 \ s" - using assms(4,6) open_contains_cball_eq by blast - define r3 where "r3=min r1 r2" - have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto - moreover have "g holomorphic_on cball z r3" - using r1(1) unfolding r3_def by auto - moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" - using r1(2) unfolding r3_def by auto - ultimately show ?thesis using that[of r3] \g z\0\ by auto - qed - - have if_0:"if f z=0 then n > 0 else n=0" - proof - - have "f\ z \ f z" - by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on) - then have "(\w. g w * (w - z) powr of_int n) \z\ f z" - apply (elim Lim_transform_within_open[where s="ball z r"]) - using r by auto - moreover have "g \z\g z" - by (metis (mono_tags, lifting) open_ball at_within_open_subset - ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE) - ultimately have "(\w. (g w * (w - z) powr of_int n) / g w) \z\ f z/g z" - apply (rule_tac tendsto_divide) - using \g z\0\ by auto - then have powr_tendsto:"(\w. (w - z) powr of_int n) \z\ f z/g z" - apply (elim Lim_transform_within_open[where s="ball z r"]) - using r by auto - - have ?thesis when "n\0" "f z=0" - proof - - have "(\w. (w - z) ^ nat n) \z\ f z/g z" - using powr_tendsto - apply (elim Lim_transform_within[where d=r]) - by (auto simp add: powr_of_int \n\0\ \r>0\) - then have *:"(\w. (w - z) ^ nat n) \z\ 0" using \f z=0\ by simp - moreover have False when "n=0" - proof - - have "(\w. (w - z) ^ nat n) \z\ 1" - using \n=0\ by auto - then show False using * using LIM_unique zero_neq_one by blast - qed - ultimately show ?thesis using that by fastforce - qed - moreover have ?thesis when "n\0" "f z\0" - proof - - have False when "n>0" - proof - - have "(\w. (w - z) ^ nat n) \z\ f z/g z" - using powr_tendsto - apply (elim Lim_transform_within[where d=r]) - by (auto simp add: powr_of_int \n\0\ \r>0\) - moreover have "(\w. (w - z) ^ nat n) \z\ 0" - using \n>0\ by (auto intro!:tendsto_eq_intros) - ultimately show False using \f z\0\ \g z\0\ using LIM_unique divide_eq_0_iff by blast - qed - then show ?thesis using that by force - qed - moreover have False when "n<0" - proof - - have "(\w. inverse ((w - z) ^ nat (- n))) \z\ f z/g z" - "(\w.((w - z) ^ nat (- n))) \z\ 0" - subgoal using powr_tendsto powr_of_int that - by (elim Lim_transform_within_open[where s=UNIV],auto) - subgoal using that by (auto intro!:tendsto_eq_intros) - done - from tendsto_mult[OF this,simplified] - have "(\x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \z\ 0" . - then have "(\x. 1::complex) \z\ 0" - by (elim Lim_transform_within_open[where s=UNIV],auto) - then show False using LIM_const_eq by fastforce - qed - ultimately show ?thesis by fastforce - qed - moreover have "f w = g w * (w-z) ^ nat n \ g w \0" when "w\cball z r" for w - proof (cases "w=z") - case True - then have "f \z\f w" - using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce - then have "(\w. g w * (w-z) ^ nat n) \z\f w" - proof (elim Lim_transform_within[OF _ \r>0\]) - fix x assume "0 < dist x z" "dist x z < r" - then have "x \ cball z r - {z}" "x\z" - unfolding cball_def by (auto simp add: dist_commute) - then have "f x = g x * (x - z) powr of_int n" - using r(4)[rule_format,of x] by simp - also have "... = g x * (x - z) ^ nat n" - apply (subst powr_of_int) - using if_0 \x\z\ by (auto split:if_splits) - finally show "f x = g x * (x - z) ^ nat n" . - qed - moreover have "(\w. g w * (w-z) ^ nat n) \z\ g w * (w-z) ^ nat n" - using True apply (auto intro!:tendsto_eq_intros) - by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball - continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that) - ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast - then show ?thesis using \g z\0\ True by auto - next - case False - then have "f w = g w * (w - z) powr of_int n \ g w \ 0" - using r(4) that by auto - then show ?thesis using False if_0 powr_of_int by (auto split:if_splits) - qed - ultimately show ?thesis using r by auto -qed - -lemma zorder_exist_pole: - fixes f::"complex \ complex" and z::complex - defines "n\zorder f z" and "g\zor_poly f z" - assumes holo: "f holomorphic_on s-{z}" and - "open s" "z\s" - and "is_pole f z" - shows "n < 0 \ g z\0 \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r - \ (\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0))" -proof - - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" - "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" - proof - - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r - \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" - proof (rule zorder_exist[of f z,folded g_def n_def]) - show "isolated_singularity_at f z" unfolding isolated_singularity_at_def - using holo assms(4,5) - by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) - show "not_essential f z" unfolding not_essential_def - using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on - by fastforce - from non_zero_neighbour_pole[OF \is_pole f z\] show "\\<^sub>F w in at z. f w \ 0" - apply (elim eventually_frequentlyE) - by auto - qed - then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" - "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" - by auto - obtain r2 where r2: "r2>0" "cball z r2 \ s" - using assms(4,5) open_contains_cball_eq by metis - define r3 where "r3=min r1 r2" - have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto - moreover have "g holomorphic_on cball z r3" - using r1(1) unfolding r3_def by auto - moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" - using r1(2) unfolding r3_def by auto - ultimately show ?thesis using that[of r3] \g z\0\ by auto - qed - - have "n<0" - proof (rule ccontr) - assume " \ n < 0" - define c where "c=(if n=0 then g z else 0)" - have [simp]:"g \z\ g z" - by (metis open_ball at_within_open ball_subset_cball centre_in_ball - continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) ) - have "\\<^sub>F x in at z. f x = g x * (x - z) ^ nat n" - unfolding eventually_at_topological - apply (rule_tac exI[where x="ball z r"]) - using r powr_of_int \\ n < 0\ by auto - moreover have "(\x. g x * (x - z) ^ nat n) \z\c" - proof (cases "n=0") - case True - then show ?thesis unfolding c_def by simp - next - case False - then have "(\x. (x - z) ^ nat n) \z\ 0" using \\ n < 0\ - by (auto intro!:tendsto_eq_intros) - from tendsto_mult[OF _ this,of g "g z",simplified] - show ?thesis unfolding c_def using False by simp - qed - ultimately have "f \z\c" using tendsto_cong by fast - then show False using \is_pole f z\ at_neq_bot not_tendsto_and_filterlim_at_infinity - unfolding is_pole_def by blast - qed - moreover have "\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0" - using r(4) \n<0\ powr_of_int - by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le) - ultimately show ?thesis using r(1-3) \g z\0\ by auto -qed - -lemma zorder_eqI: - assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" - assumes fg_eq:"\w. \w \ s;w\z\ \ f w = g w * (w - z) powr n" - shows "zorder f z = n" -proof - - have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact - moreover have "open (-{0::complex})" by auto - ultimately have "open ((g -` (-{0})) \ s)" - unfolding continuous_on_open_vimage[OF \open s\] by blast - moreover from assms have "z \ (g -` (-{0})) \ s" by auto - ultimately obtain r where r: "r > 0" "cball z r \ s \ (g -` (-{0}))" - unfolding open_contains_cball by blast - - let ?gg= "(\w. g w * (w - z) powr n)" - define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 - \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" - have "P n g r" - unfolding P_def using r assms(3,4,5) by auto - then have "\g r. P n g r" by auto - moreover have unique: "\!n. \g r. P n g r" unfolding P_def - proof (rule holomorphic_factor_puncture) - have "ball z r-{z} \ s" using r using ball_subset_cball by blast - then have "?gg holomorphic_on ball z r-{z}" - using \g holomorphic_on s\ r by (auto intro!: holomorphic_intros) - then have "f holomorphic_on ball z r - {z}" - apply (elim holomorphic_transform) - using fg_eq \ball z r-{z} \ s\ by auto - then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def - using analytic_on_open open_delete r(1) by blast - next - have "not_essential ?gg z" - proof (intro singularity_intros) - show "not_essential g z" - by (meson \continuous_on s g\ assms(1) assms(2) continuous_on_eq_continuous_at - isCont_def not_essential_def) - show " \\<^sub>F w in at z. w - z \ 0" by (simp add: eventually_at_filter) - then show "LIM w at z. w - z :> at 0" - unfolding filterlim_at by (auto intro:tendsto_eq_intros) - show "isolated_singularity_at g z" - by (meson Diff_subset open_ball analytic_on_holomorphic - assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE) - qed - then show "not_essential f z" - apply (elim not_essential_transform) - unfolding eventually_at using assms(1,2) assms(5)[symmetric] - by (metis dist_commute mem_ball openE subsetCE) - show "\\<^sub>F w in at z. f w \ 0" unfolding frequently_at - proof (rule,rule) - fix d::real assume "0 < d" - define z' where "z'=z+min d r / 2" - have "z' \ z" " dist z' z < d " - unfolding z'_def using \d>0\ \r>0\ - by (auto simp add:dist_norm) - moreover have "f z' \ 0" - proof (subst fg_eq[OF _ \z'\z\]) - have "z' \ cball z r" unfolding z'_def using \r>0\ \d>0\ by (auto simp add:dist_norm) - then show " z' \ s" using r(2) by blast - show "g z' * (z' - z) powr of_int n \ 0" - using P_def \P n g r\ \z' \ cball z r\ calculation(1) by auto - qed - ultimately show "\x\UNIV. x \ z \ dist x z < d \ f x \ 0" by auto - qed - qed - ultimately have "(THE n. \g r. P n g r) = n" - by (rule_tac the1_equality) - then show ?thesis unfolding zorder_def P_def by blast -qed - -lemma residue_pole_order: - fixes f::"complex \ complex" and z::complex - defines "n \ nat (- zorder f z)" and "h \ zor_poly f z" - assumes f_iso:"isolated_singularity_at f z" - and pole:"is_pole f z" - shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" -proof - - define g where "g \ \x. if x=z then 0 else inverse (f x)" - obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" - using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast - obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ ball z e" and h_holo: "h holomorphic_on cball z r" - and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" - proof - - obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" - "(\w\cball z r - {z}. f w = h w / (w - z) ^ n \ h w \ 0)" - using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\,folded n_def h_def] by auto - have "n>0" using \zorder f z < 0\ unfolding n_def by simp - moreover have "(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" - using \h z\0\ r(6) by blast - ultimately show ?thesis using r(3,4,5) that by blast - qed - have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" - using h_divide by simp - define c where "c \ 2 * pi * \" - define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" - define h' where "h' \ \u. h u / (u - z) ^ n" - have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" - unfolding h'_def - proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", - folded c_def Suc_pred'[OF \n>0\]]) - show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp - show "h holomorphic_on ball z r" using h_holo by auto - show " z \ ball z r" using \r>0\ by auto - qed - then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto - then have "(f has_contour_integral c * der_f) (circlepath z r)" - proof (elim has_contour_integral_eq) - fix x assume "x \ path_image (circlepath z r)" - hence "x\cball z r - {z}" using \r>0\ by auto - then show "h' x = f x" using h_divide unfolding h'_def by auto - qed - moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" - using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] - unfolding c_def by simp - ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast - hence "der_f = residue f z" unfolding c_def by auto - thus ?thesis unfolding der_f_def by auto -qed - -lemma simple_zeroI: - assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" - assumes "\w. w \ s \ f w = g w * (w - z)" - shows "zorder f z = 1" - using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto) - -lemma higher_deriv_power: - shows "(deriv ^^ j) (\w. (w - z) ^ n) w = - pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)" -proof (induction j arbitrary: w) - case 0 - thus ?case by auto -next - case (Suc j w) - have "(deriv ^^ Suc j) (\w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\w. (w - z) ^ n)) w" - by simp - also have "(deriv ^^ j) (\w. (w - z) ^ n) = - (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))" - using Suc by (intro Suc.IH ext) - also { - have "(\ has_field_derivative of_nat (n - j) * - pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)" - using Suc.prems by (auto intro!: derivative_eq_intros) - also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = - pochhammer (of_nat (Suc n - Suc j)) (Suc j)" - by (cases "Suc j \ n", subst pochhammer_rec) - (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left) - finally have "deriv (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w = - \ * (w - z) ^ (n - Suc j)" - by (rule DERIV_imp_deriv) - } - finally show ?case . -qed - -lemma zorder_zero_eqI: - assumes f_holo:"f holomorphic_on s" and "open s" "z \ s" - assumes zero: "\i. i < nat n \ (deriv ^^ i) f z = 0" - assumes nz: "(deriv ^^ nat n) f z \ 0" and "n\0" - shows "zorder f z = n" -proof - - obtain r where [simp]:"r>0" and "ball z r \ s" - using \open s\ \z\s\ openE by blast - have nz':"\w\ball z r. f w \ 0" - proof (rule ccontr) - assume "\ (\w\ball z r. f w \ 0)" - then have "eventually (\u. f u = 0) (nhds z)" - using \r>0\ unfolding eventually_nhds - apply (rule_tac x="ball z r" in exI) - by auto - then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\_. 0) z" - by (intro higher_deriv_cong_ev) auto - also have "(deriv ^^ nat n) (\_. 0) z = 0" - by (induction n) simp_all - finally show False using nz by contradiction - qed - - define zn g where "zn = zorder f z" and "g = zor_poly f z" - obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and - [simp]:"e>0" and "cball z e \ ball z r" and - g_holo:"g holomorphic_on cball z e" and - e_fac:"(\w\cball z e. f w = g w * (w - z) ^ nat zn \ g w \ 0)" - proof - - have "f holomorphic_on ball z r" - using f_holo \ball z r \ s\ by auto - from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def] - show ?thesis by blast - qed - from this(1,2,5) have "zn\0" "g z\0" - subgoal by (auto split:if_splits) - subgoal using \0 < e\ ball_subset_cball centre_in_ball e_fac by blast - done - - define A where "A = (\i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)" - have deriv_A:"(deriv ^^ i) f z = (if zn \ int i then A i else 0)" for i - proof - - have "eventually (\w. w \ ball z e) (nhds z)" - using \cball z e \ ball z r\ \e>0\ by (intro eventually_nhds_in_open) auto - hence "eventually (\w. f w = (w - z) ^ (nat zn) * g w) (nhds z)" - apply eventually_elim - by (use e_fac in auto) - hence "(deriv ^^ i) f z = (deriv ^^ i) (\w. (w - z) ^ nat zn * g w) z" - by (intro higher_deriv_cong_ev) auto - also have "\ = (\j=0..i. of_nat (i choose j) * - (deriv ^^ j) (\w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)" - using g_holo \e>0\ - by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros) - also have "\ = (\j=0..i. if j = nat zn then - of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" - proof (intro sum.cong refl, goal_cases) - case (1 j) - have "(deriv ^^ j) (\w. (w - z) ^ nat zn) z = - pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)" - by (subst higher_deriv_power) auto - also have "\ = (if j = nat zn then fact j else 0)" - by (auto simp: not_less pochhammer_0_left pochhammer_fact) - also have "of_nat (i choose j) * \ * (deriv ^^ (i - j)) g z = - (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) - * (deriv ^^ (i - nat zn)) g z else 0)" - by simp - finally show ?case . - qed - also have "\ = (if i \ zn then A i else 0)" - by (auto simp: A_def) - finally show "(deriv ^^ i) f z = \" . - qed - - have False when "nn\0\ by auto - with nz show False by auto - qed - moreover have "n\zn" - proof - - have "g z \ 0" using e_fac[rule_format,of z] \e>0\ by simp - then have "(deriv ^^ nat zn) f z \ 0" - using deriv_A[of "nat zn"] by(auto simp add:A_def) - then have "nat zn \ nat n" using zero[of "nat zn"] by linarith - moreover have "zn\0" using e_if by (auto split:if_splits) - ultimately show ?thesis using nat_le_eq_zle by blast - qed - ultimately show ?thesis unfolding zn_def by fastforce -qed - -lemma - assumes "eventually (\z. f z = g z) (at z)" "z = z'" - shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'" -proof - - define P where "P = (\ff n h r. 0 < r \ h holomorphic_on cball z r \ h z\0 - \ (\w\cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \ h w\0))" - have "(\r. P f n h r) = (\r. P g n h r)" for n h - proof - - have *: "\r. P g n h r" if "\r. P f n h r" and "eventually (\x. f x = g x) (at z)" for f g - proof - - from that(1) obtain r1 where r1_P:"P f n h r1" by auto - from that(2) obtain r2 where "r2>0" and r2_dist:"\x. x \ z \ dist x z \ r2 \ f x = g x" - unfolding eventually_at_le by auto - define r where "r=min r1 r2" - have "r>0" "h z\0" using r1_P \r2>0\ unfolding r_def P_def by auto - moreover have "h holomorphic_on cball z r" - using r1_P unfolding P_def r_def by auto - moreover have "g w = h w * (w - z) powr of_int n \ h w \ 0" when "w\cball z r - {z}" for w - proof - - have "f w = h w * (w - z) powr of_int n \ h w \ 0" - using r1_P that unfolding P_def r_def by auto - moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def - by (simp add: dist_commute) - ultimately show ?thesis by simp - qed - ultimately show ?thesis unfolding P_def by auto - qed - from assms have eq': "eventually (\z. g z = f z) (at z)" - by (simp add: eq_commute) - show ?thesis - by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']]) - qed - then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'" - using \z=z'\ unfolding P_def zorder_def zor_poly_def by auto -qed - -lemma zorder_nonzero_div_power: - assumes "open s" "z \ s" "f holomorphic_on s" "f z \ 0" "n > 0" - shows "zorder (\w. f w / (w - z) ^ n) z = - n" - apply (rule zorder_eqI[OF \open s\ \z\s\ \f holomorphic_on s\ \f z\0\]) - apply (subst powr_of_int) - using \n>0\ by (auto simp add:field_simps) - -lemma zor_poly_eq: - assumes "isolated_singularity_at f z" "not_essential f z" "\\<^sub>F w in at z. f w \ 0" - shows "eventually (\w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)" -proof - - obtain r where r:"r>0" - "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))" - using zorder_exist[OF assms] by blast - then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" - by (auto simp: field_simps powr_minus) - have "eventually (\w. w \ ball z r - {z}) (at z)" - using r eventually_at_ball'[of r z UNIV] by auto - thus ?thesis by eventually_elim (insert *, auto) -qed - -lemma zor_poly_zero_eq: - assumes "f holomorphic_on s" "open s" "connected s" "z \ s" "\w\s. f w \ 0" - shows "eventually (\w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)" -proof - - obtain r where r:"r>0" - "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))" - using zorder_exist_zero[OF assms] by auto - then have *: "\w\ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" - by (auto simp: field_simps powr_minus) - have "eventually (\w. w \ ball z r - {z}) (at z)" - using r eventually_at_ball'[of r z UNIV] by auto - thus ?thesis by eventually_elim (insert *, auto) -qed - -lemma zor_poly_pole_eq: - assumes f_iso:"isolated_singularity_at f z" "is_pole f z" - shows "eventually (\w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)" -proof - - obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" - using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast - obtain r where r:"r>0" - "(\w\cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))" - using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\] by auto - then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" - by (auto simp: field_simps) - have "eventually (\w. w \ ball z r - {z}) (at z)" - using r eventually_at_ball'[of r z UNIV] by auto - thus ?thesis by eventually_elim (insert *, auto) -qed - -lemma zor_poly_eqI: - fixes f :: "complex \ complex" and z0 :: complex - defines "n \ zorder f z0" - assumes "isolated_singularity_at f z0" "not_essential f z0" "\\<^sub>F w in at z0. f w \ 0" - assumes lim: "((\x. f (g x) * (g x - z0) powr - n) \ c) F" - assumes g: "filterlim g (at z0) F" and "F \ bot" - shows "zor_poly f z0 z0 = c" -proof - - from zorder_exist[OF assms(2-4)] obtain r where - r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" - "\w. w \ cball z0 r - {z0} \ f w = zor_poly f z0 w * (w - z0) powr n" - unfolding n_def by blast - from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" - using eventually_at_ball'[of r z0 UNIV] by auto - hence "eventually (\w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)" - by eventually_elim (insert r, auto simp: field_simps powr_minus) - moreover have "continuous_on (ball z0 r) (zor_poly f z0)" - using r by (intro holomorphic_on_imp_continuous_on) auto - with r(1,2) have "isCont (zor_poly f z0) z0" - by (auto simp: continuous_on_eq_continuous_at) - hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" - unfolding isCont_def . - ultimately have "((\w. f w * (w - z0) powr - n) \ zor_poly f z0 z0) (at z0)" - by (blast intro: Lim_transform_eventually) - hence "((\x. f (g x) * (g x - z0) powr - n) \ zor_poly f z0 z0) F" - by (rule filterlim_compose[OF _ g]) - from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . -qed - -lemma zor_poly_zero_eqI: - fixes f :: "complex \ complex" and z0 :: complex - defines "n \ zorder f z0" - assumes "f holomorphic_on A" "open A" "connected A" "z0 \ A" "\z\A. f z \ 0" - assumes lim: "((\x. f (g x) / (g x - z0) ^ nat n) \ c) F" - assumes g: "filterlim g (at z0) F" and "F \ bot" - shows "zor_poly f z0 z0 = c" -proof - - from zorder_exist_zero[OF assms(2-6)] obtain r where - r: "r > 0" "cball z0 r \ A" "zor_poly f z0 holomorphic_on cball z0 r" - "\w. w \ cball z0 r \ f w = zor_poly f z0 w * (w - z0) ^ nat n" - unfolding n_def by blast - from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" - using eventually_at_ball'[of r z0 UNIV] by auto - hence "eventually (\w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)" - by eventually_elim (insert r, auto simp: field_simps) - moreover have "continuous_on (ball z0 r) (zor_poly f z0)" - using r by (intro holomorphic_on_imp_continuous_on) auto - with r(1,2) have "isCont (zor_poly f z0) z0" - by (auto simp: continuous_on_eq_continuous_at) - hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" - unfolding isCont_def . - ultimately have "((\w. f w / (w - z0) ^ nat n) \ zor_poly f z0 z0) (at z0)" - by (blast intro: Lim_transform_eventually) - hence "((\x. f (g x) / (g x - z0) ^ nat n) \ zor_poly f z0 z0) F" - by (rule filterlim_compose[OF _ g]) - from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . -qed - -lemma zor_poly_pole_eqI: - fixes f :: "complex \ complex" and z0 :: complex - defines "n \ zorder f z0" - assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0" - assumes lim: "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ c) F" - assumes g: "filterlim g (at z0) F" and "F \ bot" - shows "zor_poly f z0 z0 = c" -proof - - obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" - proof - - have "\\<^sub>F w in at z0. f w \ 0" - using non_zero_neighbour_pole[OF \is_pole f z0\] by (auto elim:eventually_frequentlyE) - moreover have "not_essential f z0" unfolding not_essential_def using \is_pole f z0\ by simp - ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto - qed - from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" - using eventually_at_ball'[of r z0 UNIV] by auto - have "eventually (\w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)" - using zor_poly_pole_eq[OF f_iso \is_pole f z0\] unfolding n_def . - moreover have "continuous_on (ball z0 r) (zor_poly f z0)" - using r by (intro holomorphic_on_imp_continuous_on) auto - with r(1,2) have "isCont (zor_poly f z0) z0" - by (auto simp: continuous_on_eq_continuous_at) - hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" - unfolding isCont_def . - ultimately have "((\w. f w * (w - z0) ^ nat (-n)) \ zor_poly f z0 z0) (at z0)" - by (blast intro: Lim_transform_eventually) - hence "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ zor_poly f z0 z0) F" - by (rule filterlim_compose[OF _ g]) - from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . -qed - -lemma residue_simple_pole: - assumes "isolated_singularity_at f z0" - assumes "is_pole f z0" "zorder f z0 = - 1" - shows "residue f z0 = zor_poly f z0 z0" - using assms by (subst residue_pole_order) simp_all - -lemma residue_simple_pole_limit: - assumes "isolated_singularity_at f z0" - assumes "is_pole f z0" "zorder f z0 = - 1" - assumes "((\x. f (g x) * (g x - z0)) \ c) F" - assumes "filterlim g (at z0) F" "F \ bot" - shows "residue f z0 = c" -proof - - have "residue f z0 = zor_poly f z0 z0" - by (rule residue_simple_pole assms)+ - also have "\ = c" - apply (rule zor_poly_pole_eqI) - using assms by auto - finally show ?thesis . -qed - -lemma lhopital_complex_simple: - assumes "(f has_field_derivative f') (at z)" - assumes "(g has_field_derivative g') (at z)" - assumes "f z = 0" "g z = 0" "g' \ 0" "f' / g' = c" - shows "((\w. f w / g w) \ c) (at z)" -proof - - have "eventually (\w. w \ z) (at z)" - by (auto simp: eventually_at_filter) - hence "eventually (\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)" - by eventually_elim (simp add: assms field_split_simps) - moreover have "((\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \ f' / g') (at z)" - by (intro tendsto_divide has_field_derivativeD assms) - ultimately have "((\w. f w / g w) \ f' / g') (at z)" - by (blast intro: Lim_transform_eventually) - with assms show ?thesis by simp -qed - -lemma - assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" - and "open s" "connected s" "z \ s" - assumes g_deriv:"(g has_field_derivative g') (at z)" - assumes "f z \ 0" "g z = 0" "g' \ 0" - shows porder_simple_pole_deriv: "zorder (\w. f w / g w) z = - 1" - and residue_simple_pole_deriv: "residue (\w. f w / g w) z = f z / g'" -proof - - have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z" - using isolated_singularity_at_holomorphic[OF _ \open s\ \z\s\] f_holo g_holo - by (meson Diff_subset holomorphic_on_subset)+ - have [simp]:"not_essential f z" "not_essential g z" - unfolding not_essential_def using f_holo g_holo assms(3,5) - by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+ - have g_nconst:"\\<^sub>F w in at z. g w \0 " - proof (rule ccontr) - assume "\ (\\<^sub>F w in at z. g w \ 0)" - then have "\\<^sub>F w in nhds z. g w = 0" - unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ - by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball) - then have "deriv g z = deriv (\_. 0) z" - by (intro deriv_cong_ev) auto - then have "deriv g z = 0" by auto - then have "g' = 0" using g_deriv DERIV_imp_deriv by blast - then show False using \g'\0\ by auto - qed - - have "zorder (\w. f w / g w) z = zorder f z - zorder g z" - proof - - have "\\<^sub>F w in at z. f w \0 \ w\s" - apply (rule non_zero_neighbour_alt) - using assms by auto - with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" - by (elim frequently_rev_mp eventually_rev_mp,auto) - then show ?thesis using zorder_divide[of f z g] by auto - qed - moreover have "zorder f z=0" - apply (rule zorder_zero_eqI[OF f_holo \open s\ \z\s\]) - using \f z\0\ by auto - moreover have "zorder g z=1" - apply (rule zorder_zero_eqI[OF g_holo \open s\ \z\s\]) - subgoal using assms(8) by auto - subgoal using DERIV_imp_deriv assms(9) g_deriv by auto - subgoal by simp - done - ultimately show "zorder (\w. f w / g w) z = - 1" by auto - - show "residue (\w. f w / g w) z = f z / g'" - proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified]) - show "zorder (\w. f w / g w) z = - 1" by fact - show "isolated_singularity_at (\w. f w / g w) z" - by (auto intro: singularity_intros) - show "is_pole (\w. f w / g w) z" - proof (rule is_pole_divide) - have "\\<^sub>F x in at z. g x \ 0" - apply (rule non_zero_neighbour) - using g_nconst by auto - moreover have "g \z\ 0" - using DERIV_isCont assms(8) continuous_at g_deriv by force - ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp - show "isCont f z" - using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on - by auto - show "f z \ 0" by fact - qed - show "filterlim id (at z) (at z)" by (simp add: filterlim_iff) - have "((\w. (f w * (w - z)) / g w) \ f z / g') (at z)" - proof (rule lhopital_complex_simple) - show "((\w. f w * (w - z)) has_field_derivative f z) (at z)" - using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo]) - show "(g has_field_derivative g') (at z)" by fact - qed (insert assms, auto) - then show "((\w. (f w / g w) * (w - z)) \ f z / g') (at z)" - by (simp add: field_split_simps) - qed -qed - -subsection \The argument principle\ - -theorem argument_principle: - fixes f::"complex \ complex" and poles s:: "complex set" - defines "pz \ {w. f w = 0 \ w \ poles}" \ \\<^term>\pz\ is the set of poles and zeros\ - assumes "open s" and - "connected s" and - f_holo:"f holomorphic_on s-poles" and - h_holo:"h holomorphic_on s" and - "valid_path g" and - loop:"pathfinish g = pathstart g" and - path_img:"path_image g \ s - pz" and - homo:"\z. (z \ s) \ winding_number g z = 0" and - finite:"finite pz" and - poles:"\p\poles. is_pole f p" - shows "contour_integral g (\x. deriv f x * h x / f x) = 2 * pi * \ * - (\p\pz. winding_number g p * h p * zorder f p)" - (is "?L=?R") -proof - - define c where "c \ 2 * complex_of_real pi * \ " - define ff where "ff \ (\x. deriv f x * h x / f x)" - define cont where "cont \ \ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" - define avoid where "avoid \ \p e. \w\cball p e. w \ s \ (w \ p \ w \ pz)" - - have "\e>0. avoid p e \ (p\pz \ cont ff p e)" when "p\s" for p - proof - - obtain e1 where "e1>0" and e1_avoid:"avoid p e1" - using finite_cball_avoid[OF \open s\ finite] \p\s\ unfolding avoid_def by auto - have "\e2>0. cball p e2 \ ball p e1 \ cont ff p e2" when "p\pz" - proof - - define po where "po \ zorder f p" - define pp where "pp \ zor_poly f p" - define f' where "f' \ \w. pp w * (w - p) powr po" - define ff' where "ff' \ (\x. deriv f' x * h x / f' x)" - obtain r where "pp p\0" "r>0" and - "rw\cball p r-{p}. f w = pp w * (w - p) powr po \ pp w \ 0)" - proof - - have "isolated_singularity_at f p" - proof - - have "f holomorphic_on ball p e1 - {p}" - apply (intro holomorphic_on_subset[OF f_holo]) - using e1_avoid \p\pz\ unfolding avoid_def pz_def by force - then show ?thesis unfolding isolated_singularity_at_def - using \e1>0\ analytic_on_open open_delete by blast - qed - moreover have "not_essential f p" - proof (cases "is_pole f p") - case True - then show ?thesis unfolding not_essential_def by auto - next - case False - then have "p\s-poles" using \p\s\ poles unfolding pz_def by auto - moreover have "open (s-poles)" - using \open s\ - apply (elim open_Diff) - apply (rule finite_imp_closed) - using finite unfolding pz_def by simp - ultimately have "isCont f p" - using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at - by auto - then show ?thesis unfolding isCont_def not_essential_def by auto - qed - moreover have "\\<^sub>F w in at p. f w \ 0 " - proof (rule ccontr) - assume "\ (\\<^sub>F w in at p. f w \ 0)" - then have "\\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto - then obtain rr where "rr>0" "\w\ball p rr - {p}. f w =0" - unfolding eventually_at by (auto simp add:dist_commute) - then have "ball p rr - {p} \ {w\ball p rr-{p}. f w=0}" by blast - moreover have "infinite (ball p rr - {p})" using \rr>0\ using finite_imp_not_open by fastforce - ultimately have "infinite {w\ball p rr-{p}. f w=0}" using infinite_super by blast - then have "infinite pz" - unfolding pz_def infinite_super by auto - then show False using \finite pz\ by auto - qed - ultimately obtain r where "pp p \ 0" and r:"r>0" "pp holomorphic_on cball p r" - "(\w\cball p r - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" - using zorder_exist[of f p,folded po_def pp_def] by auto - define r1 where "r1=min r e1 / 2" - have "r1e1>0\ \r>0\ by auto - moreover have "r1>0" "pp holomorphic_on cball p r1" - "(\w\cball p r1 - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" - unfolding r1_def using \e1>0\ r by auto - ultimately show ?thesis using that \pp p\0\ by auto - qed - - define e2 where "e2 \ r/2" - have "e2>0" using \r>0\ unfolding e2_def by auto - define anal where "anal \ \w. deriv pp w * h w / pp w" - define prin where "prin \ \w. po * h w / (w - p)" - have "((\w. prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)" - proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) - have "ball p r \ s" - using \r avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq) - then have "cball p e2 \ s" - using \r>0\ unfolding e2_def by auto - then have "(\w. po * h w) holomorphic_on cball p e2" - using h_holo by (auto intro!: holomorphic_intros) - then show "(prin has_contour_integral c * po * h p ) (circlepath p e2)" - using Cauchy_integral_circlepath_simple[folded c_def, of "\w. po * h w"] \e2>0\ - unfolding prin_def by (auto simp add: mult.assoc) - have "anal holomorphic_on ball p r" unfolding anal_def - using pp_holo h_holo pp_po \ball p r \ s\ \pp p\0\ - by (auto intro!: holomorphic_intros) - then show "(anal has_contour_integral 0) (circlepath p e2)" - using e2_def \r>0\ - by (auto elim!: Cauchy_theorem_disc_simple) - qed - then have "cont ff' p e2" unfolding cont_def po_def - proof (elim has_contour_integral_eq) - fix w assume "w \ path_image (circlepath p e2)" - then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto - define wp where "wp \ w-p" - have "wp\0" and "pp w \0" - unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto - moreover have der_f':"deriv f' w = po * pp w * (w-p) powr (po - 1) + deriv pp w * (w-p) powr po" - proof (rule DERIV_imp_deriv) - have "(pp has_field_derivative (deriv pp w)) (at w)" - using DERIV_deriv_iff_has_field_derivative pp_holo \w\p\ - by (meson open_ball \w \ ball p r\ ball_subset_cball holomorphic_derivI holomorphic_on_subset) - then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1) - + deriv pp w * (w - p) powr of_int po) (at w)" - unfolding f'_def using \w\p\ - by (auto intro!: derivative_eq_intros DERIV_cong[OF has_field_derivative_powr_of_int]) - qed - ultimately show "prin w + anal w = ff' w" - unfolding ff'_def prin_def anal_def - apply simp - apply (unfold f'_def) - apply (fold wp_def) - apply (auto simp add:field_simps) - by (metis (no_types, lifting) diff_add_cancel mult.commute powr_add powr_to_1) - qed - then have "cont ff p e2" unfolding cont_def - proof (elim has_contour_integral_eq) - fix w assume "w \ path_image (circlepath p e2)" - then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto - have "deriv f' w = deriv f w" - proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) - show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo - by (auto intro!: holomorphic_intros) - next - have "ball p e1 - {p} \ s - poles" - using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def - by auto - then have "ball p r - {p} \ s - poles" - apply (elim dual_order.trans) - using \r by auto - then show "f holomorphic_on ball p r - {p}" using f_holo - by auto - next - show "open (ball p r - {p})" by auto - show "w \ ball p r - {p}" using \w\ball p r\ \w\p\ by auto - next - fix x assume "x \ ball p r - {p}" - then show "f' x = f x" - using pp_po unfolding f'_def by auto - qed - moreover have " f' w = f w " - using \w \ ball p r\ ball_subset_cball subset_iff pp_po \w\p\ - unfolding f'_def by auto - ultimately show "ff' w = ff w" - unfolding ff'_def ff_def by simp - qed - moreover have "cball p e2 \ ball p e1" - using \0 < r\ \r e2_def by auto - ultimately show ?thesis using \e2>0\ by auto - qed - then obtain e2 where e2:"p\pz \ e2>0 \ cball p e2 \ ball p e1 \ cont ff p e2" - by auto - define e4 where "e4 \ if p\pz then e2 else e1" - have "e4>0" using e2 \e1>0\ unfolding e4_def by auto - moreover have "avoid p e4" using e2 \e1>0\ e1_avoid unfolding e4_def avoid_def by auto - moreover have "p\pz \ cont ff p e4" - by (auto simp add: e2 e4_def) - ultimately show ?thesis by auto - qed - then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) - \ (p\pz \ cont ff p (get_e p))" - by metis - define ci where "ci \ \p. contour_integral (circlepath p (get_e p)) ff" - define w where "w \ \p. winding_number g p" - have "contour_integral g ff = (\p\pz. w p * ci p)" unfolding ci_def w_def - proof (rule Cauchy_theorem_singularities[OF \open s\ \connected s\ finite _ \valid_path g\ loop - path_img homo]) - have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto - then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo - by (auto intro!: holomorphic_intros simp add:pz_def) - next - show "\p\s. 0 < get_e p \ (\w\cball p (get_e p). w \ s \ (w \ p \ w \ pz))" - using get_e using avoid_def by blast - qed - also have "... = (\p\pz. c * w p * h p * zorder f p)" - proof (rule sum.cong[of pz pz,simplified]) - fix p assume "p \ pz" - show "w p * ci p = c * w p * h p * (zorder f p)" - proof (cases "p\s") - assume "p \ s" - have "ci p = c * h p * (zorder f p)" unfolding ci_def - apply (rule contour_integral_unique) - using get_e \p\s\ \p\pz\ unfolding cont_def by (metis mult.assoc mult.commute) - thus ?thesis by auto - next - assume "p\s" - then have "w p=0" using homo unfolding w_def by auto - then show ?thesis by auto - qed - qed - also have "... = c*(\p\pz. w p * h p * zorder f p)" - unfolding sum_distrib_left by (simp add:algebra_simps) - finally have "contour_integral g ff = c * (\p\pz. w p * h p * of_int (zorder f p))" . - then show ?thesis unfolding ff_def c_def w_def by simp -qed - -subsection \Rouche's theorem \ - -theorem Rouche_theorem: - fixes f g::"complex \ complex" and s:: "complex set" - defines "fg\(\p. f p + g p)" - defines "zeros_fg\{p. fg p = 0}" and "zeros_f\{p. f p = 0}" - assumes - "open s" and "connected s" and - "finite zeros_fg" and - "finite zeros_f" and - f_holo:"f holomorphic_on s" and - g_holo:"g holomorphic_on s" and - "valid_path \" and - loop:"pathfinish \ = pathstart \" and - path_img:"path_image \ \ s " and - path_less:"\z\path_image \. cmod(f z) > cmod(g z)" and - homo:"\z. (z \ s) \ winding_number \ z = 0" - shows "(\p\zeros_fg. winding_number \ p * zorder fg p) - = (\p\zeros_f. winding_number \ p * zorder f p)" -proof - - have path_fg:"path_image \ \ s - zeros_fg" - proof - - have False when "z\path_image \" and "f z + g z=0" for z - proof - - have "cmod (f z) > cmod (g z)" using \z\path_image \\ path_less by auto - moreover have "f z = - g z" using \f z + g z =0\ by (simp add: eq_neg_iff_add_eq_0) - then have "cmod (f z) = cmod (g z)" by auto - ultimately show False by auto - qed - then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto - qed - have path_f:"path_image \ \ s - zeros_f" - proof - - have False when "z\path_image \" and "f z =0" for z - proof - - have "cmod (g z) < cmod (f z) " using \z\path_image \\ path_less by auto - then have "cmod (g z) < 0" using \f z=0\ by auto - then show False by auto - qed - then show ?thesis unfolding zeros_f_def using path_img by auto - qed - define w where "w \ \p. winding_number \ p" - define c where "c \ 2 * complex_of_real pi * \" - define h where "h \ \p. g p / f p + 1" - obtain spikes - where "finite spikes" and spikes: "\x\{0..1} - spikes. \ differentiable at x" - using \valid_path \\ - by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) - have h_contour:"((\x. deriv h x / h x) has_contour_integral 0) \" - proof - - have outside_img:"0 \ outside (path_image (h o \))" - proof - - have "h p \ ball 1 1" when "p\path_image \" for p - proof - - have "cmod (g p/f p) <1" using path_less[rule_format,OF that] - apply (cases "cmod (f p) = 0") - by (auto simp add: norm_divide) - then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) - qed - then have "path_image (h o \) \ ball 1 1" - by (simp add: image_subset_iff path_image_compose) - moreover have " (0::complex) \ ball 1 1" by (simp add: dist_norm) - ultimately show "?thesis" - using convex_in_outside[of "ball 1 1" 0] outside_mono by blast - qed - have valid_h:"valid_path (h \ \)" - proof (rule valid_path_compose_holomorphic[OF \valid_path \\ _ _ path_f]) - show "h holomorphic_on s - zeros_f" - unfolding h_def using f_holo g_holo - by (auto intro!: holomorphic_intros simp add:zeros_f_def) - next - show "open (s - zeros_f)" using \finite zeros_f\ \open s\ finite_imp_closed - by auto - qed - have "((\z. 1/z) has_contour_integral 0) (h \ \)" - proof - - have "0 \ path_image (h \ \)" using outside_img by (simp add: outside_def) - then have "((\z. 1/z) has_contour_integral c * winding_number (h \ \) 0) (h \ \)" - using has_contour_integral_winding_number[of "h o \" 0,simplified] valid_h - unfolding c_def by auto - moreover have "winding_number (h o \) 0 = 0" - proof - - have "0 \ outside (path_image (h \ \))" using outside_img . - moreover have "path (h o \)" - using valid_h by (simp add: valid_path_imp_path) - moreover have "pathfinish (h o \) = pathstart (h o \)" - by (simp add: loop pathfinish_compose pathstart_compose) - ultimately show ?thesis using winding_number_zero_in_outside by auto - qed - ultimately show ?thesis by auto - qed - moreover have "vector_derivative (h \ \) (at x) = vector_derivative \ (at x) * deriv h (\ x)" - when "x\{0..1} - spikes" for x - proof (rule vector_derivative_chain_at_general) - show "\ differentiable at x" using that \valid_path \\ spikes by auto - next - define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" - define t where "t \ \ x" - have "f t\0" unfolding zeros_f_def t_def - by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) - moreover have "t\s" - using contra_subsetD path_image_def path_fg t_def that by fastforce - ultimately have "(h has_field_derivative der t) (at t)" - unfolding h_def der_def using g_holo f_holo \open s\ - by (auto intro!: holomorphic_derivI derivative_eq_intros) - then show "h field_differentiable at (\ x)" - unfolding t_def field_differentiable_def by blast - qed - then have " ((/) 1 has_contour_integral 0) (h \ \) - = ((\x. deriv h x / h x) has_contour_integral 0) \" - unfolding has_contour_integral - apply (intro has_integral_spike_eq[OF negligible_finite, OF \finite spikes\]) - by auto - ultimately show ?thesis by auto - qed - then have "contour_integral \ (\x. deriv h x / h x) = 0" - using contour_integral_unique by simp - moreover have "contour_integral \ (\x. deriv fg x / fg x) = contour_integral \ (\x. deriv f x / f x) - + contour_integral \ (\p. deriv h p / h p)" - proof - - have "(\p. deriv f p / f p) contour_integrable_on \" - proof (rule contour_integrable_holomorphic_simple[OF _ _ \valid_path \\ path_f]) - show "open (s - zeros_f)" using finite_imp_closed[OF \finite zeros_f\] \open s\ - by auto - then show "(\p. deriv f p / f p) holomorphic_on s - zeros_f" - using f_holo - by (auto intro!: holomorphic_intros simp add:zeros_f_def) - qed - moreover have "(\p. deriv h p / h p) contour_integrable_on \" - using h_contour - by (simp add: has_contour_integral_integrable) - ultimately have "contour_integral \ (\x. deriv f x / f x + deriv h x / h x) = - contour_integral \ (\p. deriv f p / f p) + contour_integral \ (\p. deriv h p / h p)" - using contour_integral_add[of "(\p. deriv f p / f p)" \ "(\p. deriv h p / h p)" ] - by auto - moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" - when "p\ path_image \" for p - proof - - have "fg p\0" and "f p\0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def - by auto - have "h p\0" - proof (rule ccontr) - assume "\ h p \ 0" - then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) - then have "cmod (g p/f p) = 1" by auto - moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] - apply (cases "cmod (f p) = 0") - by (auto simp add: norm_divide) - ultimately show False by auto - qed - have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def - using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \open s\] path_img that - by auto - have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" - proof - - define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" - have "p\s" using path_img that by auto - then have "(h has_field_derivative der p) (at p)" - unfolding h_def der_def using g_holo f_holo \open s\ \f p\0\ - by (auto intro!: derivative_eq_intros holomorphic_derivI) - then show ?thesis unfolding der_def using DERIV_imp_deriv by auto - qed - show ?thesis - apply (simp only:der_fg der_h) - apply (auto simp add:field_simps \h p\0\ \f p\0\ \fg p\0\) - by (auto simp add:field_simps h_def \f p\0\ fg_def) - qed - then have "contour_integral \ (\p. deriv fg p / fg p) - = contour_integral \ (\p. deriv f p / f p + deriv h p / h p)" - by (elim contour_integral_eq) - ultimately show ?thesis by auto - qed - moreover have "contour_integral \ (\x. deriv fg x / fg x) = c * (\p\zeros_fg. w p * zorder fg p)" - unfolding c_def zeros_fg_def w_def - proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo - , of _ "{}" "\_. 1",simplified]) - show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto - show "path_image \ \ s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def . - show " finite {p. fg p = 0}" using \finite zeros_fg\ unfolding zeros_fg_def . - qed - moreover have "contour_integral \ (\x. deriv f x / f x) = c * (\p\zeros_f. w p * zorder f p)" - unfolding c_def zeros_f_def w_def - proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo - , of _ "{}" "\_. 1",simplified]) - show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto - show "path_image \ \ s - {p. f p = 0}" using path_f unfolding zeros_f_def . - show " finite {p. f p = 0}" using \finite zeros_f\ unfolding zeros_f_def . - qed - ultimately have " c* (\p\zeros_fg. w p * (zorder fg p)) = c* (\p\zeros_f. w p * (zorder f p))" - by auto - then show ?thesis unfolding c_def using w_def by auto -qed - -end diff --git a/src/HOL/Analysis/Derivative.thy b/src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy +++ b/src/HOL/Analysis/Derivative.thy @@ -1,2962 +1,3475 @@ (* Title: HOL/Analysis/Derivative.thy Author: John Harrison Author: Robert Himmelmann, TU Muenchen (translation from HOL Light); tidied by LCP *) section \Derivative\ theory Derivative imports Bounded_Linear_Function Line_Segment + Convex_Euclidean_Space begin declare bounded_linear_inner_left [intro] declare has_derivative_bounded_linear[dest] subsection \Derivatives\ lemma has_derivative_add_const: "(f has_derivative f') net \ ((\x. f x + c) has_derivative f') net" by (intro derivative_eq_intros) auto subsection\<^marker>\tag unimportant\ \Derivative with composed bilinear function\ text \More explicit epsilon-delta forms.\ proposition has_derivative_within': "(f has_derivative f')(at x within s) \ bounded_linear f' \ (\e>0. \d>0. \x'\s. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" unfolding has_derivative_within Lim_within dist_norm by (simp add: diff_diff_eq) lemma has_derivative_at': "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \x'. 0 < norm (x' - x) \ norm (x' - x) < d \ norm (f x' - f x - f'(x' - x)) / norm (x' - x) < e)" using has_derivative_within' [of f f' x UNIV] by simp lemma has_derivative_componentwise_within: "(f has_derivative f') (at a within S) \ (\i \ Basis. ((\x. f x \ i) has_derivative (\x. f' x \ i)) (at a within S))" apply (simp add: has_derivative_within) apply (subst tendsto_componentwise_iff) apply (simp add: bounded_linear_componentwise_iff [symmetric] ball_conj_distrib) apply (simp add: algebra_simps) done lemma has_derivative_at_withinI: "(f has_derivative f') (at x) \ (f has_derivative f') (at x within s)" unfolding has_derivative_within' has_derivative_at' by blast lemma has_derivative_right: fixes f :: "real \ real" and y :: "real" shows "(f has_derivative ((*) y)) (at x within ({x <..} \ I)) \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x <..} \ I))" proof - have "((\t. (f t - (f x + y * (t - x))) / \t - x\) \ 0) (at x within ({x<..} \ I)) \ ((\t. (f t - f x) / (t - x) - y) \ 0) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (auto simp add: diff_divide_distrib add_divide_distrib) also have "\ \ ((\t. (f t - f x) / (t - x)) \ y) (at x within ({x<..} \ I))" by (simp add: Lim_null[symmetric]) also have "\ \ ((\t. (f x - f t) / (x - t)) \ y) (at x within ({x<..} \ I))" by (intro Lim_cong_within) (simp_all add: field_simps) finally show ?thesis by (simp add: bounded_linear_mult_right has_derivative_within) qed subsubsection \Caratheodory characterization\ lemma DERIV_caratheodory_within: "(f has_field_derivative l) (at x within S) \ (\g. (\z. f z - f x = g z * (z - x)) \ continuous (at x within S) g \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show ?rhs proof (intro exI conjI) let ?g = "(%z. if z = x then l else (f z - f x) / (z-x))" show "\z. f z - f x = ?g z * (z-x)" by simp show "continuous (at x within S) ?g" using \?lhs\ by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) show "?g x = l" by simp qed next assume ?rhs then obtain g where "(\z. f z - f x = g z * (z-x))" and "continuous (at x within S) g" and "g x = l" by blast thus ?lhs by (auto simp add: continuous_within has_field_derivative_iff cong: Lim_cong_within) qed subsection \Differentiability\ definition\<^marker>\tag important\ differentiable_on :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a set \ bool" (infix "differentiable'_on" 50) where "f differentiable_on s \ (\x\s. f differentiable (at x within s))" lemma differentiableI: "(f has_derivative f') net \ f differentiable net" unfolding differentiable_def by auto lemma differentiable_onD: "\f differentiable_on S; x \ S\ \ f differentiable (at x within S)" using differentiable_on_def by blast lemma differentiable_at_withinI: "f differentiable (at x) \ f differentiable (at x within s)" unfolding differentiable_def using has_derivative_at_withinI by blast lemma differentiable_at_imp_differentiable_on: "(\x. x \ s \ f differentiable at x) \ f differentiable_on s" by (metis differentiable_at_withinI differentiable_on_def) corollary\<^marker>\tag unimportant\ differentiable_iff_scaleR: fixes f :: "real \ 'a::real_normed_vector" shows "f differentiable F \ (\d. (f has_derivative (\x. x *\<^sub>R d)) F)" by (auto simp: differentiable_def dest: has_derivative_linear linear_imp_scaleR) lemma differentiable_on_eq_differentiable_at: "open s \ f differentiable_on s \ (\x\s. f differentiable at x)" unfolding differentiable_on_def by (metis at_within_interior interior_open) lemma differentiable_transform_within: assumes "f differentiable (at x within s)" and "0 < d" and "x \ s" and "\x'. \x'\s; dist x' x < d\ \ f x' = g x'" shows "g differentiable (at x within s)" using assms has_derivative_transform_within unfolding differentiable_def by blast lemma differentiable_on_ident [simp, derivative_intros]: "(\x. x) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_on_id [simp, derivative_intros]: "id differentiable_on S" by (simp add: id_def) lemma differentiable_on_const [simp, derivative_intros]: "(\z. c) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_mult [simp, derivative_intros]: fixes f :: "'M::real_normed_vector \ 'a::real_normed_algebra" shows "\f differentiable_on S; g differentiable_on S\ \ (\z. f z * g z) differentiable_on S" unfolding differentiable_on_def differentiable_def using differentiable_def differentiable_mult by blast lemma differentiable_on_compose: "\g differentiable_on S; f differentiable_on (g ` S)\ \ (\x. f (g x)) differentiable_on S" by (simp add: differentiable_in_compose differentiable_on_def) lemma bounded_linear_imp_differentiable_on: "bounded_linear f \ f differentiable_on S" by (simp add: differentiable_on_def bounded_linear_imp_differentiable) lemma linear_imp_differentiable_on: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" shows "linear f \ f differentiable_on S" by (simp add: differentiable_on_def linear_imp_differentiable) lemma differentiable_on_minus [simp, derivative_intros]: "f differentiable_on S \ (\z. -(f z)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_add [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z + g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_diff [simp, derivative_intros]: "\f differentiable_on S; g differentiable_on S\ \ (\z. f z - g z) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_inverse [simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" shows "f differentiable_on S \ (\x. x \ S \ f x \ 0) \ (\x. inverse (f x)) differentiable_on S" by (simp add: differentiable_on_def) lemma differentiable_on_scaleR [derivative_intros, simp]: "\f differentiable_on S; g differentiable_on S\ \ (\x. f x *\<^sub>R g x) differentiable_on S" unfolding differentiable_on_def by (blast intro: differentiable_scaleR) lemma has_derivative_sqnorm_at [derivative_intros, simp]: "((\x. (norm x)\<^sup>2) has_derivative (\x. 2 *\<^sub>R (a \ x))) (at a)" using bounded_bilinear.FDERIV [of "(\)" id id a _ id id] by (auto simp: inner_commute dot_square_norm bounded_bilinear_inner) lemma differentiable_sqnorm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "(\x. (norm x)\<^sup>2) differentiable (at a)" by (force simp add: differentiable_def intro: has_derivative_sqnorm_at) lemma differentiable_on_sqnorm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "(\x. (norm x)\<^sup>2) differentiable_on S" by (simp add: differentiable_at_imp_differentiable_on) lemma differentiable_norm_at [derivative_intros, simp]: fixes a :: "'a :: {real_normed_vector,real_inner}" shows "a \ 0 \ norm differentiable (at a)" using differentiableI has_derivative_norm by blast lemma differentiable_on_norm [derivative_intros, simp]: fixes S :: "'a :: {real_normed_vector,real_inner} set" shows "0 \ S \ norm differentiable_on S" by (metis differentiable_at_imp_differentiable_on differentiable_norm_at) subsection \Frechet derivative and Jacobian matrix\ definition "frechet_derivative f net = (SOME f'. (f has_derivative f') net)" proposition frechet_derivative_works: "f differentiable net \ (f has_derivative (frechet_derivative f net)) net" unfolding frechet_derivative_def differentiable_def unfolding some_eq_ex[of "\ f' . (f has_derivative f') net"] .. lemma linear_frechet_derivative: "f differentiable net \ linear (frechet_derivative f net)" unfolding frechet_derivative_works has_derivative_def by (auto intro: bounded_linear.linear) lemma frechet_derivative_const [simp]: "frechet_derivative (\x. c) (at a) = (\x. 0)" using differentiable_const frechet_derivative_works has_derivative_const has_derivative_unique by blast lemma frechet_derivative_id [simp]: "frechet_derivative id (at a) = id" using differentiable_def frechet_derivative_works has_derivative_id has_derivative_unique by blast lemma frechet_derivative_ident [simp]: "frechet_derivative (\x. x) (at a) = (\x. x)" by (metis eq_id_iff frechet_derivative_id) subsection \Differentiability implies continuity\ proposition differentiable_imp_continuous_within: "f differentiable (at x within s) \ continuous (at x within s) f" by (auto simp: differentiable_def intro: has_derivative_continuous) lemma differentiable_imp_continuous_on: "f differentiable_on s \ continuous_on s f" unfolding differentiable_on_def continuous_on_eq_continuous_within using differentiable_imp_continuous_within by blast lemma differentiable_on_subset: "f differentiable_on t \ s \ t \ f differentiable_on s" unfolding differentiable_on_def using differentiable_within_subset by blast lemma differentiable_on_empty: "f differentiable_on {}" unfolding differentiable_on_def by auto lemma has_derivative_continuous_on: "(\x. x \ s \ (f has_derivative f' x) (at x within s)) \ continuous_on s f" by (auto intro!: differentiable_imp_continuous_on differentiableI simp: differentiable_on_def) text \Results about neighborhoods filter.\ lemma eventually_nhds_metric_le: "eventually P (nhds a) = (\d>0. \x. dist x a \ d \ P x)" unfolding eventually_nhds_metric by (safe, rule_tac x="d / 2" in exI, auto) lemma le_nhds: "F \ nhds a \ (\S. open S \ a \ S \ eventually (\x. x \ S) F)" unfolding le_filter_def eventually_nhds by (fast elim: eventually_mono) lemma le_nhds_metric: "F \ nhds a \ (\e>0. eventually (\x. dist x a < e) F)" unfolding le_filter_def eventually_nhds_metric by (fast elim: eventually_mono) lemma le_nhds_metric_le: "F \ nhds a \ (\e>0. eventually (\x. dist x a \ e) F)" unfolding le_filter_def eventually_nhds_metric_le by (fast elim: eventually_mono) text \Several results are easier using a "multiplied-out" variant. (I got this idea from Dieudonne's proof of the chain rule).\ lemma has_derivative_within_alt: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. \d>0. \y\s. norm(y - x) < d \ norm (f y - f x - f' (y - x)) \ e * norm (y - x))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_within_alt2: "(f has_derivative f') (at x within s) \ bounded_linear f' \ (\e>0. eventually (\y. norm (f y - f x - f' (y - x)) \ e * norm (y - x)) (at x within s))" unfolding has_derivative_within filterlim_def le_nhds_metric_le eventually_filtermap eventually_at dist_norm diff_diff_eq by (force simp add: linear_0 bounded_linear.linear pos_divide_le_eq) lemma has_derivative_at_alt: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e>0. \d>0. \y. norm(y - x) < d \ norm (f y - f x - f'(y - x)) \ e * norm (y - x))" using has_derivative_within_alt[where s=UNIV] by simp subsection \The chain rule\ proposition diff_chain_within[derivative_intros]: assumes "(f has_derivative f') (at x within s)" and "(g has_derivative g') (at (f x) within (f ` s))" shows "((g \ f) has_derivative (g' \ f'))(at x within s)" using has_derivative_in_compose[OF assms] by (simp add: comp_def) lemma diff_chain_at[derivative_intros]: "(f has_derivative f') (at x) \ (g has_derivative g') (at (f x)) \ ((g \ f) has_derivative (g' \ f')) (at x)" using has_derivative_compose[of f f' x UNIV g g'] by (simp add: comp_def) lemma has_vector_derivative_within_open: "a \ S \ open S \ (f has_vector_derivative f') (at a within S) \ (f has_vector_derivative f') (at a)" by (simp only: at_within_interior interior_open) lemma field_vector_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_field_derivative g') (at (f x) within f ` S)" shows "((g \ f) has_vector_derivative (f' * g')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_diff_chain_within: assumes Df: "(f has_vector_derivative f') (at x within S)" and Dg: "(g has_derivative g') (at (f x) within f`S)" shows "((g \ f) has_vector_derivative (g' f')) (at x within S)" using diff_chain_within[OF Df[unfolded has_vector_derivative_def] Dg] linear.scaleR[OF has_derivative_linear[OF Dg]] unfolding has_vector_derivative_def o_def by (auto simp: o_def mult.commute has_vector_derivative_def) subsection\<^marker>\tag unimportant\ \Composition rules stated just for differentiability\ lemma differentiable_chain_at: "f differentiable (at x) \ g differentiable (at (f x)) \ (g \ f) differentiable (at x)" unfolding differentiable_def by (meson diff_chain_at) lemma differentiable_chain_within: "f differentiable (at x within S) \ g differentiable (at(f x) within (f ` S)) \ (g \ f) differentiable (at x within S)" unfolding differentiable_def by (meson diff_chain_within) subsection \Uniqueness of derivative\ text\<^marker>\tag important\ \ The general result is a bit messy because we need approachability of the limit point from any direction. But OK for nontrivial intervals etc. \ proposition frechet_derivative_unique_within: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes 1: "(f has_derivative f') (at x within S)" and 2: "(f has_derivative f'') (at x within S)" and S: "\i e. \i\Basis; e>0\ \ \d. 0 < \d\ \ \d\ < e \ (x + d *\<^sub>R i) \ S" shows "f' = f''" proof - note as = assms(1,2)[unfolded has_derivative_def] then interpret f': bounded_linear f' by auto from as interpret f'': bounded_linear f'' by auto have "x islimpt S" unfolding islimpt_approachable proof (intro allI impI) fix e :: real assume "e > 0" obtain d where "0 < \d\" and "\d\ < e" and "x + d *\<^sub>R (SOME i. i \ Basis) \ S" using assms(3) SOME_Basis \e>0\ by blast then show "\x'\S. x' \ x \ dist x' x < e" by (rule_tac x="x + d *\<^sub>R (SOME i. i \ Basis)" in bexI) (auto simp: dist_norm SOME_Basis nonzero_Basis) qed then have *: "netlimit (at x within S) = x" by (simp add: Lim_ident_at trivial_limit_within) show ?thesis proof (rule linear_eq_stdbasis) show "linear f'" "linear f''" unfolding linear_conv_bounded_linear using as by auto next fix i :: 'a assume i: "i \ Basis" define e where "e = norm (f' i - f'' i)" show "f' i = f'' i" proof (rule ccontr) assume "f' i \ f'' i" then have "e > 0" unfolding e_def by auto obtain d where d: "0 < d" "(\y. y\S \ 0 < dist y x \ dist y x < d \ dist ((f y - f x - f' (y - x)) /\<^sub>R norm (y - x) - (f y - f x - f'' (y - x)) /\<^sub>R norm (y - x)) (0 - 0) < e)" using tendsto_diff [OF as(1,2)[THEN conjunct2]] unfolding * Lim_within using \e>0\ by blast obtain c where c: "0 < \c\" "\c\ < d \ x + c *\<^sub>R i \ S" using assms(3) i d(1) by blast have *: "norm (- ((1 / \c\) *\<^sub>R f' (c *\<^sub>R i)) + (1 / \c\) *\<^sub>R f'' (c *\<^sub>R i)) = norm ((1 / \c\) *\<^sub>R (- (f' (c *\<^sub>R i)) + f'' (c *\<^sub>R i)))" unfolding scaleR_right_distrib by auto also have "\ = norm ((1 / \c\) *\<^sub>R (c *\<^sub>R (- (f' i) + f'' i)))" unfolding f'.scaleR f''.scaleR unfolding scaleR_right_distrib scaleR_minus_right by auto also have "\ = e" unfolding e_def using c(1) using norm_minus_cancel[of "f' i - f'' i"] by auto finally show False using c using d(2)[of "x + c *\<^sub>R i"] unfolding dist_norm unfolding f'.scaleR f''.scaleR f'.add f''.add f'.diff f''.diff scaleR_scaleR scaleR_right_diff_distrib scaleR_right_distrib using i by (auto simp: inverse_eq_divide) qed qed qed proposition frechet_derivative_unique_within_closed_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes ab: "\i. i\Basis \ a\i < b\i" and x: "x \ cbox a b" and "(f has_derivative f' ) (at x within cbox a b)" and "(f has_derivative f'') (at x within cbox a b)" shows "f' = f''" proof (rule frechet_derivative_unique_within) fix e :: real fix i :: 'a assume "e > 0" and i: "i \ Basis" then show "\d. 0 < \d\ \ \d\ < e \ x + d *\<^sub>R i \ cbox a b" proof (cases "x\i = a\i") case True with ab[of i] \e>0\ x i show ?thesis by (rule_tac x="(min (b\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) next case False moreover have "a \ i < x \ i" using False i mem_box(2) x by force moreover { have "a \ i * 2 + min (x \ i - a \ i) e \ a\i *2 + x\i - a\i" by auto also have "\ = a\i + x\i" by auto also have "\ \ 2 * (x\i)" using \a \ i < x \ i\ by auto finally have "a \ i * 2 + min (x \ i - a \ i) e \ x \ i * 2" by auto } moreover have "min (x \ i - a \ i) e \ 0" by (simp add: \0 < e\ \a \ i < x \ i\ less_eq_real_def) then have "x \ i * 2 \ b \ i * 2 + min (x \ i - a \ i) e" using i mem_box(2) x by force ultimately show ?thesis using ab[of i] \e>0\ x i by (rule_tac x="- (min (x\i - a\i) e) / 2" in exI) (auto simp add: mem_box field_simps inner_simps inner_Basis) qed qed (use assms in auto) lemma frechet_derivative_unique_within_open_interval: fixes f::"'a::euclidean_space \ 'b::real_normed_vector" assumes x: "x \ box a b" and f: "(f has_derivative f' ) (at x within box a b)" "(f has_derivative f'') (at x within box a b)" shows "f' = f''" proof - have "at x within box a b = at x" by (metis x at_within_interior interior_open open_box) with f show "f' = f''" by (simp add: has_derivative_unique) qed lemma frechet_derivative_at: "(f has_derivative f') (at x) \ f' = frechet_derivative f (at x)" using differentiable_def frechet_derivative_works has_derivative_unique by blast lemma frechet_derivative_compose: "frechet_derivative (f o g) (at x) = frechet_derivative (f) (at (g x)) o frechet_derivative g (at x)" if "g differentiable at x" "f differentiable at (g x)" by (metis diff_chain_at frechet_derivative_at frechet_derivative_works that) lemma frechet_derivative_within_cbox: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "\i. i\Basis \ a\i < b\i" and "x \ cbox a b" and "(f has_derivative f') (at x within cbox a b)" shows "frechet_derivative f (at x within cbox a b) = f'" using assms by (metis Derivative.differentiableI frechet_derivative_unique_within_closed_interval frechet_derivative_works) lemma frechet_derivative_transform_within_open: "frechet_derivative f (at x) = frechet_derivative g (at x)" if "f differentiable at x" "open X" "x \ X" "\x. x \ X \ f x = g x" by (meson frechet_derivative_at frechet_derivative_works has_derivative_transform_within_open that) subsection \Derivatives of local minima and maxima are zero\ lemma has_derivative_local_min: fixes f :: "'a::real_normed_vector \ real" assumes deriv: "(f has_derivative f') (at x)" assumes min: "eventually (\y. f x \ f y) (at x)" shows "f' = (\h. 0)" proof fix h :: 'a interpret f': bounded_linear f' using deriv by (rule has_derivative_bounded_linear) show "f' h = 0" proof (cases "h = 0") case False from min obtain d where d1: "0 < d" and d2: "\y\ball x d. f x \ f y" unfolding eventually_at by (force simp: dist_commute) have "FDERIV (\r. x + r *\<^sub>R h) 0 :> (\r. r *\<^sub>R h)" by (intro derivative_eq_intros) auto then have "FDERIV (\r. f (x + r *\<^sub>R h)) 0 :> (\k. f' (k *\<^sub>R h))" by (rule has_derivative_compose, simp add: deriv) then have "DERIV (\r. f (x + r *\<^sub>R h)) 0 :> f' h" unfolding has_field_derivative_def by (simp add: f'.scaleR mult_commute_abs) moreover have "0 < d / norm h" using d1 and \h \ 0\ by simp moreover have "\y. \0 - y\ < d / norm h \ f (x + 0 *\<^sub>R h) \ f (x + y *\<^sub>R h)" using \h \ 0\ by (auto simp add: d2 dist_norm pos_less_divide_eq) ultimately show "f' h = 0" by (rule DERIV_local_min) qed simp qed lemma has_derivative_local_max: fixes f :: "'a::real_normed_vector \ real" assumes "(f has_derivative f') (at x)" assumes "eventually (\y. f y \ f x) (at x)" shows "f' = (\h. 0)" using has_derivative_local_min [of "\x. - f x" "\h. - f' h" "x"] using assms unfolding fun_eq_iff by simp lemma differential_zero_maxmin: fixes f::"'a::real_normed_vector \ real" assumes "x \ S" and "open S" and deriv: "(f has_derivative f') (at x)" and mono: "(\y\S. f y \ f x) \ (\y\S. f x \ f y)" shows "f' = (\v. 0)" using mono proof assume "\y\S. f y \ f x" with \x \ S\ and \open S\ have "eventually (\y. f y \ f x) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_max) next assume "\y\S. f x \ f y" with \x \ S\ and \open S\ have "eventually (\y. f x \ f y) (at x)" unfolding eventually_at_topological by auto with deriv show ?thesis by (rule has_derivative_local_min) qed lemma differential_zero_maxmin_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes k: "k \ Basis" and ball: "0 < e" "(\y \ ball x e. (f y)\k \ (f x)\k) \ (\y\ball x e. (f x)\k \ (f y)\k)" and diff: "f differentiable (at x)" shows "(\j\Basis. (frechet_derivative f (at x) j \ k) *\<^sub>R j) = (0::'a)" (is "?D k = 0") proof - let ?f' = "frechet_derivative f (at x)" have "x \ ball x e" using \0 < e\ by simp moreover have "open (ball x e)" by simp moreover have "((\x. f x \ k) has_derivative (\h. ?f' h \ k)) (at x)" using bounded_linear_inner_left diff[unfolded frechet_derivative_works] by (rule bounded_linear.has_derivative) ultimately have "(\h. frechet_derivative f (at x) h \ k) = (\v. 0)" using ball(2) by (rule differential_zero_maxmin) then show ?thesis unfolding fun_eq_iff by simp qed subsection \One-dimensional mean value theorem\ lemma mvt_simple: fixes f :: "real \ real" assumes "a < b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a<.. real" assumes "a \ b" and derf: "\x. \a \ x; x \ b\ \ (f has_derivative f' x) (at x within {a..b})" shows "\x\{a..b}. f b - f a = f' x (b - a)" proof (cases "a = b") interpret bounded_linear "f' b" using assms(2) assms(1) by auto case True then show ?thesis by force next case False then show ?thesis using mvt_simple[OF _ derf] by (metis \a \ b\ atLeastAtMost_iff dual_order.order_iff_strict greaterThanLessThan_iff) qed text \A nice generalization (see Havin's proof of 5.19 from Rudin's book).\ lemma mvt_general: fixes f :: "real \ 'a::real_inner" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\x\{a<.. norm (f' x (b - a))" proof - have "\x\{a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" apply (rule mvt [OF \a < b\, where f = "\x. (f b - f a) \ f x"]) apply (intro continuous_intros contf) using derf apply (auto intro: has_derivative_inner_right) done then obtain x where x: "x \ {a<.. f b - (f b - f a) \ f a = (f b - f a) \ f' x (b - a)" .. show ?thesis proof (cases "f a = f b") case False have "norm (f b - f a) * norm (f b - f a) = (norm (f b - f a))\<^sup>2" by (simp add: power2_eq_square) also have "\ = (f b - f a) \ (f b - f a)" unfolding power2_norm_eq_inner .. also have "\ = (f b - f a) \ f' x (b - a)" using x(2) by (simp only: inner_diff_right) also have "\ \ norm (f b - f a) * norm (f' x (b - a))" by (rule norm_cauchy_schwarz) finally show ?thesis using False x(1) by (auto simp add: mult_left_cancel) next case True then show ?thesis using \a < b\ by (rule_tac x="(a + b) /2" in bexI) auto qed qed subsection \More general bound theorems\ proposition differentiable_bound_general: fixes f :: "real \ 'a::real_normed_vector" assumes "a < b" and f_cont: "continuous_on {a..b} f" and phi_cont: "continuous_on {a..b} \" and f': "\x. a < x \ x < b \ (f has_vector_derivative f' x) (at x)" and phi': "\x. a < x \ x < b \ (\ has_vector_derivative \' x) (at x)" and bnd: "\x. a < x \ x < b \ norm (f' x) \ \' x" shows "norm (f b - f a) \ \ b - \ a" proof - { fix x assume x: "a < x" "x < b" have "0 \ norm (f' x)" by simp also have "\ \ \' x" using x by (auto intro!: bnd) finally have "0 \ \' x" . } note phi'_nonneg = this note f_tendsto = assms(2)[simplified continuous_on_def, rule_format] note phi_tendsto = assms(3)[simplified continuous_on_def, rule_format] { fix e::real assume "e > 0" define e2 where "e2 = e / 2" with \e > 0\ have "e2 > 0" by simp let ?le = "\x1. norm (f x1 - f a) \ \ x1 - \ a + e * (x1 - a) + e" define A where "A = {x2. a \ x2 \ x2 \ b \ (\x1\{a ..< x2}. ?le x1)}" have A_subset: "A \ {a..b}" by (auto simp: A_def) { fix x2 assume a: "a \ x2" "x2 \ b" and le: "\x1\{a..e > 0\ proof cases assume "x2 \ a" with a have "a < x2" by simp have "at x2 within {a <.. bot" using \a < x2\ by (auto simp: trivial_limit_within islimpt_in_closure) moreover have "((\x1. (\ x1 - \ a) + e * (x1 - a) + e) \ (\ x2 - \ a) + e * (x2 - a) + e) (at x2 within {a <..x1. norm (f x1 - f a)) \ norm (f x2 - f a)) (at x2 within {a <..x. x > a) (at x2 within {a <.. A" using assms by (auto simp: A_def) hence [simp]: "A \ {}" by auto have A_ivl: "\x1 x2. x2 \ A \ x1 \ {a ..x2} \ x1 \ A" by (simp add: A_def) have [simp]: "bdd_above A" by (auto simp: A_def) define y where "y = Sup A" have "y \ b" unfolding y_def by (simp add: cSup_le_iff) (simp add: A_def) have leI: "\x x1. a \ x1 \ x \ A \ x1 < x \ ?le x1" by (auto simp: A_def intro!: le_cont) have y_all_le: "\x1\{a.. y" by (metis \a \ A\ \bdd_above A\ cSup_upper y_def) have "y \ A" using y_all_le \a \ y\ \y \ b\ by (auto simp: A_def) hence "A = {a .. y}" using A_subset by (auto simp: subset_iff y_def cSup_upper intro: A_ivl) from le_cont[OF \a \ y\ \y \ b\ y_all_le] have le_y: "?le y" . have "y = b" proof (cases "a = y") case True with \a < b\ have "y < b" by simp with \a = y\ f_cont phi_cont \e2 > 0\ have 1: "\\<^sub>F x in at y within {y..b}. dist (f x) (f y) < e2" and 2: "\\<^sub>F x in at y within {y..b}. dist (\ x) (\ y) < e2" by (auto simp: continuous_on_def tendsto_iff) have 3: "eventually (\x. y < x) (at y within {y..b})" by (auto simp: eventually_at_filter) have 4: "eventually (\x::real. x < b) (at y within {y..b})" using _ \y < b\ by (rule order_tendstoD) (auto intro!: tendsto_eq_intros) from 1 2 3 4 have eventually_le: "eventually (\x. ?le x) (at y within {y .. b})" proof eventually_elim case (elim x1) have "norm (f x1 - f a) = norm (f x1 - f y)" by (simp add: \a = y\) also have "norm (f x1 - f y) \ e2" using elim \a = y\ by (auto simp : dist_norm intro!: less_imp_le) also have "\ \ e2 + (\ x1 - \ a + e2 + e * (x1 - a))" using \0 < e\ elim by (intro add_increasing2[OF add_nonneg_nonneg order.refl]) (auto simp: \a = y\ dist_norm intro!: mult_nonneg_nonneg) also have "\ = \ x1 - \ a + e * (x1 - a) + e" by (simp add: e2_def) finally show "?le x1" . qed from this[unfolded eventually_at_topological] \?le y\ obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y..b} \ ?le x" by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min b (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a = y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" by (simp add: d'_def) fix x1 assume "x1 \ {a.. S" "x1 \ {y..b}" by (auto simp: \a = y\ d'_def dist_real_def intro!: d ) thus "?le x1" by (rule S) qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp then show "y = b" using \d > 0\ \y < b\ by (simp add: d'_def) next case False with \a \ y\ have "a < y" by simp show "y = b" proof (rule ccontr) assume "y \ b" hence "y < b" using \y \ b\ by simp let ?F = "at y within {y.. has_vector_derivative \' y) ?F" using \a < y\ \y < b\ by (auto simp add: at_within_open[of _ "{a<..\<^sub>F x1 in ?F. norm (f x1 - f y - (x1 - y) *\<^sub>R f' y) \ e2 * \x1 - y\" "\\<^sub>F x1 in ?F. norm (\ x1 - \ y - (x1 - y) *\<^sub>R \' y) \ e2 * \x1 - y\" using \e2 > 0\ by (auto simp: has_derivative_within_alt2 has_vector_derivative_def) moreover have "\\<^sub>F x1 in ?F. y \ x1" "\\<^sub>F x1 in ?F. x1 < b" by (auto simp: eventually_at_filter) ultimately have "\\<^sub>F x1 in ?F. norm (f x1 - f y) \ (\ x1 - \ y) + e * \x1 - y\" (is "\\<^sub>F x1 in ?F. ?le' x1") proof eventually_elim case (elim x1) from norm_triangle_ineq2[THEN order_trans, OF elim(1)] have "norm (f x1 - f y) \ norm (f' y) * \x1 - y\ + e2 * \x1 - y\" by (simp add: ac_simps) also have "norm (f' y) \ \' y" using bnd \a < y\ \y < b\ by simp also have "\' y * \x1 - y\ \ \ x1 - \ y + e2 * \x1 - y\" using elim by (simp add: ac_simps) finally have "norm (f x1 - f y) \ \ x1 - \ y + e2 * \x1 - y\ + e2 * \x1 - y\" by (auto simp: mult_right_mono) thus ?case by (simp add: e2_def) qed moreover have "?le' y" by simp ultimately obtain S where S: "open S" "y \ S" "\x. x\S \ x \ {y.. ?le' x" unfolding eventually_at_topological by metis from \open S\ obtain d where d: "\x. dist x y < d \ x \ S" "d > 0" by (force simp: dist_commute open_dist ball_def dest!: bspec[OF _ \y \ S\]) define d' where "d' = min ((y + b)/2) (y + (d/2))" have "d' \ A" unfolding A_def proof safe show "a \ d'" using \a < y\ \0 < d\ \y < b\ by (simp add: d'_def) show "d' \ b" using \y < b\ by (simp add: d'_def min_def) fix x1 assume x1: "x1 \ {a..y \ A\ local.leI x1 by auto next case False hence x1': "x1 \ S" "x1 \ {y.. norm (f x1 - f y) + norm (f y - f a)" by (rule order_trans[OF _ norm_triangle_ineq]) simp also note S(3)[OF x1'] also note le_y finally show "?le x1" using False by (auto simp: algebra_simps) qed qed hence "d' \ y" unfolding y_def by (rule cSup_upper) simp thus False using \d > 0\ \y < b\ by (simp add: d'_def min_def split: if_split_asm) qed qed with le_y have "norm (f b - f a) \ \ b - \ a + e * (b - a + 1)" by (simp add: algebra_simps) } note * = this show ?thesis proof (rule field_le_epsilon) fix e::real assume "e > 0" then show "norm (f b - f a) \ \ b - \ a + e" using *[of "e / (b - a + 1)"] \a < b\ by simp qed qed lemma differentiable_bound: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\x. x\S \ (f has_derivative f' x) (at x within S)" and B: "\x. x \ S \ onorm (f' x) \ B" and x: "x \ S" and y: "y \ S" shows "norm (f x - f y) \ B * norm (x - y)" proof - let ?p = "\u. x + u *\<^sub>R (y - x)" let ?\ = "\h. h * B * norm (x - y)" have *: "x + u *\<^sub>R (y - x) \ S" if "u \ {0..1}" for u proof - have "u *\<^sub>R y = u *\<^sub>R (y - x) + u *\<^sub>R x" by (simp add: scale_right_diff_distrib) then show "x + u *\<^sub>R (y - x) \ S" using that \convex S\ x y by (simp add: convex_alt) (metis pth_b(2) pth_c(1) scaleR_collapse) qed have "\z. z \ (\u. x + u *\<^sub>R (y - x)) ` {0..1} \ (f has_derivative f' z) (at z within (\u. x + u *\<^sub>R (y - x)) ` {0..1})" by (auto intro: * has_derivative_within_subset [OF derf]) then have "continuous_on (?p ` {0..1}) f" unfolding continuous_on_eq_continuous_within by (meson has_derivative_continuous) with * have 1: "continuous_on {0 .. 1} (f \ ?p)" by (intro continuous_intros)+ { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret linear "(f' ?u)" using u by (auto intro!: has_derivative_linear derf *) have "(f \ ?p has_derivative (f' ?u) \ (\u. 0 + u *\<^sub>R (y - x))) (at u within box 0 1)" by (intro derivative_intros has_derivative_within_subset [OF derf]) (use u * in auto) hence "((f \ ?p) has_vector_derivative f' ?u (y - x)) (at u)" by (simp add: at_within_open[OF u open_greaterThanLessThan] scaleR has_vector_derivative_def o_def) } note 2 = this have 3: "continuous_on {0..1} ?\" by (rule continuous_intros)+ have 4: "(?\ has_vector_derivative B * norm (x - y)) (at u)" for u by (auto simp: has_vector_derivative_def intro!: derivative_eq_intros) { fix u::real assume u: "u \{0 <..< 1}" let ?u = "?p u" interpret bounded_linear "(f' ?u)" using u by (auto intro!: has_derivative_bounded_linear derf *) have "norm (f' ?u (y - x)) \ onorm (f' ?u) * norm (y - x)" by (rule onorm) (rule bounded_linear) also have "onorm (f' ?u) \ B" using u by (auto intro!: assms(3)[rule_format] *) finally have "norm ((f' ?u) (y - x)) \ B * norm (x - y)" by (simp add: mult_right_mono norm_minus_commute) } note 5 = this have "norm (f x - f y) = norm ((f \ (\u. x + u *\<^sub>R (y - x))) 1 - (f \ (\u. x + u *\<^sub>R (y - x))) 0)" by (auto simp add: norm_minus_commute) also from differentiable_bound_general[OF zero_less_one 1, OF 3 2 4 5] have "norm ((f \ ?p) 1 - (f \ ?p) 0) \ B * norm (x - y)" by simp finally show ?thesis . qed lemma field_differentiable_bound: fixes S :: "'a::real_normed_field set" assumes cvs: "convex S" and df: "\z. z \ S \ (f has_field_derivative f' z) (at z within S)" and dn: "\z. z \ S \ norm (f' z) \ B" and "x \ S" "y \ S" shows "norm(f x - f y) \ B * norm(x - y)" apply (rule differentiable_bound [OF cvs]) apply (erule df [unfolded has_field_derivative_def]) apply (rule onorm_le, simp_all add: norm_mult mult_right_mono assms) done lemma differentiable_bound_segment: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes "\t. t \ {0..1} \ x0 + t *\<^sub>R a \ G" assumes f': "\x. x \ G \ (f has_derivative f' x) (at x within G)" assumes B: "\x. x \ {0..1} \ onorm (f' (x0 + x *\<^sub>R a)) \ B" shows "norm (f (x0 + a) - f x0) \ norm a * B" proof - let ?G = "(\x. x0 + x *\<^sub>R a) ` {0..1}" have "?G = (+) x0 ` (\x. x *\<^sub>R a) ` {0..1}" by auto also have "convex \" by (intro convex_translation convex_scaled convex_real_interval) finally have "convex ?G" . moreover have "?G \ G" "x0 \ ?G" "x0 + a \ ?G" using assms by (auto intro: image_eqI[where x=1]) ultimately show ?thesis using has_derivative_subset[OF f' \?G \ G\] B differentiable_bound[of "(\x. x0 + x *\<^sub>R a) ` {0..1}" f f' B "x0 + a" x0] by (force simp: ac_simps) qed lemma differentiable_bound_linearization: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector" assumes S: "\t. t \ {0..1} \ a + t *\<^sub>R (b - a) \ S" assumes f'[derivative_intros]: "\x. x \ S \ (f has_derivative f' x) (at x within S)" assumes B: "\x. x \ S \ onorm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - f' x0 (b - a)) \ norm (b - a) * B" proof - define g where [abs_def]: "g x = f x - f' x0 x" for x have g: "\x. x \ S \ (g has_derivative (\i. f' x i - f' x0 i)) (at x within S)" unfolding g_def using assms by (auto intro!: derivative_eq_intros bounded_linear.has_derivative[OF has_derivative_bounded_linear, OF f']) from B have "\x\{0..1}. onorm (\i. f' (a + x *\<^sub>R (b - a)) i - f' x0 i) \ B" using assms by (auto simp: fun_diff_def) with differentiable_bound_segment[OF S g] \x0 \ S\ show ?thesis by (simp add: g_def field_simps linear_diff[OF has_derivative_linear[OF f']]) qed lemma vector_differentiable_bound_linearization: fixes f::"real \ 'b::real_normed_vector" assumes f': "\x. x \ S \ (f has_vector_derivative f' x) (at x within S)" assumes "closed_segment a b \ S" assumes B: "\x. x \ S \ norm (f' x - f' x0) \ B" assumes "x0 \ S" shows "norm (f b - f a - (b - a) *\<^sub>R f' x0) \ norm (b - a) * B" using assms by (intro differentiable_bound_linearization[of a b S f "\x h. h *\<^sub>R f' x" x0 B]) (force simp: closed_segment_real_eq has_vector_derivative_def scaleR_diff_right[symmetric] mult.commute[of B] intro!: onorm_le mult_left_mono)+ text \In particular.\ lemma has_derivative_zero_constant: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" shows "\c. \x\s. f x = c" proof - { fix x y assume "x \ s" "y \ s" then have "norm (f x - f y) \ 0 * norm (x - y)" using assms by (intro differentiable_bound[of s]) (auto simp: onorm_zero) then have "f x = f y" by simp } then show ?thesis by metis qed lemma has_field_derivative_zero_constant: assumes "convex s" "\x. x \ s \ (f has_field_derivative 0) (at x within s)" shows "\c. \x\s. f (x) = (c :: 'a :: real_normed_field)" proof (rule has_derivative_zero_constant) have A: "(*) 0 = (\_. 0 :: 'a)" by (intro ext) simp fix x assume "x \ s" thus "(f has_derivative (\h. 0)) (at x within s)" using assms(2)[of x] by (simp add: has_field_derivative_def A) qed fact lemma has_vector_derivative_zero_constant: assumes "convex s" assumes "\x. x \ s \ (f has_vector_derivative 0) (at x within s)" obtains c where "\x. x \ s \ f x = c" using has_derivative_zero_constant[of s f] assms by (auto simp: has_vector_derivative_def) lemma has_derivative_zero_unique: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex s" and "\x. x \ s \ (f has_derivative (\h. 0)) (at x within s)" and "x \ s" "y \ s" shows "f x = f y" using has_derivative_zero_constant[OF assms(1,2)] assms(3-) by force lemma has_derivative_zero_unique_connected: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open s" "connected s" assumes f: "\x. x \ s \ (f has_derivative (\x. 0)) (at x)" assumes "x \ s" "y \ s" shows "f x = f y" proof (rule connected_local_const[where f=f, OF \connected s\ \x\s\ \y\s\]) show "\a\s. eventually (\b. f a = f b) (at a within s)" proof fix a assume "a \ s" with \open s\ obtain e where "0 < e" "ball a e \ s" by (rule openE) then have "\c. \x\ball a e. f x = c" by (intro has_derivative_zero_constant) (auto simp: at_within_open[OF _ open_ball] f convex_ball) with \0 have "\x\ball a e. f a = f x" by auto then show "eventually (\b. f a = f b) (at a within s)" using \0 unfolding eventually_at_topological by (intro exI[of _ "ball a e"]) auto qed qed subsection \Differentiability of inverse function (most basic form)\ lemma has_derivative_inverse_basic: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes derf: "(f has_derivative f') (at (g y))" and ling': "bounded_linear g'" and "g' \ f' = id" and contg: "continuous (at y) g" and "open T" and "y \ T" and fg: "\z. z \ T \ f (g z) = z" shows "(g has_derivative g') (at y)" proof - interpret f': bounded_linear f' using assms unfolding has_derivative_def by auto interpret g': bounded_linear g' using assms by auto obtain C where C: "0 < C" "\x. norm (g' x) \ norm x * C" using bounded_linear.pos_bounded[OF assms(2)] by blast have lem1: "\e>0. \d>0. \z. norm (z - y) < d \ norm (g z - g y - g'(z - y)) \ e * norm (g z - g y)" proof (intro allI impI) fix e :: real assume "e > 0" with C(1) have *: "e / C > 0" by auto obtain d0 where "0 < d0" and d0: "\u. norm (u - g y) < d0 \ norm (f u - f (g y) - f' (u - g y)) \ e / C * norm (u - g y)" using derf * unfolding has_derivative_at_alt by blast obtain d1 where "0 < d1" and d1: "\x. \0 < dist x y; dist x y < d1\ \ dist (g x) (g y) < d0" using contg \0 < d0\ unfolding continuous_at Lim_at by blast obtain d2 where "0 < d2" and d2: "\u. dist u y < d2 \ u \ T" using \open T\ \y \ T\ unfolding open_dist by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF \0 < d1\ \0 < d2\] by blast show "\d>0. \z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < d" then have "z \ T" using d2 d unfolding dist_norm by auto have "norm (g z - g y - g' (z - y)) \ norm (g' (f (g z) - y - f' (g z - g y)))" unfolding g'.diff f'.diff unfolding assms(3)[unfolded o_def id_def, THEN fun_cong] fg[OF \z\T\] by (simp add: norm_minus_commute) also have "\ \ norm (f (g z) - y - f' (g z - g y)) * C" by (rule C(2)) also have "\ \ (e / C) * norm (g z - g y) * C" proof - have "norm (g z - g y) < d0" by (metis as cancel_comm_monoid_add_class.diff_cancel d(2) \0 < d0\ d1 diff_gt_0_iff_gt diff_strict_mono dist_norm dist_self zero_less_dist_iff) then show ?thesis by (metis C(1) \y \ T\ d0 fg real_mult_le_cancel_iff1) qed also have "\ \ e * norm (g z - g y)" using C by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (g z - g y)" by simp qed (use d in auto) qed have *: "(0::real) < 1 / 2" by auto obtain d where "0 < d" and d: "\z. norm (z - y) < d \ norm (g z - g y - g' (z - y)) \ 1/2 * norm (g z - g y)" using lem1 * by blast define B where "B = C * 2" have "B > 0" unfolding B_def using C by auto have lem2: "norm (g z - g y) \ B * norm (z - y)" if z: "norm(z - y) < d" for z proof - have "norm (g z - g y) \ norm(g' (z - y)) + norm ((g z - g y) - g'(z - y))" by (rule norm_triangle_sub) also have "\ \ norm (g' (z - y)) + 1 / 2 * norm (g z - g y)" by (rule add_left_mono) (use d z in auto) also have "\ \ norm (z - y) * C + 1 / 2 * norm (g z - g y)" by (rule add_right_mono) (use C in auto) finally show "norm (g z - g y) \ B * norm (z - y)" unfolding B_def by (auto simp add: field_simps) qed show ?thesis unfolding has_derivative_at_alt proof (intro conjI assms allI impI) fix e :: real assume "e > 0" then have *: "e / B > 0" by (metis \B > 0\ divide_pos_pos) obtain d' where "0 < d'" and d': "\z. norm (z - y) < d' \ norm (g z - g y - g' (z - y)) \ e / B * norm (g z - g y)" using lem1 * by blast obtain k where k: "0 < k" "k < d" "k < d'" using field_lbound_gt_zero[OF \0 < d\ \0 < d'\] by blast show "\d>0. \ya. norm (ya - y) < d \ norm (g ya - g y - g' (ya - y)) \ e * norm (ya - y)" proof (intro exI allI impI conjI) fix z assume as: "norm (z - y) < k" then have "norm (g z - g y - g' (z - y)) \ e / B * norm(g z - g y)" using d' k by auto also have "\ \ e * norm (z - y)" unfolding times_divide_eq_left pos_divide_le_eq[OF \B>0\] using lem2[of z] k as \e > 0\ by (auto simp add: field_simps) finally show "norm (g z - g y - g' (z - y)) \ e * norm (z - y)" by simp qed (use k in auto) qed qed text\<^marker>\tag unimportant\\Inverse function theorem for complex derivatives\ lemma has_field_derivative_inverse_basic: shows "DERIV f (g y) :> f' \ f' \ 0 \ continuous (at y) g \ open t \ y \ t \ (\z. z \ t \ f (g z) = z) \ DERIV g y :> inverse (f')" unfolding has_field_derivative_def apply (rule has_derivative_inverse_basic) apply (auto simp: bounded_linear_mult_right) done text \Simply rewrite that based on the domain point x.\ lemma has_derivative_inverse_basic_x: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" and "continuous (at (f x)) g" and "g (f x) = x" and "open T" and "f x \ T" and "\y. y \ T \ f (g y) = y" shows "(g has_derivative g') (at (f x))" by (rule has_derivative_inverse_basic) (use assms in auto) text \This is the version in Dieudonne', assuming continuity of f and g.\ lemma has_derivative_inverse_dieudonne: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "open S" and "open (f ` S)" and "continuous_on S f" and "continuous_on (f ` S) g" and "\x. x \ S \ g (f x) = x" and "x \ S" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" apply (rule has_derivative_inverse_basic_x[OF assms(7-9) _ _ assms(2)]) using assms(3-6) unfolding continuous_on_eq_continuous_at[OF assms(1)] continuous_on_eq_continuous_at[OF assms(2)] apply auto done text \Here's the simplest way of not assuming much about g.\ proposition has_derivative_inverse: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "compact S" and "x \ S" and fx: "f x \ interior (f ` S)" and "continuous_on S f" and gf: "\y. y \ S \ g (f y) = y" and "(f has_derivative f') (at x)" and "bounded_linear g'" and "g' \ f' = id" shows "(g has_derivative g') (at (f x))" proof - have *: "\y. y \ interior (f ` S) \ f (g y) = y" by (metis gf image_iff interior_subset subsetCE) show ?thesis apply (rule has_derivative_inverse_basic_x[OF assms(6-8), where T = "interior (f ` S)"]) apply (rule continuous_on_interior[OF _ fx]) apply (rule continuous_on_inv) apply (simp_all add: assms *) done qed text \Invertible derivative continuous at a point implies local injectivity. It's only for this we need continuity of the derivative, except of course if we want the fact that the inverse derivative is also continuous. So if we know for some other reason that the inverse function exists, it's OK.\ proposition has_derivative_locally_injective: fixes f :: "'n::euclidean_space \ 'm::euclidean_space" assumes "a \ S" and "open S" and bling: "bounded_linear g'" and "g' \ f' a = id" and derf: "\x. x \ S \ (f has_derivative f' x) (at x)" and "\e. e > 0 \ \d>0. \x. dist a x < d \ onorm (\v. f' x v - f' a v) < e" obtains r where "r > 0" "ball a r \ S" "inj_on f (ball a r)" proof - interpret bounded_linear g' using assms by auto note f'g' = assms(4)[unfolded id_def o_def,THEN cong] have "g' (f' a (\Basis)) = (\Basis)" "(\Basis) \ (0::'n)" using f'g' by auto then have *: "0 < onorm g'" unfolding onorm_pos_lt[OF assms(3)] by fastforce define k where "k = 1 / onorm g' / 2" have *: "k > 0" unfolding k_def using * by auto obtain d1 where d1: "0 < d1" "\x. dist a x < d1 \ onorm (\v. f' x v - f' a v) < k" using assms(6) * by blast from \open S\ obtain d2 where "d2 > 0" "ball a d2 \ S" using \a\S\ .. obtain d2 where d2: "0 < d2" "ball a d2 \ S" using \0 < d2\ \ball a d2 \ S\ by blast obtain d where d: "0 < d" "d < d1" "d < d2" using field_lbound_gt_zero[OF d1(1) d2(1)] by blast show ?thesis proof show "0 < d" by (fact d) show "ball a d \ S" using \d < d2\ \ball a d2 \ S\ by auto show "inj_on f (ball a d)" unfolding inj_on_def proof (intro strip) fix x y assume as: "x \ ball a d" "y \ ball a d" "f x = f y" define ph where [abs_def]: "ph w = w - g' (f w - f x)" for w have ph':"ph = g' \ (\w. f' a w - (f w - f x))" unfolding ph_def o_def by (simp add: diff f'g') have "norm (ph x - ph y) \ (1 / 2) * norm (x - y)" proof (rule differentiable_bound[OF convex_ball _ _ as(1-2)]) fix u assume u: "u \ ball a d" then have "u \ S" using d d2 by auto have *: "(\v. v - g' (f' u v)) = g' \ (\w. f' a w - f' u w)" unfolding o_def and diff using f'g' by auto have blin: "bounded_linear (f' a)" using \a \ S\ derf by blast show "(ph has_derivative (\v. v - g' (f' u v))) (at u within ball a d)" unfolding ph' * comp_def by (rule \u \ S\ derivative_eq_intros has_derivative_at_withinI [OF derf] bounded_linear.has_derivative [OF blin] bounded_linear.has_derivative [OF bling] |simp)+ have **: "bounded_linear (\x. f' u x - f' a x)" "bounded_linear (\x. f' a x - f' u x)" using \u \ S\ blin bounded_linear_sub derf by auto then have "onorm (\v. v - g' (f' u v)) \ onorm g' * onorm (\w. f' a w - f' u w)" by (simp add: "*" bounded_linear_axioms onorm_compose) also have "\ \ onorm g' * k" apply (rule mult_left_mono) using d1(2)[of u] using onorm_neg[where f="\x. f' u x - f' a x"] d u onorm_pos_le[OF bling] apply (auto simp: algebra_simps) done also have "\ \ 1 / 2" unfolding k_def by auto finally show "onorm (\v. v - g' (f' u v)) \ 1 / 2" . qed moreover have "norm (ph y - ph x) = norm (y - x)" by (simp add: as(3) ph_def) ultimately show "x = y" unfolding norm_minus_commute by auto qed qed qed subsection \Uniformly convergent sequence of derivatives\ lemma has_derivative_sequence_lipschitz_lemma: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\n x h. \n\N; x \ S\ \ norm (f' n x h - g' x h) \ e * norm h" and "0 \ e" shows "\m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof clarify fix m n x y assume as: "N \ m" "N \ n" "x \ S" "y \ S" show "norm ((f m x - f n x) - (f m y - f n y)) \ 2 * e * norm (x - y)" proof (rule differentiable_bound[where f'="\x h. f' m x h - f' n x h", OF \convex S\ _ _ as(3-4)]) fix x assume "x \ S" show "((\a. f m a - f n a) has_derivative (\h. f' m x h - f' n x h)) (at x within S)" by (rule derivative_intros derf \x\S\)+ show "onorm (\h. f' m x h - f' n x h) \ 2 * e" proof (rule onorm_bound) fix h have "norm (f' m x h - f' n x h) \ norm (f' m x h - g' x h) + norm (f' n x h - g' x h)" using norm_triangle_ineq[of "f' m x h - g' x h" "- f' n x h + g' x h"] by (auto simp add: algebra_simps norm_minus_commute) also have "\ \ e * norm h + e * norm h" using nle[OF \N \ m\ \x \ S\, of h] nle[OF \N \ n\ \x \ S\, of h] by (auto simp add: field_simps) finally show "norm (f' m x h - f' n x h) \ 2 * e * norm h" by auto qed (simp add: \0 \ e\) qed qed lemma has_derivative_sequence_Lipschitz: fixes f :: "nat \ 'a::real_normed_vector \ 'b::real_normed_vector" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "e > 0" shows "\N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" proof - have *: "2 * (e/2) = e" using \e > 0\ by auto obtain N where "\n\N. \x\S. \h. norm (f' n x h - g' x h) \ (e/2) * norm h" using nle \e > 0\ unfolding eventually_sequentially by (metis less_divide_eq_numeral1(1) mult_zero_left) then show "\N. \m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" apply (rule_tac x=N in exI) apply (rule has_derivative_sequence_lipschitz_lemma[where e="e/2", unfolded *]) using assms \e > 0\ apply auto done qed proposition has_derivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and derf: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and nle: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" and "x0 \ S" and lim: "((\n. f n x0) \ l) sequentially" shows "\g. \x\S. (\n. f n x) \ g x \ (g has_derivative g'(x)) (at x within S)" proof - have lem1: "\e. e > 0 \ \N. \m\N. \n\N. \x\S. \y\S. norm ((f m x - f n x) - (f m y - f n y)) \ e * norm (x - y)" using assms(1,2,3) by (rule has_derivative_sequence_Lipschitz) have "\g. \x\S. ((\n. f n x) \ g x) sequentially" proof (intro ballI bchoice) fix x assume "x \ S" show "\y. (\n. f n x) \ y" unfolding convergent_eq_Cauchy proof (cases "x = x0") case True then show "Cauchy (\n. f n x)" using LIMSEQ_imp_Cauchy[OF lim] by auto next case False show "Cauchy (\n. f n x)" unfolding Cauchy_def proof (intro allI impI) fix e :: real assume "e > 0" hence *: "e / 2 > 0" "e / 2 / norm (x - x0) > 0" using False by auto obtain M where M: "\m\M. \n\M. dist (f m x0) (f n x0) < e / 2" using LIMSEQ_imp_Cauchy[OF lim] * unfolding Cauchy_def by blast obtain N where N: "\m\N. \n\N. \u\S. \y\S. norm (f m u - f n u - (f m y - f n y)) \ e / 2 / norm (x - x0) * norm (u - y)" using lem1 *(2) by blast show "\M. \m\M. \n\M. dist (f m x) (f n x) < e" proof (intro exI allI impI) fix m n assume as: "max M N \m" "max M N\n" have "dist (f m x) (f n x) \ norm (f m x0 - f n x0) + norm (f m x - f n x - (f m x0 - f n x0))" unfolding dist_norm by (rule norm_triangle_sub) also have "\ \ norm (f m x0 - f n x0) + e / 2" using N \x\S\ \x0\S\ as False by fastforce also have "\ < e / 2 + e / 2" by (rule add_strict_right_mono) (use as M in \auto simp: dist_norm\) finally show "dist (f m x) (f n x) < e" by auto qed qed qed qed then obtain g where g: "\x\S. (\n. f n x) \ g x" .. have lem2: "\N. \n\N. \x\S. \y\S. norm ((f n x - f n y) - (g x - g y)) \ e * norm (x - y)" if "e > 0" for e proof - obtain N where N: "\m\N. \n\N. \x\S. \y\S. norm (f m x - f n x - (f m y - f n y)) \ e * norm (x - y)" using lem1 \e > 0\ by blast show "\N. \n\N. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" proof (intro exI ballI allI impI) fix n x y assume as: "N \ n" "x \ S" "y \ S" have "((\m. norm (f n x - f n y - (f m x - f m y))) \ norm (f n x - f n y - (g x - g y))) sequentially" by (intro tendsto_intros g[rule_format] as) moreover have "eventually (\m. norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)) sequentially" unfolding eventually_sequentially proof (intro exI allI impI) fix m assume "N \ m" then show "norm (f n x - f n y - (f m x - f m y)) \ e * norm (x - y)" using N as by (auto simp add: algebra_simps) qed ultimately show "norm (f n x - f n y - (g x - g y)) \ e * norm (x - y)" by (simp add: tendsto_upperbound) qed qed have "\x\S. ((\n. f n x) \ g x) sequentially \ (g has_derivative g' x) (at x within S)" unfolding has_derivative_within_alt2 proof (intro ballI conjI allI impI) fix x assume "x \ S" then show "(\n. f n x) \ g x" by (simp add: g) have tog': "(\n. f' n x u) \ g' x u" for u unfolding filterlim_def le_nhds_metric_le eventually_filtermap dist_norm proof (intro allI impI) fix e :: real assume "e > 0" show "eventually (\n. norm (f' n x u - g' x u) \ e) sequentially" proof (cases "u = 0") case True have "eventually (\n. norm (f' n x u - g' x u) \ e * norm u) sequentially" using nle \0 < e\ \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u = 0\ \0 < e\ by (auto elim: eventually_mono) next case False with \0 < e\ have "0 < e / norm u" by simp then have "eventually (\n. norm (f' n x u - g' x u) \ e / norm u * norm u) sequentially" using nle \x \ S\ by (fast elim: eventually_mono) then show ?thesis using \u \ 0\ by simp qed qed show "bounded_linear (g' x)" proof fix x' y z :: 'a fix c :: real note lin = assms(2)[rule_format,OF \x\S\,THEN has_derivative_bounded_linear] show "g' x (c *\<^sub>R x') = c *\<^sub>R g' x x'" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_cmul] apply (intro tendsto_intros tog') done show "g' x (y + z) = g' x y + g' x z" apply (rule tendsto_unique[OF trivial_limit_sequentially tog']) unfolding lin[THEN bounded_linear.linear, THEN linear_add] apply (rule tendsto_add) apply (rule tog')+ done obtain N where N: "\h. norm (f' N x h - g' x h) \ 1 * norm h" using nle \x \ S\ unfolding eventually_sequentially by (fast intro: zero_less_one) have "bounded_linear (f' N x)" using derf \x \ S\ by fast from bounded_linear.bounded [OF this] obtain K where K: "\h. norm (f' N x h) \ norm h * K" .. { fix h have "norm (g' x h) = norm (f' N x h - (f' N x h - g' x h))" by simp also have "\ \ norm (f' N x h) + norm (f' N x h - g' x h)" by (rule norm_triangle_ineq4) also have "\ \ norm h * K + 1 * norm h" using N K by (fast intro: add_mono) finally have "norm (g' x h) \ norm h * (K + 1)" by (simp add: ring_distribs) } then show "\K. \h. norm (g' x h) \ norm h * K" by fast qed show "eventually (\y. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)) (at x within S)" if "e > 0" for e proof - have *: "e / 3 > 0" using that by auto obtain N1 where N1: "\n\N1. \x\S. \h. norm (f' n x h - g' x h) \ e / 3 * norm h" using nle * unfolding eventually_sequentially by blast obtain N2 where N2[rule_format]: "\n\N2. \x\S. \y\S. norm (f n x - f n y - (g x - g y)) \ e / 3 * norm (x - y)" using lem2 * by blast let ?N = "max N1 N2" have "eventually (\y. norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)) (at x within S)" using derf[unfolded has_derivative_within_alt2] and \x \ S\ and * by fast moreover have "eventually (\y. y \ S) (at x within S)" unfolding eventually_at by (fast intro: zero_less_one) ultimately show "\\<^sub>F y in at x within S. norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" proof (rule eventually_elim2) fix y assume "y \ S" assume "norm (f ?N y - f ?N x - f' ?N x (y - x)) \ e / 3 * norm (y - x)" moreover have "norm (g y - g x - (f ?N y - f ?N x)) \ e / 3 * norm (y - x)" using N2[OF _ \y \ S\ \x \ S\] by (simp add: norm_minus_commute) ultimately have "norm (g y - g x - f' ?N x (y - x)) \ 2 * e / 3 * norm (y - x)" using norm_triangle_le[of "g y - g x - (f ?N y - f ?N x)" "f ?N y - f ?N x - f' ?N x (y - x)" "2 * e / 3 * norm (y - x)"] by (auto simp add: algebra_simps) moreover have " norm (f' ?N x (y - x) - g' x (y - x)) \ e / 3 * norm (y - x)" using N1 \x \ S\ by auto ultimately show "norm (g y - g x - g' x (y - x)) \ e * norm (y - x)" using norm_triangle_le[of "g y - g x - f' (max N1 N2) x (y - x)" "f' (max N1 N2) x (y - x) - g' x (y - x)"] by (auto simp add: algebra_simps) qed qed qed then show ?thesis by fast qed text \Can choose to line up antiderivatives if we want.\ lemma has_antiderivative_sequence: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and der: "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and no: "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof (cases "S = {}") case False then obtain a where "a \ S" by auto have *: "\P Q. \g. \x\S. P g x \ Q g x \ \g. \x\S. Q g x" by auto show ?thesis apply (rule *) apply (rule has_derivative_sequence [OF \convex S\ _ no, of "\n x. f n x + (f 0 a - f n a)"]) apply (metis assms(2) has_derivative_add_const) using \a \ S\ apply auto done qed auto lemma has_antiderivative_limit: fixes g' :: "'a::real_normed_vector \ 'a \ 'b::banach" assumes "convex S" and "\e. e>0 \ \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm (f' x h - g' x h) \ e * norm h)" shows "\g. \x\S. (g has_derivative g' x) (at x within S)" proof - have *: "\n. \f f'. \x\S. (f has_derivative (f' x)) (at x within S) \ (\h. norm(f' x h - g' x h) \ inverse (real (Suc n)) * norm h)" by (simp add: assms(2)) obtain f where *: "\x. \f'. \xa\S. (f x has_derivative f' xa) (at xa within S) \ (\h. norm (f' xa h - g' xa h) \ inverse (real (Suc x)) * norm h)" using * by metis obtain f' where f': "\x. \z\S. (f x has_derivative f' x z) (at z within S) \ (\h. norm (f' x z h - g' z h) \ inverse (real (Suc x)) * norm h)" using * by metis show ?thesis proof (rule has_antiderivative_sequence[OF \convex S\, of f f']) fix e :: real assume "e > 0" obtain N where N: "inverse (real (Suc N)) < e" using reals_Archimedean[OF \e>0\] .. show "\\<^sub>F n in sequentially. \x\S. \h. norm (f' n x h - g' x h) \ e * norm h" unfolding eventually_sequentially proof (intro exI allI ballI impI) fix n x h assume n: "N \ n" and x: "x \ S" have *: "inverse (real (Suc n)) \ e" apply (rule order_trans[OF _ N[THEN less_imp_le]]) using n apply (auto simp add: field_simps) done show "norm (f' n x h - g' x h) \ e * norm h" by (meson "*" mult_right_mono norm_ge_zero order.trans x f') qed qed (use f' in auto) qed subsection \Differentiation of a series\ proposition has_derivative_series: fixes f :: "nat \ 'a::real_normed_vector \ 'b::banach" assumes "convex S" and "\n x. x \ S \ ((f n) has_derivative (f' n x)) (at x within S)" and "\e. e>0 \ \\<^sub>F n in sequentially. \x\S. \h. norm (sum (\i. f' i x h) {.. e * norm h" and "x \ S" and "(\n. f n x) sums l" shows "\g. \x\S. (\n. f n x) sums (g x) \ (g has_derivative g' x) (at x within S)" unfolding sums_def apply (rule has_derivative_sequence[OF assms(1) _ assms(3)]) apply (metis assms(2) has_derivative_sum) using assms(4-5) unfolding sums_def apply auto done lemma has_field_derivative_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniform_limit S (\n x. \i S" "summable (\n. f n x0)" shows "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" unfolding has_field_derivative_def proof (rule has_derivative_series) show "\\<^sub>F n in sequentially. \x\S. \h. norm ((\i e * norm h" if "e > 0" for e unfolding eventually_sequentially proof - from that assms(3) obtain N where N: "\n x. n \ N \ x \ S \ norm ((\i N" "x \ S" have "norm ((\iii e" by simp hence "norm ((\i e * norm h" by (intro mult_right_mono) simp_all finally have "norm ((\i e * norm h" . } thus "\N. \n\N. \x\S. \h. norm ((\i e * norm h" by blast qed qed (use assms in \auto simp: has_field_derivative_def\) lemma has_field_derivative_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x within S)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" "x \ interior S" shows "summable (\n. f n x)" "((\x. \n. f n x) has_field_derivative (\n. f' n x)) (at x)" proof - from \x \ interior S\ have "x \ S" using interior_subset by blast define g' where [abs_def]: "g' x = (\i. f' i x)" for x from assms(3) have "uniform_limit S (\n x. \ix. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g(1)[OF \x \ S\] show "summable (\n. f n x)" by (simp add: sums_iff) from g(2)[OF \x \ S\] \x \ interior S\ have "(g has_field_derivative g' x) (at x)" by (simp add: at_within_interior[of x S]) also have "(g has_field_derivative g' x) (at x) \ ((\x. \n. f n x) has_field_derivative g' x) (at x)" using eventually_nhds_in_nhd[OF \x \ interior S\] interior_subset[of S] g(1) by (intro DERIV_cong_ev) (auto elim!: eventually_mono simp: sums_iff) finally show "((\x. \n. f n x) has_field_derivative g' x) (at x)" . qed lemma differentiable_series: fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" and x: "x \ S" shows "summable (\n. f n x)" and "(\x. \n. f n x) differentiable (at x)" proof - from assms(4) obtain g' where A: "uniform_limit S (\n x. \iopen S\ have S: "at x within S = at x" by (rule at_within_open) have "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g[OF x] show "summable (\n. f n x)" by (auto simp: summable_def) from g(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)" by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative (*) (g' x)) (at x)" by (rule has_derivative_transform_within_open[OF g' \open S\ x]) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def differentiable_def has_field_derivative_def) qed lemma differentiable_series': fixes f :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" shows "(\x. \n. f n x) differentiable (at x0)" using differentiable_series[OF assms, of x0] \x0 \ S\ by blast+ subsection \Derivative as a vector\ text \Considering derivative \<^typ>\real \ 'b::real_normed_vector\ as a vector.\ definition "vector_derivative f net = (SOME f'. (f has_vector_derivative f') net)" lemma vector_derivative_unique_within: assumes not_bot: "at x within S \ bot" and f': "(f has_vector_derivative f') (at x within S)" and f'': "(f has_vector_derivative f'') (at x within S)" shows "f' = f''" proof - have "(\x. x *\<^sub>R f') = (\x. x *\<^sub>R f'')" proof (rule frechet_derivative_unique_within, simp_all) show "\d. d \ 0 \ \d\ < e \ x + d \ S" if "0 < e" for e proof - from that obtain x' where "x' \ S" "x' \ x" "\x' - x\ < e" using islimpt_approachable_real[of x S] not_bot by (auto simp add: trivial_limit_within) then show ?thesis using eq_iff_diff_eq_0 by fastforce qed qed (use f' f'' in \auto simp: has_vector_derivative_def\) then show ?thesis unfolding fun_eq_iff by (metis scaleR_one) qed lemma vector_derivative_unique_at: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f'') (at x) \ f' = f''" by (rule vector_derivative_unique_within) auto lemma differentiableI_vector: "(f has_vector_derivative y) F \ f differentiable F" by (auto simp: differentiable_def has_vector_derivative_def) proposition vector_derivative_works: "f differentiable net \ (f has_vector_derivative (vector_derivative f net)) net" (is "?l = ?r") proof assume ?l obtain f' where f': "(f has_derivative f') net" using \?l\ unfolding differentiable_def .. then interpret bounded_linear f' by auto show ?r unfolding vector_derivative_def has_vector_derivative_def by (rule someI[of _ "f' 1"]) (simp add: scaleR[symmetric] f') qed (auto simp: vector_derivative_def has_vector_derivative_def differentiable_def) lemma vector_derivative_within: assumes not_bot: "at x within S \ bot" and y: "(f has_vector_derivative y) (at x within S)" shows "vector_derivative f (at x within S) = y" using y by (intro vector_derivative_unique_within[OF not_bot vector_derivative_works[THEN iffD1] y]) (auto simp: differentiable_def has_vector_derivative_def) lemma frechet_derivative_eq_vector_derivative: assumes "f differentiable (at x)" shows "(frechet_derivative f (at x)) = (\r. r *\<^sub>R vector_derivative f (at x))" using assms by (auto simp: differentiable_iff_scaleR vector_derivative_def has_vector_derivative_def intro: someI frechet_derivative_at [symmetric]) lemma has_real_derivative: fixes f :: "real \ real" assumes "(f has_derivative f') F" obtains c where "(f has_real_derivative c) F" proof - obtain c where "f' = (\x. x * c)" by (metis assms has_derivative_bounded_linear real_bounded_linear) then show ?thesis by (metis assms that has_field_derivative_def mult_commute_abs) qed lemma has_real_derivative_iff: fixes f :: "real \ real" shows "(\c. (f has_real_derivative c) F) = (\D. (f has_derivative D) F)" by (metis has_field_derivative_def has_real_derivative) lemma has_vector_derivative_cong_ev: assumes *: "eventually (\x. x \ S \ f x = g x) (nhds x)" "f x = g x" shows "(f has_vector_derivative f') (at x within S) = (g has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def has_derivative_def using * apply (cases "at x within S \ bot") apply (intro refl conj_cong filterlim_cong) apply (auto simp: Lim_ident_at eventually_at_filter elim: eventually_mono) done lemma islimpt_closure_open: fixes s :: "'a::perfect_space set" assumes "open s" and t: "t = closure s" "x \ t" shows "x islimpt t" proof cases assume "x \ s" { fix T assume "x \ T" "open T" then have "open (s \ T)" using \open s\ by auto then have "s \ T \ {x}" using not_open_singleton[of x] by auto with \x \ T\ \x \ s\ have "\y\t. y \ T \ y \ x" using closure_subset[of s] by (auto simp: t) } then show ?thesis by (auto intro!: islimptI) next assume "x \ s" with t show ?thesis unfolding t closure_def by (auto intro: islimpt_subset) qed lemma vector_derivative_unique_within_closed_interval: assumes ab: "a < b" "x \ cbox a b" assumes D: "(f has_vector_derivative f') (at x within cbox a b)" "(f has_vector_derivative f'') (at x within cbox a b)" shows "f' = f''" using ab by (intro vector_derivative_unique_within[OF _ D]) (auto simp: trivial_limit_within intro!: islimpt_closure_open[where s="{a <..< b}"]) lemma vector_derivative_at: "(f has_vector_derivative f') (at x) \ vector_derivative f (at x) = f'" by (intro vector_derivative_within at_neq_bot) lemma has_vector_derivative_id_at [simp]: "vector_derivative (\x. x) (at a) = 1" by (simp add: vector_derivative_at) lemma vector_derivative_minus_at [simp]: "f differentiable at a \ vector_derivative (\x. - f x) (at a) = - vector_derivative f (at a)" by (simp add: vector_derivative_at has_vector_derivative_minus vector_derivative_works [symmetric]) lemma vector_derivative_add_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x + g x) (at a) = vector_derivative f (at a) + vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_add vector_derivative_works [symmetric]) lemma vector_derivative_diff_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x - g x) (at a) = vector_derivative f (at a) - vector_derivative g (at a)" by (simp add: vector_derivative_at has_vector_derivative_diff vector_derivative_works [symmetric]) lemma vector_derivative_mult_at [simp]: fixes f g :: "real \ 'a :: real_normed_algebra" shows "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x * g x) (at a) = f a * vector_derivative g (at a) + vector_derivative f (at a) * g a" by (simp add: vector_derivative_at has_vector_derivative_mult vector_derivative_works [symmetric]) lemma vector_derivative_scaleR_at [simp]: "\f differentiable at a; g differentiable at a\ \ vector_derivative (\x. f x *\<^sub>R g x) (at a) = f a *\<^sub>R vector_derivative g (at a) + vector_derivative f (at a) *\<^sub>R g a" apply (rule vector_derivative_at) apply (rule has_vector_derivative_scaleR) apply (auto simp: vector_derivative_works has_vector_derivative_def has_field_derivative_def mult_commute_abs) done lemma vector_derivative_within_cbox: assumes ab: "a < b" "x \ cbox a b" assumes f: "(f has_vector_derivative f') (at x within cbox a b)" shows "vector_derivative f (at x within cbox a b) = f'" by (intro vector_derivative_unique_within_closed_interval[OF ab _ f] vector_derivative_works[THEN iffD1] differentiableI_vector) fact lemma vector_derivative_within_closed_interval: fixes f::"real \ 'a::euclidean_space" assumes "a < b" and "x \ {a..b}" assumes "(f has_vector_derivative f') (at x within {a..b})" shows "vector_derivative f (at x within {a..b}) = f'" using assms vector_derivative_within_cbox by fastforce lemma has_vector_derivative_within_subset: "(f has_vector_derivative f') (at x within S) \ T \ S \ (f has_vector_derivative f') (at x within T)" by (auto simp: has_vector_derivative_def intro: has_derivative_within_subset) lemma has_vector_derivative_at_within: "(f has_vector_derivative f') (at x) \ (f has_vector_derivative f') (at x within S)" unfolding has_vector_derivative_def by (rule has_derivative_at_withinI) lemma has_vector_derivative_weaken: fixes x D and f g S T assumes f: "(f has_vector_derivative D) (at x within T)" and "x \ S" "S \ T" and "\x. x \ S \ f x = g x" shows "(g has_vector_derivative D) (at x within S)" proof - have "(f has_vector_derivative D) (at x within S) \ (g has_vector_derivative D) (at x within S)" unfolding has_vector_derivative_def has_derivative_iff_norm using assms by (intro conj_cong Lim_cong_within refl) auto then show ?thesis using has_vector_derivative_within_subset[OF f \S \ T\] by simp qed lemma has_vector_derivative_transform_within: assumes "(f has_vector_derivative f') (at x within S)" and "0 < d" and "x \ S" and "\x'. \x'\S; dist x' x < d\ \ f x' = g x'" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within) lemma has_vector_derivative_transform_within_open: assumes "(f has_vector_derivative f') (at x)" and "open S" and "x \ S" and "\y. y\S \ f y = g y" shows "(g has_vector_derivative f') (at x)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform_within_open) lemma has_vector_derivative_transform: assumes "x \ S" "\x. x \ S \ g x = f x" assumes f': "(f has_vector_derivative f') (at x within S)" shows "(g has_vector_derivative f') (at x within S)" using assms unfolding has_vector_derivative_def by (rule has_derivative_transform) lemma vector_diff_chain_at: assumes "(f has_vector_derivative f') (at x)" and "(g has_vector_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x)" using assms has_vector_derivative_at_within has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_diff_chain_within: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at (f x) within f ` s)" shows "((g \ f) has_vector_derivative (f' *\<^sub>R g')) (at x within s)" using assms has_vector_derivative_def vector_derivative_diff_chain_within by blast lemma vector_derivative_const_at [simp]: "vector_derivative (\x. c) (at a) = 0" by (simp add: vector_derivative_at) lemma vector_derivative_at_within_ivl: "(f has_vector_derivative f') (at x) \ a \ x \ x \ b \ a vector_derivative f (at x within {a..b}) = f'" using has_vector_derivative_at_within vector_derivative_within_cbox by fastforce lemma vector_derivative_chain_at: assumes "f differentiable at x" "(g differentiable at (f x))" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" by (metis vector_diff_chain_at vector_derivative_at vector_derivative_works assms) lemma field_vector_diff_chain_at: (*thanks to Wenda Li*) assumes Df: "(f has_vector_derivative f') (at x)" and Dg: "(g has_field_derivative g') (at (f x))" shows "((g \ f) has_vector_derivative (f' * g')) (at x)" using diff_chain_at[OF Df[unfolded has_vector_derivative_def] Dg [unfolded has_field_derivative_def]] by (auto simp: o_def mult.commute has_vector_derivative_def) lemma vector_derivative_chain_within: assumes "at x within S \ bot" "f differentiable (at x within S)" "(g has_derivative g') (at (f x) within f ` S)" shows "vector_derivative (g \ f) (at x within S) = g' (vector_derivative f (at x within S)) " apply (rule vector_derivative_within [OF \at x within S \ bot\]) apply (rule vector_derivative_diff_chain_within) using assms(2-3) vector_derivative_works by auto subsection \Field differentiability\ definition\<^marker>\tag important\ field_differentiable :: "['a \ 'a::real_normed_field, 'a filter] \ bool" (infixr "(field'_differentiable)" 50) where "f field_differentiable F \ \f'. (f has_field_derivative f') F" lemma field_differentiable_imp_differentiable: "f field_differentiable F \ f differentiable F" unfolding field_differentiable_def differentiable_def using has_field_derivative_imp_has_derivative by auto lemma field_differentiable_imp_continuous_at: "f field_differentiable (at x within S) \ continuous (at x within S) f" by (metis DERIV_continuous field_differentiable_def) lemma field_differentiable_within_subset: "\f field_differentiable (at x within S); T \ S\ \ f field_differentiable (at x within T)" by (metis DERIV_subset field_differentiable_def) lemma field_differentiable_at_within: "\f field_differentiable (at x)\ \ f field_differentiable (at x within S)" unfolding field_differentiable_def by (metis DERIV_subset top_greatest) lemma field_differentiable_linear [simp,derivative_intros]: "((*) c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def mult_commute_abs by (force intro: has_derivative_mult_right) lemma field_differentiable_const [simp,derivative_intros]: "(\z. c) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_const has_field_derivative_imp_has_derivative by blast lemma field_differentiable_ident [simp,derivative_intros]: "(\z. z) field_differentiable F" unfolding field_differentiable_def has_field_derivative_def using DERIV_ident has_field_derivative_def by blast lemma field_differentiable_id [simp,derivative_intros]: "id field_differentiable F" unfolding id_def by (rule field_differentiable_ident) lemma field_differentiable_minus [derivative_intros]: "f field_differentiable F \ (\z. - (f z)) field_differentiable F" unfolding field_differentiable_def by (metis field_differentiable_minus) lemma field_differentiable_add [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z + g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_add) lemma field_differentiable_add_const [simp,derivative_intros]: "(+) c field_differentiable F" by (simp add: field_differentiable_add) lemma field_differentiable_sum [derivative_intros]: "(\i. i \ I \ (f i) field_differentiable F) \ (\z. \i\I. f i z) field_differentiable F" by (induct I rule: infinite_finite_induct) (auto intro: field_differentiable_add field_differentiable_const) lemma field_differentiable_diff [derivative_intros]: assumes "f field_differentiable F" "g field_differentiable F" shows "(\z. f z - g z) field_differentiable F" using assms unfolding field_differentiable_def by (metis field_differentiable_diff) lemma field_differentiable_inverse [derivative_intros]: assumes "f field_differentiable (at a within S)" "f a \ 0" shows "(\z. inverse (f z)) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_inverse_fun) lemma field_differentiable_mult [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" shows "(\z. f z * g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_mult [of f _ a S g]) lemma field_differentiable_divide [derivative_intros]: assumes "f field_differentiable (at a within S)" "g field_differentiable (at a within S)" "g a \ 0" shows "(\z. f z / g z) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_divide [of f _ a S g]) lemma field_differentiable_power [derivative_intros]: assumes "f field_differentiable (at a within S)" shows "(\z. f z ^ n) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_power) lemma field_differentiable_transform_within: "0 < d \ x \ S \ (\x'. x' \ S \ dist x' x < d \ f x' = g x') \ f field_differentiable (at x within S) \ g field_differentiable (at x within S)" unfolding field_differentiable_def has_field_derivative_def by (blast intro: has_derivative_transform_within) lemma field_differentiable_compose_within: assumes "f field_differentiable (at a within S)" "g field_differentiable (at (f a) within f`S)" shows "(g o f) field_differentiable (at a within S)" using assms unfolding field_differentiable_def by (metis DERIV_image_chain) lemma field_differentiable_compose: "f field_differentiable at z \ g field_differentiable at (f z) \ (g o f) field_differentiable at z" by (metis field_differentiable_at_within field_differentiable_compose_within) lemma field_differentiable_within_open: "\a \ S; open S\ \ f field_differentiable at a within S \ f field_differentiable at a" unfolding field_differentiable_def by (metis at_within_open) lemma exp_scaleR_has_vector_derivative_right: "((\t. exp (t *\<^sub>R A)) has_vector_derivative exp (t *\<^sub>R A) * A) (at t within T)" unfolding has_vector_derivative_def proof (rule has_derivativeI) let ?F = "at t within (T \ {t - 1 <..< t + 1})" have *: "at t within T = ?F" by (rule at_within_nhd[where S="{t - 1 <..< t + 1}"]) auto let ?e = "\i x. (inverse (1 + real i) * inverse (fact i) * (x - t) ^ i) *\<^sub>R (A * A ^ i)" have "\\<^sub>F n in sequentially. \x\T \ {t - 1<.. norm (A ^ (n + 1) /\<^sub>R fact (n + 1))" apply (auto simp: algebra_split_simps intro!: eventuallyI) apply (rule mult_left_mono) apply (auto simp add: field_simps power_abs intro!: divide_right_mono power_le_one) done then have "uniform_limit (T \ {t - 1<..n x. \ix. \i. ?e i x) sequentially" by (rule Weierstrass_m_test_ev) (intro summable_ignore_initial_segment summable_norm_exp) moreover have "\\<^sub>F x in sequentially. x > 0" by (metis eventually_gt_at_top) then have "\\<^sub>F n in sequentially. ((\x. \i A) ?F" by eventually_elim (auto intro!: tendsto_eq_intros simp: power_0_left if_distrib if_distribR cong: if_cong) ultimately have [tendsto_intros]: "((\x. \i. ?e i x) \ A) ?F" by (auto intro!: swap_uniform_limit[where f="\n x. \i < n. ?e i x" and F = sequentially]) have [tendsto_intros]: "((\x. if x = t then 0 else 1) \ 1) ?F" by (rule tendsto_eventually) (simp add: eventually_at_filter) have "((\y. ((y - t) / abs (y - t)) *\<^sub>R ((\n. ?e n y) - A)) \ 0) (at t within T)" unfolding * by (rule tendsto_norm_zero_cancel) (auto intro!: tendsto_eq_intros) moreover have "\\<^sub>F x in at t within T. x \ t" by (simp add: eventually_at_filter) then have "\\<^sub>F x in at t within T. ((x - t) / \x - t\) *\<^sub>R ((\n. ?e n x) - A) = (exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" proof eventually_elim case (elim x) have "(exp ((x - t) *\<^sub>R A) - 1 - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = ((\n. (x - t) *\<^sub>R ?e n x) - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t)" unfolding exp_first_term by (simp add: ac_simps) also have "summable (\n. ?e n x)" proof - from elim have "?e n x = (((x - t) *\<^sub>R A) ^ (n + 1)) /\<^sub>R fact (n + 1) /\<^sub>R (x - t)" for n by simp then show ?thesis by (auto simp only: intro!: summable_scaleR_right summable_ignore_initial_segment summable_exp_generic) qed then have "(\n. (x - t) *\<^sub>R ?e n x) = (x - t) *\<^sub>R (\n. ?e n x)" by (rule suminf_scaleR_right[symmetric]) also have "(\ - (x - t) *\<^sub>R A) /\<^sub>R norm (x - t) = (x - t) *\<^sub>R ((\n. ?e n x) - A) /\<^sub>R norm (x - t)" by (simp add: algebra_simps) finally show ?case by simp (simp add: field_simps) qed ultimately have "((\y. (exp ((y - t) *\<^sub>R A) - 1 - (y - t) *\<^sub>R A) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) from tendsto_mult_right_zero[OF this, where c="exp (t *\<^sub>R A)"] show "((\y. (exp (y *\<^sub>R A) - exp (t *\<^sub>R A) - (y - t) *\<^sub>R (exp (t *\<^sub>R A) * A)) /\<^sub>R norm (y - t)) \ 0) (at t within T)" by (rule Lim_transform_eventually) (auto simp: algebra_simps field_split_simps exp_add_commuting[symmetric]) qed (rule bounded_linear_scaleR_left) lemma exp_times_scaleR_commute: "exp (t *\<^sub>R A) * A = A * exp (t *\<^sub>R A)" using exp_times_arg_commute[symmetric, of "t *\<^sub>R A"] by (auto simp: algebra_simps) lemma exp_scaleR_has_vector_derivative_left: "((\t. exp (t *\<^sub>R A)) has_vector_derivative A * exp (t *\<^sub>R A)) (at t)" using exp_scaleR_has_vector_derivative_right[of A t] by (simp add: exp_times_scaleR_commute) lemma field_differentiable_series: fixes f :: "nat \ 'a::{real_normed_field,banach} \ 'a" assumes "convex S" "open S" assumes "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" assumes "uniformly_convergent_on S (\n x. \i S" "summable (\n. f n x0)" and x: "x \ S" shows "(\x. \n. f n x) field_differentiable (at x)" proof - from assms(4) obtain g' where A: "uniform_limit S (\n x. \iopen S\ have S: "at x within S = at x" by (rule at_within_open) have "\g. \x\S. (\n. f n x) sums g x \ (g has_field_derivative g' x) (at x within S)" by (intro has_field_derivative_series[of S f f' g' x0] assms A has_field_derivative_at_within) then obtain g where g: "\x. x \ S \ (\n. f n x) sums g x" "\x. x \ S \ (g has_field_derivative g' x) (at x within S)" by blast from g(2)[OF x] have g': "(g has_derivative (*) (g' x)) (at x)" by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative (*) (g' x)) (at x)" by (rule has_derivative_transform_within_open[OF g' \open S\ x]) (insert g, auto simp: sums_iff) thus "(\x. \n. f n x) field_differentiable (at x)" unfolding differentiable_def by (auto simp: summable_def field_differentiable_def has_field_derivative_def) qed subsubsection\<^marker>\tag unimportant\\Caratheodory characterization\ lemma field_differentiable_caratheodory_at: "f field_differentiable (at z) \ (\g. (\w. f(w) - f(z) = g(w) * (w - z)) \ continuous (at z) g)" using CARAT_DERIV [of f] by (simp add: field_differentiable_def has_field_derivative_def) lemma field_differentiable_caratheodory_within: "f field_differentiable (at z within s) \ (\g. (\w. f(w) - f(z) = g(w) * (w - z)) \ continuous (at z within s) g)" using DERIV_caratheodory_within [of f] by (simp add: field_differentiable_def has_field_derivative_def) subsection \Field derivative\ definition\<^marker>\tag important\ deriv :: "('a \ 'a::real_normed_field) \ 'a \ 'a" where "deriv f x \ SOME D. DERIV f x :> D" lemma DERIV_imp_deriv: "DERIV f x :> f' \ deriv f x = f'" unfolding deriv_def by (metis some_equality DERIV_unique) lemma DERIV_deriv_iff_has_field_derivative: "DERIV f x :> deriv f x \ (\f'. (f has_field_derivative f') (at x))" by (auto simp: has_field_derivative_def DERIV_imp_deriv) lemma DERIV_deriv_iff_real_differentiable: fixes x :: real shows "DERIV f x :> deriv f x \ f differentiable at x" unfolding differentiable_def by (metis DERIV_imp_deriv has_real_derivative_iff) lemma deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "deriv f x = deriv g y" proof - have "(\D. (f has_field_derivative D) (at x)) = (\D. (g has_field_derivative D) (at y))" by (intro ext DERIV_cong_ev refl assms) thus ?thesis by (simp add: deriv_def assms) qed lemma higher_deriv_cong_ev: assumes "eventually (\x. f x = g x) (nhds x)" "x = y" shows "(deriv ^^ n) f x = (deriv ^^ n) g y" proof - from assms(1) have "eventually (\x. (deriv ^^ n) f x = (deriv ^^ n) g x) (nhds x)" proof (induction n arbitrary: f g) case (Suc n) from Suc.prems have "eventually (\y. eventually (\z. f z = g z) (nhds y)) (nhds x)" by (simp add: eventually_eventually) hence "eventually (\x. deriv f x = deriv g x) (nhds x)" by eventually_elim (rule deriv_cong_ev, simp_all) thus ?case by (auto intro!: deriv_cong_ev Suc simp: funpow_Suc_right simp del: funpow.simps) qed auto from eventually_nhds_x_imp_x[OF this] assms(2) show ?thesis by simp qed lemma real_derivative_chain: fixes x :: real shows "f differentiable at x \ g differentiable at (f x) \ deriv (g o f) x = deriv g (f x) * deriv f x" by (metis DERIV_deriv_iff_real_differentiable DERIV_chain DERIV_imp_deriv) lemma field_derivative_eq_vector_derivative: "(deriv f x) = vector_derivative f (at x)" by (simp add: mult.commute deriv_def vector_derivative_def has_vector_derivative_def has_field_derivative_def) proposition field_differentiable_derivI: "f field_differentiable (at x) \ (f has_field_derivative deriv f x) (at x)" by (simp add: field_differentiable_def DERIV_deriv_iff_has_field_derivative) lemma vector_derivative_chain_at_general: assumes "f differentiable at x" "g field_differentiable at (f x)" shows "vector_derivative (g \ f) (at x) = vector_derivative f (at x) * deriv g (f x)" apply (rule vector_derivative_at [OF field_vector_diff_chain_at]) using assms vector_derivative_works by (auto simp: field_differentiable_derivI) +lemma DERIV_deriv_iff_field_differentiable: + "DERIV f x :> deriv f x \ f field_differentiable at x" + unfolding field_differentiable_def by (metis DERIV_imp_deriv) + +lemma deriv_chain: + "f field_differentiable at x \ g field_differentiable at (f x) + \ deriv (g o f) x = deriv g (f x) * deriv f x" + by (metis DERIV_deriv_iff_field_differentiable DERIV_chain DERIV_imp_deriv) + +lemma deriv_linear [simp]: "deriv (\w. c * w) = (\z. c)" + by (metis DERIV_imp_deriv DERIV_cmult_Id) + +lemma deriv_uminus [simp]: "deriv (\w. -w) = (\z. -1)" + using deriv_linear[of "-1"] by (simp del: deriv_linear) + +lemma deriv_ident [simp]: "deriv (\w. w) = (\z. 1)" + by (metis DERIV_imp_deriv DERIV_ident) + +lemma deriv_id [simp]: "deriv id = (\z. 1)" + by (simp add: id_def) + +lemma deriv_const [simp]: "deriv (\w. c) = (\z. 0)" + by (metis DERIV_imp_deriv DERIV_const) + +lemma deriv_add [simp]: + "\f field_differentiable at z; g field_differentiable at z\ + \ deriv (\w. f w + g w) z = deriv f z + deriv g z" + unfolding DERIV_deriv_iff_field_differentiable[symmetric] + by (auto intro!: DERIV_imp_deriv derivative_intros) + +lemma deriv_diff [simp]: + "\f field_differentiable at z; g field_differentiable at z\ + \ deriv (\w. f w - g w) z = deriv f z - deriv g z" + unfolding DERIV_deriv_iff_field_differentiable[symmetric] + by (auto intro!: DERIV_imp_deriv derivative_intros) + +lemma deriv_mult [simp]: + "\f field_differentiable at z; g field_differentiable at z\ + \ deriv (\w. f w * g w) z = f z * deriv g z + deriv f z * g z" + unfolding DERIV_deriv_iff_field_differentiable[symmetric] + by (auto intro!: DERIV_imp_deriv derivative_eq_intros) + +lemma deriv_cmult: + "f field_differentiable at z \ deriv (\w. c * f w) z = c * deriv f z" + by simp + +lemma deriv_cmult_right: + "f field_differentiable at z \ deriv (\w. f w * c) z = deriv f z * c" + by simp + +lemma deriv_inverse [simp]: + "\f field_differentiable at z; f z \ 0\ + \ deriv (\w. inverse (f w)) z = - deriv f z / f z ^ 2" + unfolding DERIV_deriv_iff_field_differentiable[symmetric] + by (safe intro!: DERIV_imp_deriv derivative_eq_intros) (auto simp: field_split_simps power2_eq_square) + +lemma deriv_divide [simp]: + "\f field_differentiable at z; g field_differentiable at z; g z \ 0\ + \ deriv (\w. f w / g w) z = (deriv f z * g z - f z * deriv g z) / g z ^ 2" + by (simp add: field_class.field_divide_inverse field_differentiable_inverse) + (simp add: field_split_simps power2_eq_square) + +lemma deriv_cdivide_right: + "f field_differentiable at z \ deriv (\w. f w / c) z = deriv f z / c" + by (simp add: field_class.field_divide_inverse) + +lemma deriv_compose_linear: + "f field_differentiable at (c * z) \ deriv (\w. f (c * w)) z = c * deriv f (c * z)" +apply (rule DERIV_imp_deriv) + unfolding DERIV_deriv_iff_field_differentiable [symmetric] + by (metis (full_types) DERIV_chain2 DERIV_cmult_Id mult.commute) + + +lemma nonzero_deriv_nonconstant: + assumes df: "DERIV f \ :> df" and S: "open S" "\ \ S" and "df \ 0" + shows "\ f constant_on S" +unfolding constant_on_def +by (metis \df \ 0\ has_field_derivative_transform_within_open [OF df S] DERIV_const DERIV_unique) + + subsection \Relation between convexity and derivative\ (* TODO: Generalise to real vector spaces? *) proposition convex_on_imp_above_tangent: assumes convex: "convex_on A f" and connected: "connected A" assumes c: "c \ interior A" and x : "x \ A" assumes deriv: "(f has_field_derivative f') (at c within A)" shows "f x - f c \ f' * (x - c)" proof (cases x c rule: linorder_cases) assume xc: "x > c" let ?A' = "interior A \ {c<..}" from c have "c \ interior A \ closure {c<..}" by auto also have "\ \ closure (interior A \ {c<..})" by (intro open_Int_closure_subset) auto finally have "at c within ?A' \ bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_right_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_right c)" proof eventually_elim fix y assume y: "y \ {c<.. (f x - f c) / (x - c) * (y - c) + f c" using interior_subset[of A] by (intro convex_onD_Icc' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) / (x - c) * (y - c)" by simp thus "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_upperbound) thus ?thesis using xc by (simp add: field_simps) next assume xc: "x < c" let ?A' = "interior A \ {.. interior A \ closure {.. \ closure (interior A \ {.. bot" by (subst at_within_eq_bot_iff) auto moreover from deriv have "((\y. (f y - f c) / (y - c)) \ f') (at c within ?A')" unfolding has_field_derivative_iff using interior_subset[of A] by (blast intro: tendsto_mono at_le) moreover from eventually_at_left_real[OF xc] have "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at_left c)" proof eventually_elim fix y assume y: "y \ {x<.. (f x - f c) / (c - x) * (c - y) + f c" using interior_subset[of A] by (intro convex_onD_Icc'' convex_on_subset[OF convex] connected_contains_Icc) auto hence "f y - f c \ (f x - f c) * ((c - y) / (c - x))" by simp also have "(c - y) / (c - x) = (y - c) / (x - c)" using y xc by (simp add: field_simps) finally show "(f y - f c) / (y - c) \ (f x - f c) / (x - c)" using y xc by (simp add: field_split_simps) qed hence "eventually (\y. (f y - f c) / (y - c) \ (f x - f c) / (x - c)) (at c within ?A')" by (blast intro: filter_leD at_le) ultimately have "f' \ (f x - f c) / (x - c)" by (simp add: tendsto_lowerbound) thus ?thesis using xc by (simp add: field_simps) qed simp_all subsection \Partial derivatives\ lemma eventually_at_Pair_within_TimesI1: fixes x::"'a::metric_space" assumes "\\<^sub>F x' in at x within X. P x'" assumes "P x" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "x \ S" "\x'. x' \ X \ x' \ S \ P x'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P x'" unfolding eventually_at_topological by (auto intro!: exI[where x="S \ UNIV"] S open_Times) qed lemma eventually_at_Pair_within_TimesI2: fixes x::"'a::metric_space" assumes "\\<^sub>F y' in at y within Y. P y'" "P y" shows "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" proof - from assms[unfolded eventually_at_topological] obtain S where S: "open S" "y \ S" "\y'. y' \ Y \ y' \ S \ P y'" by metis show "\\<^sub>F (x', y') in at (x, y) within X \ Y. P y'" unfolding eventually_at_topological by (auto intro!: exI[where x="UNIV \ S"] S open_Times) qed proposition has_derivative_partialsI: fixes f::"'a::real_normed_vector \ 'b::real_normed_vector \ 'c::real_normed_vector" assumes fx: "((\x. f x y) has_derivative fx) (at x within X)" assumes fy: "\x y. x \ X \ y \ Y \ ((\y. f x y) has_derivative blinfun_apply (fy x y)) (at y within Y)" assumes fy_cont[unfolded continuous_within]: "continuous (at (x, y) within X \ Y) (\(x, y). fy x y)" assumes "y \ Y" "convex Y" shows "((\(x, y). f x y) has_derivative (\(tx, ty). fx tx + fy x y ty)) (at (x, y) within X \ Y)" proof (safe intro!: has_derivativeI tendstoI, goal_cases) case (2 e') interpret fx: bounded_linear "fx" using fx by (rule has_derivative_bounded_linear) define e where "e = e' / 9" have "e > 0" using \e' > 0\ by (simp add: e_def) from fy_cont[THEN tendstoD, OF \e > 0\] have "\\<^sub>F (x', y') in at (x, y) within X \ Y. dist (fy x' y') (fy x y) < e" by (auto simp: split_beta') from this[unfolded eventually_at] obtain d' where "d' > 0" "\x' y'. x' \ X \ y' \ Y \ (x', y') \ (x, y) \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" by auto then have d': "x' \ X \ y' \ Y \ dist (x', y') (x, y) < d' \ dist (fy x' y') (fy x y) < e" for x' y' using \0 < e\ by (cases "(x', y') = (x, y)") auto define d where "d = d' / sqrt 2" have "d > 0" using \0 < d'\ by (simp add: d_def) have d: "x' \ X \ y' \ Y \ dist x' x < d \ dist y' y < d \ dist (fy x' y') (fy x y) < e" for x' y' by (auto simp: dist_prod_def d_def intro!: d' real_sqrt_sum_squares_less) let ?S = "ball y d \ Y" have "convex ?S" by (auto intro!: convex_Int \convex Y\) { fix x'::'a and y'::'b assume x': "x' \ X" and y': "y' \ Y" assume dx': "dist x' x < d" and dy': "dist y' y < d" have "norm (fy x' y' - fy x' y) \ dist (fy x' y') (fy x y) + dist (fy x' y) (fy x y)" by norm also have "dist (fy x' y') (fy x y) < e" by (rule d; fact) also have "dist (fy x' y) (fy x y) < e" by (auto intro!: d simp: dist_prod_def x' \d > 0\ \y \ Y\ dx') finally have "norm (fy x' y' - fy x' y) < e + e" by arith then have "onorm (blinfun_apply (fy x' y') - blinfun_apply (fy x' y)) < e + e" by (auto simp: norm_blinfun.rep_eq blinfun.diff_left[abs_def] fun_diff_def) } note onorm = this have ev_mem: "\\<^sub>F (x', y') in at (x, y) within X \ Y. (x', y') \ X \ Y" using \y \ Y\ by (auto simp: eventually_at intro!: zero_less_one) moreover have ev_dist: "\\<^sub>F xy in at (x, y) within X \ Y. dist xy (x, y) < d" if "d > 0" for d using eventually_at_ball[OF that] by (rule eventually_elim2) (auto simp: dist_commute intro!: eventually_True) note ev_dist[OF \0 < d\] ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" proof (eventually_elim, safe) fix x' y' assume "x' \ X" and y': "y' \ Y" assume dist: "dist (x', y') (x, y) < d" then have dx: "dist x' x < d" and dy: "dist y' y < d" unfolding dist_prod_def fst_conv snd_conv atomize_conj by (metis le_less_trans real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) { fix t::real assume "t \ {0 .. 1}" then have "y + t *\<^sub>R (y' - y) \ closed_segment y y'" by (auto simp: closed_segment_def algebra_simps intro!: exI[where x=t]) also have "\ \ ball y d \ Y" using \y \ Y\ \0 < d\ dy y' by (intro \convex ?S\[unfolded convex_contains_segment, rule_format, of y y']) (auto simp: dist_commute) finally have "y + t *\<^sub>R (y' - y) \ ?S" . } note seg = this have "\x. x \ ball y d \ Y \ onorm (blinfun_apply (fy x' x) - blinfun_apply (fy x' y)) \ e + e" by (safe intro!: onorm less_imp_le \x' \ X\ dx) (auto simp: dist_commute \0 < d\ \y \ Y\) with seg has_derivative_within_subset[OF assms(2)[OF \x' \ X\]] show "norm (f x' y' - f x' y - (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" by (rule differentiable_bound_linearization[where S="?S"]) (auto intro!: \0 < d\ \y \ Y\) qed moreover let ?le = "\x'. norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" from fx[unfolded has_derivative_within, THEN conjunct2, THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. ?le x'" by eventually_elim (simp, simp add: dist_norm field_split_simps split: if_split_asm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. ?le x'" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps) moreover have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((x', y') - (x, y)) \ 0" unfolding norm_eq_zero right_minus_eq by (auto simp: eventually_at intro!: zero_less_one) moreover from fy_cont[THEN tendstoD, OF \0 < e\] have "\\<^sub>F x' in at x within X. norm (fy x' y - fy x y) < e" unfolding eventually_at using \y \ Y\ by (auto simp: dist_prod_def dist_norm) then have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm (fy x' y - fy x y) < e" by (rule eventually_at_Pair_within_TimesI1) (simp add: blinfun.bilinear_simps \0 < e\) ultimately have "\\<^sub>F (x', y') in at (x, y) within X \ Y. norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" apply eventually_elim proof safe fix x' y' have "norm (f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) \ norm (f x' y' - f x' y - fy x' y (y' - y)) + norm (fy x y (y' - y) - fy x' y (y' - y)) + norm (f x' y - f x y - fx (x' - x))" by norm also assume nz: "norm ((x', y') - (x, y)) \ 0" and nfy: "norm (fy x' y - fy x y) < e" assume "norm (f x' y' - f x' y - blinfun_apply (fy x' y) (y' - y)) \ norm (y' - y) * (e + e)" also assume "norm (f x' y - f x y - (fx) (x' - x)) \ norm (x' - x) * e" also have "norm ((fy x y) (y' - y) - (fy x' y) (y' - y)) \ norm ((fy x y) - (fy x' y)) * norm (y' - y)" by (auto simp: blinfun.bilinear_simps[symmetric] intro!: norm_blinfun) also have "\ \ (e + e) * norm (y' - y)" using \e > 0\ nfy by (auto simp: norm_minus_commute intro!: mult_right_mono) also have "norm (x' - x) * e \ norm (x' - x) * (e + e)" using \0 < e\ by simp also have "norm (y' - y) * (e + e) + (e + e) * norm (y' - y) + norm (x' - x) * (e + e) \ (norm (y' - y) + norm (x' - x)) * (4 * e)" using \e > 0\ by (simp add: algebra_simps) also have "\ \ 2 * norm ((x', y') - (x, y)) * (4 * e)" using \0 < e\ real_sqrt_sum_squares_ge1[of "norm (x' - x)" "norm (y' - y)"] real_sqrt_sum_squares_ge2[of "norm (y' - y)" "norm (x' - x)"] by (auto intro!: mult_right_mono simp: norm_prod_def simp del: real_sqrt_sum_squares_ge1 real_sqrt_sum_squares_ge2) also have "\ \ norm ((x', y') - (x, y)) * (8 * e)" by simp also have "\ < norm ((x', y') - (x, y)) * e'" using \0 < e'\ nz by (auto simp: e_def) finally show "norm ((f x' y' - f x y - (fx (x' - x) + fy x y (y' - y))) /\<^sub>R norm ((x', y') - (x, y))) < e'" by (simp add: dist_norm) (auto simp add: field_split_simps) qed then show ?case by eventually_elim (auto simp: dist_norm field_simps) next from has_derivative_bounded_linear[OF fx] obtain fxb where "fx = blinfun_apply fxb" by (metis bounded_linear_Blinfun_apply) then show "bounded_linear (\(tx, ty). fx tx + blinfun_apply (fy x y) ty)" by (auto intro!: bounded_linear_intros simp: split_beta') qed subsection\<^marker>\tag unimportant\ \Differentiable case distinction\ lemma has_derivative_within_If_eq: "((\x. if P x then f x else g x) has_derivative f') (at x within S) = (bounded_linear f' \ ((\y.(if P y then (f y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x) else (g y - ((if P x then f x else g x) + f' (y - x)))/\<^sub>R norm (y - x))) \ 0) (at x within S))" (is "_ = (_ \ (?if \ 0) _)") proof - have "(\y. (1 / norm (y - x)) *\<^sub>R ((if P y then f y else g y) - ((if P x then f x else g x) + f' (y - x)))) = ?if" by (auto simp: inverse_eq_divide) thus ?thesis by (auto simp: has_derivative_within) qed lemma has_derivative_If_within_closures: assumes f': "x \ S \ (closure S \ closure T) \ (f has_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" assumes x_in: "x \ S \ T" shows "((\x. if x \ S then f x else g x) has_derivative (if x \ S then f' x else g' x)) (at x within (S \ T))" proof - from f' x_in interpret f': bounded_linear "if x \ S then f' x else (\x. 0)" by (auto simp add: has_derivative_within) from g' interpret g': bounded_linear "if x \ T then g' x else (\x. 0)" by (auto simp add: has_derivative_within) have bl: "bounded_linear (if x \ S then f' x else g' x)" using f'.scaleR f'.bounded f'.add g'.scaleR g'.bounded g'.add x_in by (unfold_locales; force) show ?thesis using f' g' closure_subset[of T] closure_subset[of S] unfolding has_derivative_within_If_eq by (intro conjI bl tendsto_If_within_closures x_in) (auto simp: has_derivative_within inverse_eq_divide connect connect' subsetD) qed lemma has_vector_derivative_If_within_closures: assumes x_in: "x \ S \ T" assumes "u = S \ T" assumes f': "x \ S \ (closure S \ closure T) \ (f has_vector_derivative f' x) (at x within S \ (closure S \ closure T))" assumes g': "x \ T \ (closure S \ closure T) \ (g has_vector_derivative g' x) (at x within T \ (closure S \ closure T))" assumes connect: "x \ closure S \ x \ closure T \ f x = g x" assumes connect': "x \ closure S \ x \ closure T \ f' x = g' x" shows "((\x. if x \ S then f x else g x) has_vector_derivative (if x \ S then f' x else g' x)) (at x within u)" unfolding has_vector_derivative_def assms using x_in apply (intro has_derivative_If_within_closures[where ?f' = "\x a. a *\<^sub>R f' x" and ?g' = "\x a. a *\<^sub>R g' x", THEN has_derivative_eq_rhs]) subgoal by (rule f'[unfolded has_vector_derivative_def]; assumption) subgoal by (rule g'[unfolded has_vector_derivative_def]; assumption) by (auto simp: assms) subsection\<^marker>\tag important\\The Inverse Function Theorem\ lemma linear_injective_contraction: assumes "linear f" "c < 1" and le: "\x. norm (f x - x) \ c * norm x" shows "inj f" unfolding linear_injective_0[OF \linear f\] proof safe fix x assume "f x = 0" with le [of x] have "norm x \ c * norm x" by simp then show "x = 0" using \c < 1\ by (simp add: mult_le_cancel_right1) qed text\From an online proof by J. Michael Boardman, Department of Mathematics, Johns Hopkins University\ lemma inverse_function_theorem_scaled: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative blinfun_apply (f' x)) (at x)" and contf: "continuous_on U f'" and "0 \ U" and [simp]: "f 0 = 0" and id: "f' 0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "0 \ U'" "open V" "0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - obtain d1 where "cball 0 d1 \ U" "d1 > 0" using \open U\ \0 \ U\ open_contains_cball by blast obtain d2 where d2: "\x. \x \ U; dist x 0 \ d2\ \ dist (f' x) (f' 0) < 1/2" "0 < d2" using continuous_onE [OF contf, of 0 "1/2"] by (metis \0 \ U\ half_gt_zero_iff zero_less_one) obtain \ where le: "\x. norm x \ \ \ dist (f' x) id_blinfun \ 1/2" and "0 < \" and subU: "cball 0 \ \ U" proof show "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) show "cball 0 (min d1 d2) \ U" using \cball 0 d1 \ U\ by auto show "dist (f' x) id_blinfun \ 1/2" if "norm x \ min d1 d2" for x using \cball 0 d1 \ U\ d2 that id by fastforce qed let ?D = "cball 0 \" define V:: "'a set" where "V \ ball 0 (\/2)" have 4: "norm (f (x + h) - f x - h) \ 1/2 * norm h" if "x \ ?D" "x+h \ ?D" for x h proof - let ?w = "\x. f x - x" have B: "\x. x \ ?D \ onorm (blinfun_apply (f' x - id_blinfun)) \ 1/2" by (metis dist_norm le mem_cball_0 norm_blinfun.rep_eq) have "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x)" by (rule derivative_eq_intros derf subsetD [OF subU] | force simp: blinfun.diff_left)+ then have Dw: "\x. x \ ?D \ (?w has_derivative (blinfun_apply (f' x - id_blinfun))) (at x within ?D)" using has_derivative_at_withinI by blast have "norm (?w (x+h) - ?w x) \ (1/2) * norm h" using differentiable_bound [OF convex_cball Dw B] that by fastforce then show ?thesis by (auto simp: algebra_simps) qed have for_g: "\!x. norm x < \ \ f x = y" if y: "norm y < \/2" for y proof - let ?u = "\x. x + (y - f x)" have *: "norm (?u x) < \" if "x \ ?D" for x proof - have fxx: "norm (f x - x) \ \/2" using 4 [of 0 x] \0 < \\ \f 0 = 0\ that by auto have "norm (?u x) \ norm y + norm (f x - x)" by (metis add.commute add_diff_eq norm_minus_commute norm_triangle_ineq) also have "\ < \/2 + \/2" using fxx y by auto finally show ?thesis by simp qed have "\!x \ ?D. ?u x = x" proof (rule banach_fix) show "cball 0 \ \ {}" using \0 < \\ by auto show "(\x. x + (y - f x)) ` cball 0 \ \ cball 0 \" using * by force have "dist (x + (y - f x)) (xh + (y - f xh)) * 2 \ dist x xh" if "norm x \ \" and "norm xh \ \" for x xh using that 4 [of x "xh-x"] by (auto simp: dist_norm norm_minus_commute algebra_simps) then show "\x\cball 0 \. \ya\cball 0 \. dist (x + (y - f x)) (ya + (y - f ya)) \ (1/2) * dist x ya" by auto qed (auto simp: complete_eq_closed) then show ?thesis by (metis "*" add_cancel_right_right eq_iff_diff_eq_0 le_less mem_cball_0) qed define g where "g \ \y. THE x. norm x < \ \ f x = y" have g: "norm (g y) < \ \ f (g y) = y" if "norm y < \/2" for y unfolding g_def using that theI' [OF for_g] by meson then have fg[simp]: "f (g y) = y" if "y \ V" for y using that by (auto simp: V_def) have 5: "norm (g y' - g y) \ 2 * norm (y' - y)" if "y \ V" "y' \ V" for y y' proof - have no: "norm (g y) \ \" "norm (g y') \ \" and [simp]: "f (g y) = y" using that g unfolding V_def by force+ have "norm (g y' - g y) \ norm (g y' - g y - (y' - y)) + norm (y' - y)" by (simp add: add.commute norm_triangle_sub) also have "\ \ (1/2) * norm (g y' - g y) + norm (y' - y)" using 4 [of "g y" "g y' - g y"] that no by (simp add: g norm_minus_commute V_def) finally show ?thesis by auto qed have contg: "continuous_on V g" proof fix y::'a and e::real assume "0 < e" and y: "y \ V" show "\d>0. \x'\V. dist x' y < d \ dist (g x') (g y) \ e" proof (intro exI conjI ballI impI) show "0 < e/2" by (simp add: \0 < e\) qed (use 5 y in \force simp: dist_norm\) qed show thesis proof define U' where "U' \ (f -` V) \ ball 0 \" have contf: "continuous_on U f" using derf has_derivative_at_withinI by (fast intro: has_derivative_continuous_on) then have "continuous_on (ball 0 \) f" by (meson ball_subset_cball continuous_on_subset subU) then show "open U'" by (simp add: U'_def V_def Int_commute continuous_open_preimage) show "0 \ U'" "U' \ U" "open V" "0 \ V" using \0 < \\ subU by (auto simp: U'_def V_def) show hom: "homeomorphism U' V f g" proof show "continuous_on U' f" using \U' \ U\ contf continuous_on_subset by blast show "continuous_on V g" using contg by blast show "f ` U' \ V" using U'_def by blast show "g ` V \ U'" by (simp add: U'_def V_def g image_subset_iff) show "g (f x) = x" if "x \ U'" for x by (metis that fg Int_iff U'_def V_def for_g g mem_ball_0 vimage_eq) show "f (g y) = y" if "y \ V" for y using that by (simp add: g V_def) qed show bij: "bij (blinfun_apply (f'(g y)))" if "y \ V" for y proof - have inj: "inj (blinfun_apply (f' (g y)))" proof (rule linear_injective_contraction) show "linear (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by blast next fix x have "norm (blinfun_apply (f' (g y)) x - x) = norm (blinfun_apply (f' (g y) - id_blinfun) x)" by (simp add: blinfun.diff_left) also have "\ \ norm (f' (g y) - id_blinfun) * norm x" by (rule norm_blinfun) also have "\ \ (1/2) * norm x" proof (rule mult_right_mono) show "norm (f' (g y) - id_blinfun) \ 1/2" using that g [of y] le by (auto simp: V_def dist_norm) qed auto finally show "norm (blinfun_apply (f' (g y)) x - x) \ (1/2) * norm x" . qed auto moreover have "surj (blinfun_apply (f' (g y)))" using blinfun.bounded_linear_right bounded_linear_def by (blast intro!: linear_inj_imp_surj [OF _ inj]) ultimately show ?thesis using bijI by blast qed define g' where "g' \ \y. inv (blinfun_apply (f'(g y)))" show "(g has_derivative g' y) (at y)" if "y \ V" for y proof - have gy: "g y \ U" using g subU that unfolding V_def by fastforce obtain e where e: "\h. f (g y + h) = y + blinfun_apply (f' (g y)) h + e h" and e0: "(\h. norm (e h) / norm h) \0\ 0" using iffD1 [OF has_derivative_iff_Ex derf [OF gy]] \y \ V\ by auto have [simp]: "e 0 = 0" using e [of 0] that by simp let ?INV = "inv (blinfun_apply (f' (g y)))" have inj: "inj (blinfun_apply (f' (g y)))" using bij bij_betw_def that by blast have "(g has_derivative g' y) (at y within V)" unfolding has_derivative_at_within_iff_Ex [OF \y \ V\ \open V\] proof show blinv: "bounded_linear (g' y)" unfolding g'_def using derf gy inj inj_linear_imp_inv_bounded_linear by blast define eg where "eg \ \k. - ?INV (e (g (y+k) - g y))" have "g (y+k) = g y + g' y k + eg k" if "y + k \ V" for k proof - have "?INV k = ?INV (blinfun_apply (f' (g y)) (g (y+k) - g y) + e (g (y+k) - g y))" using e [of "g(y+k) - g y"] that by simp then have "g (y+k) = g y + ?INV k - ?INV (e (g (y+k) - g y))" using inj blinv by (simp add: linear_simps g'_def) then show ?thesis by (auto simp: eg_def g'_def) qed moreover have "(\k. norm (eg k) / norm k) \0\ 0" proof (rule Lim_null_comparison) let ?g = "\k. 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" show "\\<^sub>F k in at 0. norm (norm (eg k) / norm k) \ ?g k" unfolding eventually_at_topological proof (intro exI conjI ballI impI) show "open ((+)(-y) ` V)" using \open V\ open_translation by blast show "0 \ (+)(-y) ` V" by (simp add: that) show "norm (norm (eg k) / norm k) \ 2 * onorm (inv (blinfun_apply (f' (g y)))) * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" if "k \ (+)(-y) ` V" "k \ 0" for k proof - have "y+k \ V" using that by auto have "norm (norm (eg k) / norm k) \ onorm ?INV * norm (e (g (y+k) - g y)) / norm k" using blinv g'_def onorm by (force simp: eg_def divide_simps) also have "\ = (norm (g (y+k) - g y) / norm k) * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" by (simp add: divide_simps) also have "\ \ 2 * (onorm ?INV * (norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)))" apply (rule mult_right_mono) using 5 [of y "y+k"] \y \ V\ \y + k \ V\ onorm_pos_le [OF blinv] apply (auto simp: divide_simps zero_le_mult_iff zero_le_divide_iff g'_def) done finally show "norm (norm (eg k) / norm k) \ 2 * onorm ?INV * norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)" by simp qed qed have 1: "(\h. norm (e h) / norm h) \0\ (norm (e 0) / norm 0)" using e0 by auto have 2: "(\k. g (y+k) - g y) \0\ 0" using contg \open V\ \y \ V\ LIM_offset_zero_iff LIM_zero_iff at_within_open continuous_on_def by fastforce from tendsto_compose [OF 1 2, simplified] have "(\k. norm (e (g (y+k) - g y)) / norm (g (y+k) - g y)) \0\ 0" . from tendsto_mult_left [OF this] show "?g \0\ 0" by auto qed ultimately show "\e. (\k. y + k \ V \ g (y+k) = g y + g' y k + e k) \ (\k. norm (e k) / norm k) \0\ 0" by blast qed then show ?thesis by (metis \open V\ at_within_open that) qed show "g' y = inv (blinfun_apply (f' (g y)))" if "y \ V" for y by (simp add: g'_def) qed qed text\We need all this to justify the scaling and translations.\ theorem inverse_function_theorem: fixes f::"'a::euclidean_space \ 'a" and f'::"'a \ ('a \\<^sub>L 'a)" assumes "open U" and derf: "\x. x \ U \ (f has_derivative (blinfun_apply (f' x))) (at x)" and contf: "continuous_on U f'" and "x0 \ U" and invf: "invf o\<^sub>L f' x0 = id_blinfun" obtains U' V g g' where "open U'" "U' \ U" "x0 \ U'" "open V" "f x0 \ V" "homeomorphism U' V f g" "\y. y \ V \ (g has_derivative (g' y)) (at y)" "\y. y \ V \ g' y = inv (blinfun_apply (f'(g y)))" "\y. y \ V \ bij (blinfun_apply (f'(g y)))" proof - have apply1 [simp]: "\i. blinfun_apply invf (blinfun_apply (f' x0) i) = i" by (metis blinfun_apply_blinfun_compose blinfun_apply_id_blinfun invf) have apply2 [simp]: "\i. blinfun_apply (f' x0) (blinfun_apply invf i) = i" by (metis apply1 bij_inv_eq_iff blinfun_bij1 invf) have [simp]: "(range (blinfun_apply invf)) = UNIV" using apply1 surjI by blast let ?f = "invf \ (\x. (f \ (+)x0)x - f x0)" let ?f' = "\x. invf o\<^sub>L (f' (x + x0))" obtain U' V g g' where "open U'" and U': "U' \ (+)(-x0) ` U" "0 \ U'" and "open V" "0 \ V" and hom: "homeomorphism U' V ?f g" and derg: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv (?f'(g y))" and bij: "\y. y \ V \ bij (?f'(g y))" proof (rule inverse_function_theorem_scaled [of "(+)(-x0) ` U" ?f "?f'"]) show ope: "open ((+) (- x0) ` U)" using \open U\ open_translation by blast show "(?f has_derivative blinfun_apply (?f' x)) (at x)" if "x \ (+) (- x0) ` U" for x using that apply clarify apply (rule derf derivative_eq_intros | simp add: blinfun_compose.rep_eq)+ done have YY: "(\x. f' (x + x0)) \u-x0\ f' u" if "f' \u\ f' u" "u \ U" for u using that LIM_offset [where k = x0] by (auto simp: algebra_simps) then have "continuous_on ((+) (- x0) ` U) (\x. f' (x + x0))" using contf \open U\ Lim_at_imp_Lim_at_within by (fastforce simp: continuous_on_def at_within_open_NO_MATCH ope) then show "continuous_on ((+) (- x0) ` U) ?f'" by (intro continuous_intros) simp qed (auto simp: invf \x0 \ U\) show thesis proof let ?U' = "(+)x0 ` U'" let ?V = "((+)(f x0) \ f' x0) ` V" let ?g = "(+)x0 \ g \ invf \ (+)(- f x0)" let ?g' = "\y. inv (blinfun_apply (f' (?g y)))" show oU': "open ?U'" by (simp add: \open U'\ open_translation) show subU: "?U' \ U" using ComplI \U' \ (+) (- x0) ` U\ by auto show "x0 \ ?U'" by (simp add: \0 \ U'\) show "open ?V" using blinfun_bij2 [OF invf] by (metis \open V\ bij_is_surj blinfun.bounded_linear_right bounded_linear_def image_comp open_surjective_linear_image open_translation) show "f x0 \ ?V" using \0 \ V\ image_iff by fastforce show "homeomorphism ?U' ?V f ?g" proof show "continuous_on ?U' f" by (meson subU continuous_on_eq_continuous_at derf has_derivative_continuous oU' subsetD) have "?f ` U' \ V" using hom homeomorphism_image1 by blast then show "f ` ?U' \ ?V" unfolding image_subset_iff by (clarsimp simp: image_def) (metis apply2 add.commute diff_add_cancel) show "?g ` ?V \ ?U'" using hom invf by (auto simp: image_def homeomorphism_def) show "?g (f x) = x" if "x \ ?U'" for x using that hom homeomorphism_apply1 by fastforce have "continuous_on V g" using hom homeomorphism_def by blast then show "continuous_on ?V ?g" by (intro continuous_intros) (auto elim!: continuous_on_subset) have fg: "?f (g x) = x" if "x \ V" for x using hom homeomorphism_apply2 that by blast show "f (?g y) = y" if "y \ ?V" for y using that fg by (simp add: image_iff) (metis apply2 add.commute diff_add_cancel) qed show "(?g has_derivative ?g' y) (at y)" "bij (blinfun_apply (f' (?g y)))" if "y \ ?V" for y proof - have 1: "bij (blinfun_apply invf)" using blinfun_bij1 invf by blast then have 2: "bij (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x by (metis add.commute bij bij_betw_comp_iff2 blinfun_compose.rep_eq that top_greatest) then show "bij (blinfun_apply (f' (?g y)))" using that by auto have "g' x \ blinfun_apply invf = inv (blinfun_apply (f' (x0 + g x)))" if "x \ V" for x using that by (simp add: g' o_inv_distrib blinfun_compose.rep_eq 1 2 add.commute bij_is_inj flip: o_assoc) then show "(?g has_derivative ?g' y) (at y)" using that invf by clarsimp (rule derg derivative_eq_intros | simp flip: id_def)+ qed qed auto qed + +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) + + +subsection\The concept of continuously differentiable\ + +text \ +John Harrison writes as follows: + +``The usual assumption in complex analysis texts is that a path \\\ should be piecewise +continuously differentiable, which ensures that the path integral exists at least for any continuous +f, since all piecewise continuous functions are integrable. However, our notion of validity is +weaker, just piecewise differentiability\ldots{} [namely] continuity plus differentiability except on a +finite set\ldots{} [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to +the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this +can integrate all derivatives.'' + +"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. +Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. + +And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably +difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem +asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ + +definition\<^marker>\tag important\ C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" + (infix "C1'_differentiable'_on" 50) + where + "f C1_differentiable_on S \ + (\D. (\x \ S. (f has_vector_derivative (D x)) (at x)) \ continuous_on S D)" + +lemma C1_differentiable_on_eq: + "f C1_differentiable_on S \ + (\x \ S. f differentiable at x) \ continuous_on S (\x. vector_derivative f (at x))" + (is "?lhs = ?rhs") +proof + assume ?lhs + then show ?rhs + unfolding C1_differentiable_on_def + by (metis (no_types, lifting) continuous_on_eq differentiableI_vector vector_derivative_at) +next + assume ?rhs + then show ?lhs + using C1_differentiable_on_def vector_derivative_works by fastforce +qed + +lemma C1_differentiable_on_subset: + "f C1_differentiable_on T \ S \ T \ f C1_differentiable_on S" + unfolding C1_differentiable_on_def continuous_on_eq_continuous_within + by (blast intro: continuous_within_subset) + +lemma C1_differentiable_compose: + assumes fg: "f C1_differentiable_on S" "g C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" + shows "(g \ f) C1_differentiable_on S" +proof - + have "\x. x \ S \ g \ f differentiable at x" + by (meson C1_differentiable_on_eq assms differentiable_chain_at imageI) + moreover have "continuous_on S (\x. vector_derivative (g \ f) (at x))" + proof (rule continuous_on_eq [of _ "\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))"]) + show "continuous_on S (\x. vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)))" + using fg + apply (clarsimp simp add: C1_differentiable_on_eq) + apply (rule Limits.continuous_on_scaleR, assumption) + by (metis (mono_tags, lifting) continuous_at_imp_continuous_on continuous_on_compose continuous_on_cong differentiable_imp_continuous_within o_def) + show "\x. x \ S \ vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x)) = vector_derivative (g \ f) (at x)" + by (metis (mono_tags, hide_lams) C1_differentiable_on_eq fg imageI vector_derivative_chain_at) + qed + ultimately show ?thesis + by (simp add: C1_differentiable_on_eq) +qed + +lemma C1_diff_imp_diff: "f C1_differentiable_on S \ f differentiable_on S" + by (simp add: C1_differentiable_on_eq differentiable_at_imp_differentiable_on) + +lemma C1_differentiable_on_ident [simp, derivative_intros]: "(\x. x) C1_differentiable_on S" + by (auto simp: C1_differentiable_on_eq) + +lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" + by (auto simp: C1_differentiable_on_eq) + +lemma C1_differentiable_on_add [simp, derivative_intros]: + "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x + g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_minus [simp, derivative_intros]: + "f C1_differentiable_on S \ (\x. - f x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_diff [simp, derivative_intros]: + "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x - g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) + +lemma C1_differentiable_on_mult [simp, derivative_intros]: + fixes f g :: "real \ 'a :: real_normed_algebra" + shows "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x * g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq + by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) + +lemma C1_differentiable_on_scaleR [simp, derivative_intros]: + "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x *\<^sub>R g x) C1_differentiable_on S" + unfolding C1_differentiable_on_eq + by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ + + +definition\<^marker>\tag important\ piecewise_C1_differentiable_on + (infixr "piecewise'_C1'_differentiable'_on" 50) + where "f piecewise_C1_differentiable_on i \ + continuous_on i f \ + (\S. finite S \ (f C1_differentiable_on (i - S)))" + +lemma C1_differentiable_imp_piecewise: + "f C1_differentiable_on S \ f piecewise_C1_differentiable_on S" + by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) + +lemma piecewise_C1_imp_differentiable: + "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" + by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def + C1_differentiable_on_def differentiable_def has_vector_derivative_def + intro: has_derivative_at_withinI) + +lemma piecewise_C1_differentiable_compose: + assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" + shows "(g \ f) piecewise_C1_differentiable_on S" +proof - + have "continuous_on S (\x. g (f x))" + by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) + moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" + proof - + obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" + using fg by (auto simp: piecewise_C1_differentiable_on_def) + obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" + using fg by (auto simp: piecewise_C1_differentiable_on_def) + show ?thesis + proof (intro exI conjI) + show "finite (F \ (\x\G. S \ f-`{x}))" + using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) + show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" + apply (rule C1_differentiable_compose) + apply (blast intro: C1_differentiable_on_subset [OF F]) + apply (blast intro: C1_differentiable_on_subset [OF G]) + by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) + qed + qed + ultimately show ?thesis + by (simp add: piecewise_C1_differentiable_on_def) +qed + +lemma piecewise_C1_differentiable_on_subset: + "f piecewise_C1_differentiable_on S \ T \ S \ f piecewise_C1_differentiable_on T" + by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) + +lemma C1_differentiable_imp_continuous_on: + "f C1_differentiable_on S \ continuous_on S f" + unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within + using differentiable_at_withinI differentiable_imp_continuous_within by blast + +lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" + unfolding C1_differentiable_on_def + by auto + +lemma piecewise_C1_differentiable_affine: + fixes m::real + assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` S)" + shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on S" +proof (cases "m = 0") + case True + then show ?thesis + unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def) +next + case False + have *: "\x. finite (S \ {y. m * y + c = x})" + using False not_finite_existsD by fastforce + show ?thesis + apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) + apply (rule * assms derivative_intros | simp add: False vimage_def)+ + done +qed + +lemma piecewise_C1_differentiable_cases: + fixes c::real + assumes "f piecewise_C1_differentiable_on {a..c}" + "g piecewise_C1_differentiable_on {c..b}" + "a \ c" "c \ b" "f c = g c" + shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" +proof - + obtain S T where st: "f C1_differentiable_on ({a..c} - S)" + "g C1_differentiable_on ({c..b} - T)" + "finite S" "finite T" + using assms + by (force simp: piecewise_C1_differentiable_on_def) + then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" + using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], + OF closed_real_atLeastAtMost [of c b], + of f g "\x. x\c"] assms + by (force simp: ivl_disj_un_two_touch) + { fix x + assume x: "x \ {a..b} - insert c (S \ T)" + have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") + proof (cases x c rule: le_cases) + case le show ?diff_fg + apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) + using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) + next + case ge show ?diff_fg + apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) + using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) + qed + } + then have "(\x \ {a..b} - insert c (S \ T). (\x. if x \ c then f x else g x) differentiable at x)" + by auto + moreover + { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" + and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" + have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + proof - + have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" + if "a < x" "x < c" "x \ S" for x + proof - + have f: "f differentiable at x" + by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) + show ?thesis + using that + apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) + apply (auto simp: dist_norm vector_derivative_works [symmetric] f) + done + qed + then show ?thesis + by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) + qed + moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + proof - + have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" + if "c < x" "x < b" "x \ T" for x + proof - + have g: "g differentiable at x" + by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) + show ?thesis + using that + apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) + apply (auto simp: dist_norm vector_derivative_works [symmetric] g) + done + qed + then show ?thesis + by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) + qed + ultimately have "continuous_on ({a<.. T)) + (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + by (rule continuous_on_subset [OF continuous_on_open_Un], auto) + } note * = this + have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" + using st + by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) + ultimately have "\S. finite S \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - S)" + apply (rule_tac x="{a,b,c} \ S \ T" in exI) + using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) + with cab show ?thesis + by (simp add: piecewise_C1_differentiable_on_def) +qed + +lemma piecewise_C1_differentiable_neg: + "f piecewise_C1_differentiable_on S \ (\x. -(f x)) piecewise_C1_differentiable_on S" + unfolding piecewise_C1_differentiable_on_def + by (auto intro!: continuous_on_minus C1_differentiable_on_minus) + +lemma piecewise_C1_differentiable_add: + assumes "f piecewise_C1_differentiable_on i" + "g piecewise_C1_differentiable_on i" + shows "(\x. f x + g x) piecewise_C1_differentiable_on i" +proof - + obtain S t where st: "finite S" "finite t" + "f C1_differentiable_on (i-S)" + "g C1_differentiable_on (i-t)" + using assms by (auto simp: piecewise_C1_differentiable_on_def) + then have "finite (S \ t) \ (\x. f x + g x) C1_differentiable_on i - (S \ t)" + by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) + moreover have "continuous_on i f" "continuous_on i g" + using assms piecewise_C1_differentiable_on_def by auto + ultimately show ?thesis + by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) +qed + +lemma piecewise_C1_differentiable_diff: + "\f piecewise_C1_differentiable_on S; g piecewise_C1_differentiable_on S\ + \ (\x. f x - g x) piecewise_C1_differentiable_on S" + unfolding diff_conv_add_uminus + by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) + end diff --git a/src/HOL/Analysis/FPS_Convergence.thy b/src/HOL/Analysis/FPS_Convergence.thy --- a/src/HOL/Analysis/FPS_Convergence.thy +++ b/src/HOL/Analysis/FPS_Convergence.thy @@ -1,1311 +1,1056 @@ (* Title: HOL/Analysis/FPS_Convergence.thy Author: Manuel Eberl, TU München Connection of formal power series and actual convergent power series on Banach spaces (most notably the complex numbers). *) section \Convergence of Formal Power Series\ theory FPS_Convergence imports - Conformal_Mappings Generalised_Binomial_Theorem "HOL-Computational_Algebra.Formal_Power_Series" begin +text \ + In this theory, we will connect formal power series (which are algebraic objects) with analytic + functions. This will become more important in complex analysis, and indeed some of the less + trivial results will only be proven there. +\ + subsection\<^marker>\tag unimportant\ \Balls with extended real radius\ +(* TODO: This should probably go somewhere else *) + text \ The following is a variant of \<^const>\ball\ that also allows an infinite radius. \ definition eball :: "'a :: metric_space \ ereal \ 'a set" where "eball z r = {z'. ereal (dist z z') < r}" lemma in_eball_iff [simp]: "z \ eball z0 r \ ereal (dist z0 z) < r" by (simp add: eball_def) lemma eball_ereal [simp]: "eball z (ereal r) = ball z r" by auto lemma eball_inf [simp]: "eball z \ = UNIV" by auto lemma eball_empty [simp]: "r \ 0 \ eball z r = {}" proof safe fix x assume "r \ 0" "x \ eball z r" hence "dist z x < r" by simp also have "\ \ ereal 0" using \r \ 0\ by (simp add: zero_ereal_def) finally show "x \ {}" by simp qed lemma eball_conv_UNION_balls: "eball z r = (\r'\{r'. ereal r' < r}. ball z r')" by (cases r) (use dense gt_ex in force)+ lemma eball_mono: "r \ r' \ eball z r \ eball z r'" by auto lemma ball_eball_mono: "ereal r \ r' \ ball z r \ eball z r'" using eball_mono[of "ereal r" r'] by simp lemma open_eball [simp, intro]: "open (eball z r)" by (cases r) auto subsection \Basic properties of convergent power series\ definition\<^marker>\tag important\ fps_conv_radius :: "'a :: {banach, real_normed_div_algebra} fps \ ereal" where "fps_conv_radius f = conv_radius (fps_nth f)" definition\<^marker>\tag important\ eval_fps :: "'a :: {banach, real_normed_div_algebra} fps \ 'a \ 'a" where "eval_fps f z = (\n. fps_nth f n * z ^ n)" -definition\<^marker>\tag important\ fps_expansion :: "(complex \ complex) \ complex \ complex fps" where - "fps_expansion f z0 = Abs_fps (\n. (deriv ^^ n) f z0 / fact n)" - lemma norm_summable_fps: fixes f :: "'a :: {banach, real_normed_div_algebra} fps" shows "norm z < fps_conv_radius f \ summable (\n. norm (fps_nth f n * z ^ n))" by (rule abs_summable_in_conv_radius) (simp_all add: fps_conv_radius_def) lemma summable_fps: fixes f :: "'a :: {banach, real_normed_div_algebra} fps" shows "norm z < fps_conv_radius f \ summable (\n. fps_nth f n * z ^ n)" by (rule summable_in_conv_radius) (simp_all add: fps_conv_radius_def) theorem sums_eval_fps: fixes f :: "'a :: {banach, real_normed_div_algebra} fps" assumes "norm z < fps_conv_radius f" shows "(\n. fps_nth f n * z ^ n) sums eval_fps f z" using assms unfolding eval_fps_def fps_conv_radius_def by (intro summable_sums summable_in_conv_radius) simp_all -lemma - fixes r :: ereal - assumes "f holomorphic_on eball z0 r" - shows conv_radius_fps_expansion: "fps_conv_radius (fps_expansion f z0) \ r" - and eval_fps_expansion: "\z. z \ eball z0 r \ eval_fps (fps_expansion f z0) (z - z0) = f z" - and eval_fps_expansion': "\z. norm z < r \ eval_fps (fps_expansion f z0) z = f (z0 + z)" -proof - - have "(\n. fps_nth (fps_expansion f z0) n * (z - z0) ^ n) sums f z" - if "z \ ball z0 r'" "ereal r' < r" for z r' - proof - - from that(2) have "ereal r' \ r" by simp - from assms(1) and this have "f holomorphic_on ball z0 r'" - by (rule holomorphic_on_subset[OF _ ball_eball_mono]) - from holomorphic_power_series [OF this that(1)] - show ?thesis by (simp add: fps_expansion_def) - qed - hence *: "(\n. fps_nth (fps_expansion f z0) n * (z - z0) ^ n) sums f z" - if "z \ eball z0 r" for z - using that by (subst (asm) eball_conv_UNION_balls) blast - show "fps_conv_radius (fps_expansion f z0) \ r" unfolding fps_conv_radius_def - proof (rule conv_radius_geI_ex) - fix r' :: real assume r': "r' > 0" "ereal r' < r" - thus "\z. norm z = r' \ summable (\n. fps_nth (fps_expansion f z0) n * z ^ n)" - using *[of "z0 + of_real r'"] - by (intro exI[of _ "of_real r'"]) (auto simp: summable_def dist_norm) - qed - show "eval_fps (fps_expansion f z0) (z - z0) = f z" if "z \ eball z0 r" for z - using *[OF that] by (simp add: eval_fps_def sums_iff) - show "eval_fps (fps_expansion f z0) z = f (z0 + z)" if "ereal (norm z) < r" for z - using *[of "z0 + z"] and that by (simp add: eval_fps_def sums_iff dist_norm) -qed - lemma continuous_on_eval_fps: fixes f :: "'a :: {banach, real_normed_div_algebra} fps" shows "continuous_on (eball 0 (fps_conv_radius f)) (eval_fps f)" proof (subst continuous_on_eq_continuous_at [OF open_eball], safe) fix x :: 'a assume x: "x \ eball 0 (fps_conv_radius f)" define r where "r = (if fps_conv_radius f = \ then norm x + 1 else (norm x + real_of_ereal (fps_conv_radius f)) / 2)" have r: "norm x < r \ ereal r < fps_conv_radius f" using x by (cases "fps_conv_radius f") (auto simp: r_def eball_def split: if_splits) have "continuous_on (cball 0 r) (\x. \i. fps_nth f i * (x - 0) ^ i)" by (rule powser_continuous_suminf) (insert r, auto simp: fps_conv_radius_def) hence "continuous_on (cball 0 r) (eval_fps f)" by (simp add: eval_fps_def) thus "isCont (eval_fps f) x" by (rule continuous_on_interior) (use r in auto) qed lemma continuous_on_eval_fps' [continuous_intros]: assumes "continuous_on A g" assumes "g ` A \ eball 0 (fps_conv_radius f)" shows "continuous_on A (\x. eval_fps f (g x))" using continuous_on_compose2[OF continuous_on_eval_fps assms] . lemma has_field_derivative_powser: fixes z :: "'a :: {banach, real_normed_field}" assumes "ereal (norm z) < conv_radius f" shows "((\z. \n. f n * z ^ n) has_field_derivative (\n. diffs f n * z ^ n)) (at z within A)" proof - define K where "K = (if conv_radius f = \ then norm z + 1 else (norm z + real_of_ereal (conv_radius f)) / 2)" have K: "norm z < K \ ereal K < conv_radius f" using assms by (cases "conv_radius f") (auto simp: K_def) have "0 \ norm z" by simp also from K have "\ < K" by simp finally have K_pos: "K > 0" by simp have "summable (\n. f n * of_real K ^ n)" using K and K_pos by (intro summable_in_conv_radius) auto moreover from K and K_pos have "norm z < norm (of_real K :: 'a)" by auto ultimately show ?thesis by (rule has_field_derivative_at_within [OF termdiffs_strong]) qed lemma has_field_derivative_eval_fps: fixes z :: "'a :: {banach, real_normed_field}" assumes "norm z < fps_conv_radius f" shows "(eval_fps f has_field_derivative eval_fps (fps_deriv f) z) (at z within A)" proof - have "(eval_fps f has_field_derivative eval_fps (Abs_fps (diffs (fps_nth f))) z) (at z within A)" using assms unfolding eval_fps_def fps_nth_Abs_fps fps_conv_radius_def by (intro has_field_derivative_powser) auto also have "Abs_fps (diffs (fps_nth f)) = fps_deriv f" by (simp add: fps_eq_iff diffs_def) finally show ?thesis . qed lemma holomorphic_on_eval_fps [holomorphic_intros]: fixes z :: "'a :: {banach, real_normed_field}" assumes "A \ eball 0 (fps_conv_radius f)" shows "eval_fps f holomorphic_on A" proof (rule holomorphic_on_subset [OF _ assms]) show "eval_fps f holomorphic_on eball 0 (fps_conv_radius f)" proof (subst holomorphic_on_open [OF open_eball], safe, goal_cases) case (1 x) thus ?case by (intro exI[of _ "eval_fps (fps_deriv f) x"]) (auto intro: has_field_derivative_eval_fps) qed qed lemma analytic_on_eval_fps: fixes z :: "'a :: {banach, real_normed_field}" assumes "A \ eball 0 (fps_conv_radius f)" shows "eval_fps f analytic_on A" proof (rule analytic_on_subset [OF _ assms]) show "eval_fps f analytic_on eball 0 (fps_conv_radius f)" using holomorphic_on_eval_fps[of "eball 0 (fps_conv_radius f)"] by (subst analytic_on_open) auto qed lemma continuous_eval_fps [continuous_intros]: fixes z :: "'a::{real_normed_field,banach}" assumes "norm z < fps_conv_radius F" shows "continuous (at z within A) (eval_fps F)" proof - from ereal_dense2[OF assms] obtain K :: real where K: "norm z < K" "K < fps_conv_radius F" by auto have "0 \ norm z" by simp also have "norm z < K" by fact finally have "K > 0" . from K and \K > 0\ have "summable (\n. fps_nth F n * of_real K ^ n)" by (intro summable_fps) auto from this have "isCont (eval_fps F) z" unfolding eval_fps_def by (rule isCont_powser) (use K in auto) thus "continuous (at z within A) (eval_fps F)" by (simp add: continuous_at_imp_continuous_within) qed subsection\<^marker>\tag unimportant\ \Lower bounds on radius of convergence\ lemma fps_conv_radius_deriv: fixes f :: "'a :: {banach, real_normed_field} fps" shows "fps_conv_radius (fps_deriv f) \ fps_conv_radius f" unfolding fps_conv_radius_def proof (rule conv_radius_geI_ex) fix r :: real assume r: "r > 0" "ereal r < conv_radius (fps_nth f)" define K where "K = (if conv_radius (fps_nth f) = \ then r + 1 else (real_of_ereal (conv_radius (fps_nth f)) + r) / 2)" have K: "r < K \ ereal K < conv_radius (fps_nth f)" using r by (cases "conv_radius (fps_nth f)") (auto simp: K_def) have "summable (\n. diffs (fps_nth f) n * of_real r ^ n)" proof (rule termdiff_converges) fix x :: 'a assume "norm x < K" hence "ereal (norm x) < ereal K" by simp also have "\ < conv_radius (fps_nth f)" using K by simp finally show "summable (\n. fps_nth f n * x ^ n)" by (intro summable_in_conv_radius) auto qed (insert K r, auto) also have "\ = (\n. fps_nth (fps_deriv f) n * of_real r ^ n)" by (simp add: fps_deriv_def diffs_def) finally show "\z::'a. norm z = r \ summable (\n. fps_nth (fps_deriv f) n * z ^ n)" using r by (intro exI[of _ "of_real r"]) auto qed lemma eval_fps_at_0: "eval_fps f 0 = fps_nth f 0" by (simp add: eval_fps_def) lemma fps_conv_radius_norm [simp]: "fps_conv_radius (Abs_fps (\n. norm (fps_nth f n))) = fps_conv_radius f" by (simp add: fps_conv_radius_def) lemma fps_conv_radius_const [simp]: "fps_conv_radius (fps_const c) = \" proof - have "fps_conv_radius (fps_const c) = conv_radius (\_. 0 :: 'a)" unfolding fps_conv_radius_def by (intro conv_radius_cong eventually_mono[OF eventually_gt_at_top[of 0]]) auto thus ?thesis by simp qed lemma fps_conv_radius_0 [simp]: "fps_conv_radius 0 = \" by (simp only: fps_const_0_eq_0 [symmetric] fps_conv_radius_const) lemma fps_conv_radius_1 [simp]: "fps_conv_radius 1 = \" by (simp only: fps_const_1_eq_1 [symmetric] fps_conv_radius_const) lemma fps_conv_radius_numeral [simp]: "fps_conv_radius (numeral n) = \" by (simp add: numeral_fps_const) lemma fps_conv_radius_fps_X_power [simp]: "fps_conv_radius (fps_X ^ n) = \" proof - have "fps_conv_radius (fps_X ^ n :: 'a fps) = conv_radius (\_. 0 :: 'a)" unfolding fps_conv_radius_def by (intro conv_radius_cong eventually_mono[OF eventually_gt_at_top[of n]]) (auto simp: fps_X_power_iff) thus ?thesis by simp qed lemma fps_conv_radius_fps_X [simp]: "fps_conv_radius fps_X = \" using fps_conv_radius_fps_X_power[of 1] by (simp only: power_one_right) lemma fps_conv_radius_shift [simp]: "fps_conv_radius (fps_shift n f) = fps_conv_radius f" by (simp add: fps_conv_radius_def fps_shift_def conv_radius_shift) lemma fps_conv_radius_cmult_left: "c \ 0 \ fps_conv_radius (fps_const c * f) = fps_conv_radius f" unfolding fps_conv_radius_def by (simp add: conv_radius_cmult_left) lemma fps_conv_radius_cmult_right: "c \ 0 \ fps_conv_radius (f * fps_const c) = fps_conv_radius f" unfolding fps_conv_radius_def by (simp add: conv_radius_cmult_right) lemma fps_conv_radius_uminus [simp]: "fps_conv_radius (-f) = fps_conv_radius f" using fps_conv_radius_cmult_left[of "-1" f] by (simp flip: fps_const_neg) lemma fps_conv_radius_add: "fps_conv_radius (f + g) \ min (fps_conv_radius f) (fps_conv_radius g)" unfolding fps_conv_radius_def using conv_radius_add_ge[of "fps_nth f" "fps_nth g"] by simp lemma fps_conv_radius_diff: "fps_conv_radius (f - g) \ min (fps_conv_radius f) (fps_conv_radius g)" using fps_conv_radius_add[of f "-g"] by simp lemma fps_conv_radius_mult: "fps_conv_radius (f * g) \ min (fps_conv_radius f) (fps_conv_radius g)" using conv_radius_mult_ge[of "fps_nth f" "fps_nth g"] by (simp add: fps_mult_nth fps_conv_radius_def atLeast0AtMost) lemma fps_conv_radius_power: "fps_conv_radius (f ^ n) \ fps_conv_radius f" proof (induction n) case (Suc n) hence "fps_conv_radius f \ min (fps_conv_radius f) (fps_conv_radius (f ^ n))" by simp also have "\ \ fps_conv_radius (f * f ^ n)" by (rule fps_conv_radius_mult) finally show ?case by simp qed simp_all context begin lemma natfun_inverse_bound: fixes f :: "'a :: {real_normed_field} fps" assumes "fps_nth f 0 = 1" and "\ > 0" and summable: "summable (\n. norm (fps_nth f (Suc n)) * \ ^ Suc n)" and le: "(\n. norm (fps_nth f (Suc n)) * \ ^ Suc n) \ 1" shows "norm (natfun_inverse f n) \ inverse (\ ^ n)" proof (induction n rule: less_induct) case (less m) show ?case proof (cases m) case 0 thus ?thesis using assms by (simp add: field_split_simps norm_inverse norm_divide) next case [simp]: (Suc n) have "norm (natfun_inverse f (Suc n)) = norm (\i = Suc 0..Suc n. fps_nth f i * natfun_inverse f (Suc n - i))" (is "_ = norm ?S") using assms by (simp add: field_simps norm_mult norm_divide del: sum.cl_ivl_Suc) also have "norm ?S \ (\i = Suc 0..Suc n. norm (fps_nth f i * natfun_inverse f (Suc n - i)))" by (rule norm_sum) also have "\ \ (\i = Suc 0..Suc n. norm (fps_nth f i) / \ ^ (Suc n - i))" proof (intro sum_mono, goal_cases) case (1 i) have "norm (fps_nth f i * natfun_inverse f (Suc n - i)) = norm (fps_nth f i) * norm (natfun_inverse f (Suc n - i))" by (simp add: norm_mult) also have "\ \ norm (fps_nth f i) * inverse (\ ^ (Suc n - i))" using 1 by (intro mult_left_mono less.IH) auto also have "\ = norm (fps_nth f i) / \ ^ (Suc n - i)" by (simp add: field_split_simps) finally show ?case . qed also have "\ = (\i = Suc 0..Suc n. norm (fps_nth f i) * \ ^ i) / \ ^ Suc n" by (subst sum_divide_distrib, rule sum.cong) (insert \\ > 0\, auto simp: field_simps power_diff) also have "(\i = Suc 0..Suc n. norm (fps_nth f i) * \ ^ i) = (\i=0..n. norm (fps_nth f (Suc i)) * \ ^ (Suc i))" by (subst sum.atLeast_Suc_atMost_Suc_shift) simp_all also have "{0..n} = {..i< Suc n. norm (fps_nth f (Suc i)) * \ ^ (Suc i)) \ (\n. norm (fps_nth f (Suc n)) * \ ^ (Suc n))" using \\ > 0\ by (intro sum_le_suminf ballI mult_nonneg_nonneg zero_le_power summable) auto also have "\ \ 1" by fact finally show ?thesis using \\ > 0\ by (simp add: divide_right_mono field_split_simps) qed qed private lemma fps_conv_radius_inverse_pos_aux: fixes f :: "'a :: {banach, real_normed_field} fps" assumes "fps_nth f 0 = 1" "fps_conv_radius f > 0" shows "fps_conv_radius (inverse f) > 0" proof - let ?R = "fps_conv_radius f" define h where "h = Abs_fps (\n. norm (fps_nth f n))" have [simp]: "fps_conv_radius h = ?R" by (simp add: h_def) have "continuous_on (eball 0 (fps_conv_radius h)) (eval_fps h)" by (intro continuous_on_eval_fps) hence *: "open (eval_fps h -` A \ eball 0 ?R)" if "open A" for A using that by (subst (asm) continuous_on_open_vimage) auto have "open (eval_fps h -` {..<2} \ eball 0 ?R)" by (rule *) auto moreover have "0 \ eval_fps h -` {..<2} \ eball 0 ?R" using assms by (auto simp: eball_def zero_ereal_def eval_fps_at_0 h_def) ultimately obtain \ where \: "\ > 0" "ball 0 \ \ eval_fps h -` {..<2} \ eball 0 ?R" by (subst (asm) open_contains_ball_eq) blast+ define \ where "\ = real_of_ereal (min (ereal \ / 2) (?R / 2))" have \: "0 < \ \ \ < \ \ ereal \ < ?R" using \\ > 0\ and assms by (cases ?R) (auto simp: \_def min_def) have summable: "summable (\n. norm (fps_nth f n) * \ ^ n)" using \ by (intro summable_in_conv_radius) (simp_all add: fps_conv_radius_def) hence "(\n. norm (fps_nth f n) * \ ^ n) sums eval_fps h \" by (simp add: eval_fps_def summable_sums h_def) hence "(\n. norm (fps_nth f (Suc n)) * \ ^ Suc n) sums (eval_fps h \ - 1)" by (subst sums_Suc_iff) (auto simp: assms) moreover { from \ have "\ \ ball 0 \" by auto also have "\ \ eval_fps h -` {..<2} \ eball 0 ?R" by fact finally have "eval_fps h \ < 2" by simp } ultimately have le: "(\n. norm (fps_nth f (Suc n)) * \ ^ Suc n) \ 1" by (simp add: sums_iff) from summable have summable: "summable (\n. norm (fps_nth f (Suc n)) * \ ^ Suc n)" by (subst summable_Suc_iff) have "0 < \" using \ by blast also have "\ = inverse (limsup (\n. ereal (inverse \)))" using \ by (subst Limsup_const) auto also have "\ \ conv_radius (natfun_inverse f)" unfolding conv_radius_def proof (intro ereal_inverse_antimono Limsup_mono eventually_mono[OF eventually_gt_at_top[of 0]]) fix n :: nat assume n: "n > 0" have "root n (norm (natfun_inverse f n)) \ root n (inverse (\ ^ n))" using n assms \ le summable by (intro real_root_le_mono natfun_inverse_bound) auto also have "\ = inverse \" using n \ by (simp add: power_inverse [symmetric] real_root_pos2) finally show "ereal (inverse \) \ ereal (root n (norm (natfun_inverse f n)))" by (subst ereal_less_eq) next have "0 = limsup (\n. 0::ereal)" by (rule Limsup_const [symmetric]) auto also have "\ \ limsup (\n. ereal (root n (norm (natfun_inverse f n))))" by (intro Limsup_mono) (auto simp: real_root_ge_zero) finally show "0 \ \" by simp qed also have "\ = fps_conv_radius (inverse f)" using assms by (simp add: fps_conv_radius_def fps_inverse_def) finally show ?thesis by (simp add: zero_ereal_def) qed lemma fps_conv_radius_inverse_pos: fixes f :: "'a :: {banach, real_normed_field} fps" assumes "fps_nth f 0 \ 0" and "fps_conv_radius f > 0" shows "fps_conv_radius (inverse f) > 0" proof - let ?c = "fps_nth f 0" have "fps_conv_radius (inverse f) = fps_conv_radius (fps_const ?c * inverse f)" using assms by (subst fps_conv_radius_cmult_left) auto also have "fps_const ?c * inverse f = inverse (fps_const (inverse ?c) * f)" using assms by (simp add: fps_inverse_mult fps_const_inverse) also have "fps_conv_radius \ > 0" using assms by (intro fps_conv_radius_inverse_pos_aux) (auto simp: fps_conv_radius_cmult_left) finally show ?thesis . qed end lemma fps_conv_radius_exp [simp]: fixes c :: "'a :: {banach, real_normed_field}" shows "fps_conv_radius (fps_exp c) = \" unfolding fps_conv_radius_def proof (rule conv_radius_inftyI'') fix z :: 'a have "(\n. norm (c * z) ^ n /\<^sub>R fact n) sums exp (norm (c * z))" by (rule exp_converges) also have "(\n. norm (c * z) ^ n /\<^sub>R fact n) = (\n. norm (fps_nth (fps_exp c) n * z ^ n))" by (rule ext) (simp add: norm_divide norm_mult norm_power field_split_simps power_mult_distrib) finally have "summable \" by (simp add: sums_iff) thus "summable (\n. fps_nth (fps_exp c) n * z ^ n)" by (rule summable_norm_cancel) qed subsection \Evaluating power series\ theorem eval_fps_deriv: assumes "norm z < fps_conv_radius f" shows "eval_fps (fps_deriv f) z = deriv (eval_fps f) z" by (intro DERIV_imp_deriv [symmetric] has_field_derivative_eval_fps assms) theorem fps_nth_conv_deriv: fixes f :: "complex fps" assumes "fps_conv_radius f > 0" shows "fps_nth f n = (deriv ^^ n) (eval_fps f) 0 / fact n" using assms proof (induction n arbitrary: f) case 0 thus ?case by (simp add: eval_fps_def) next case (Suc n f) have "(deriv ^^ Suc n) (eval_fps f) 0 = (deriv ^^ n) (deriv (eval_fps f)) 0" unfolding funpow_Suc_right o_def .. also have "eventually (\z::complex. z \ eball 0 (fps_conv_radius f)) (nhds 0)" using Suc.prems by (intro eventually_nhds_in_open) (auto simp: zero_ereal_def) hence "eventually (\z. deriv (eval_fps f) z = eval_fps (fps_deriv f) z) (nhds 0)" by eventually_elim (simp add: eval_fps_deriv) hence "(deriv ^^ n) (deriv (eval_fps f)) 0 = (deriv ^^ n) (eval_fps (fps_deriv f)) 0" by (intro higher_deriv_cong_ev refl) also have "\ / fact n = fps_nth (fps_deriv f) n" using Suc.prems fps_conv_radius_deriv[of f] by (intro Suc.IH [symmetric]) auto also have "\ / of_nat (Suc n) = fps_nth f (Suc n)" by (simp add: fps_deriv_def del: of_nat_Suc) finally show ?case by (simp add: field_split_simps) qed theorem eval_fps_eqD: fixes f g :: "complex fps" assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0" assumes "eventually (\z. eval_fps f z = eval_fps g z) (nhds 0)" shows "f = g" proof (rule fps_ext) fix n :: nat have "fps_nth f n = (deriv ^^ n) (eval_fps f) 0 / fact n" using assms by (intro fps_nth_conv_deriv) also have "(deriv ^^ n) (eval_fps f) 0 = (deriv ^^ n) (eval_fps g) 0" by (intro higher_deriv_cong_ev refl assms) also have "\ / fact n = fps_nth g n" using assms by (intro fps_nth_conv_deriv [symmetric]) finally show "fps_nth f n = fps_nth g n" . qed lemma eval_fps_const [simp]: fixes c :: "'a :: {banach, real_normed_div_algebra}" shows "eval_fps (fps_const c) z = c" proof - have "(\n::nat. if n \ {0} then c else 0) sums (\n\{0::nat}. c)" by (rule sums_If_finite_set) auto also have "?this \ (\n::nat. fps_nth (fps_const c) n * z ^ n) sums (\n\{0::nat}. c)" by (intro sums_cong) auto also have "(\n\{0::nat}. c) = c" by simp finally show ?thesis by (simp add: eval_fps_def sums_iff) qed lemma eval_fps_0 [simp]: "eval_fps (0 :: 'a :: {banach, real_normed_div_algebra} fps) z = 0" by (simp only: fps_const_0_eq_0 [symmetric] eval_fps_const) lemma eval_fps_1 [simp]: "eval_fps (1 :: 'a :: {banach, real_normed_div_algebra} fps) z = 1" by (simp only: fps_const_1_eq_1 [symmetric] eval_fps_const) lemma eval_fps_numeral [simp]: "eval_fps (numeral n :: 'a :: {banach, real_normed_div_algebra} fps) z = numeral n" by (simp only: numeral_fps_const eval_fps_const) lemma eval_fps_X_power [simp]: "eval_fps (fps_X ^ m :: 'a :: {banach, real_normed_div_algebra} fps) z = z ^ m" proof - have "(\n::nat. if n \ {m} then z ^ n else 0 :: 'a) sums (\n\{m::nat}. z ^ n)" by (rule sums_If_finite_set) auto also have "?this \ (\n::nat. fps_nth (fps_X ^ m) n * z ^ n) sums (\n\{m::nat}. z ^ n)" by (intro sums_cong) (auto simp: fps_X_power_iff) also have "(\n\{m::nat}. z ^ n) = z ^ m" by simp finally show ?thesis by (simp add: eval_fps_def sums_iff) qed lemma eval_fps_X [simp]: "eval_fps (fps_X :: 'a :: {banach, real_normed_div_algebra} fps) z = z" using eval_fps_X_power[of 1 z] by (simp only: power_one_right) lemma eval_fps_minus: fixes f :: "'a :: {banach, real_normed_div_algebra} fps" assumes "norm z < fps_conv_radius f" shows "eval_fps (-f) z = -eval_fps f z" using assms unfolding eval_fps_def by (subst suminf_minus [symmetric]) (auto intro!: summable_fps) lemma eval_fps_add: fixes f g :: "'a :: {banach, real_normed_div_algebra} fps" assumes "norm z < fps_conv_radius f" "norm z < fps_conv_radius g" shows "eval_fps (f + g) z = eval_fps f z + eval_fps g z" using assms unfolding eval_fps_def by (subst suminf_add) (auto simp: ring_distribs intro!: summable_fps) lemma eval_fps_diff: fixes f g :: "'a :: {banach, real_normed_div_algebra} fps" assumes "norm z < fps_conv_radius f" "norm z < fps_conv_radius g" shows "eval_fps (f - g) z = eval_fps f z - eval_fps g z" using assms unfolding eval_fps_def by (subst suminf_diff) (auto simp: ring_distribs intro!: summable_fps) lemma eval_fps_mult: fixes f g :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps" assumes "norm z < fps_conv_radius f" "norm z < fps_conv_radius g" shows "eval_fps (f * g) z = eval_fps f z * eval_fps g z" proof - have "eval_fps f z * eval_fps g z = (\k. \i\k. fps_nth f i * fps_nth g (k - i) * (z ^ i * z ^ (k - i)))" unfolding eval_fps_def proof (subst Cauchy_product) show "summable (\k. norm (fps_nth f k * z ^ k))" "summable (\k. norm (fps_nth g k * z ^ k))" by (rule norm_summable_fps assms)+ qed (simp_all add: algebra_simps) also have "(\k. \i\k. fps_nth f i * fps_nth g (k - i) * (z ^ i * z ^ (k - i))) = (\k. \i\k. fps_nth f i * fps_nth g (k - i) * z ^ k)" by (intro ext sum.cong refl) (simp add: power_add [symmetric]) also have "suminf \ = eval_fps (f * g) z" by (simp add: eval_fps_def fps_mult_nth atLeast0AtMost sum_distrib_right) finally show ?thesis .. qed lemma eval_fps_shift: fixes f :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps" assumes "n \ subdegree f" "norm z < fps_conv_radius f" shows "eval_fps (fps_shift n f) z = (if z = 0 then fps_nth f n else eval_fps f z / z ^ n)" proof (cases "z = 0") case False have "eval_fps (fps_shift n f * fps_X ^ n) z = eval_fps (fps_shift n f) z * z ^ n" using assms by (subst eval_fps_mult) simp_all also from assms have "fps_shift n f * fps_X ^ n = f" by (simp add: fps_shift_times_fps_X_power) finally show ?thesis using False by (simp add: field_simps) qed (simp_all add: eval_fps_at_0) lemma eval_fps_exp [simp]: fixes c :: "'a :: {banach, real_normed_field}" shows "eval_fps (fps_exp c) z = exp (c * z)" unfolding eval_fps_def exp_def by (simp add: eval_fps_def exp_def scaleR_conv_of_real field_split_simps power_mult_distrib) -lemma - fixes f :: "complex fps" and r :: ereal - assumes "\z. ereal (norm z) < r \ eval_fps f z \ 0" - shows fps_conv_radius_inverse: "fps_conv_radius (inverse f) \ min r (fps_conv_radius f)" - and eval_fps_inverse: "\z. ereal (norm z) < fps_conv_radius f \ ereal (norm z) < r \ - eval_fps (inverse f) z = inverse (eval_fps f z)" -proof - - define R where "R = min (fps_conv_radius f) r" - have *: "fps_conv_radius (inverse f) \ min r (fps_conv_radius f) \ - (\z\eball 0 (min (fps_conv_radius f) r). eval_fps (inverse f) z = inverse (eval_fps f z))" - proof (cases "min r (fps_conv_radius f) > 0") - case True - define f' where "f' = fps_expansion (\z. inverse (eval_fps f z)) 0" - have holo: "(\z. inverse (eval_fps f z)) holomorphic_on eball 0 (min r (fps_conv_radius f))" - using assms by (intro holomorphic_intros) auto - from holo have radius: "fps_conv_radius f' \ min r (fps_conv_radius f)" - unfolding f'_def by (rule conv_radius_fps_expansion) - have eval_f': "eval_fps f' z = inverse (eval_fps f z)" - if "norm z < fps_conv_radius f" "norm z < r" for z - using that unfolding f'_def by (subst eval_fps_expansion'[OF holo]) auto - - have "f * f' = 1" - proof (rule eval_fps_eqD) - from radius and True have "0 < min (fps_conv_radius f) (fps_conv_radius f')" - by (auto simp: min_def split: if_splits) - also have "\ \ fps_conv_radius (f * f')" by (rule fps_conv_radius_mult) - finally show "\ > 0" . - next - from True have "R > 0" by (auto simp: R_def) - hence "eventually (\z. z \ eball 0 R) (nhds 0)" - by (intro eventually_nhds_in_open) (auto simp: zero_ereal_def) - thus "eventually (\z. eval_fps (f * f') z = eval_fps 1 z) (nhds 0)" - proof eventually_elim - case (elim z) - hence "eval_fps (f * f') z = eval_fps f z * eval_fps f' z" - using radius by (intro eval_fps_mult) - (auto simp: R_def min_def split: if_splits intro: less_trans) - also have "eval_fps f' z = inverse (eval_fps f z)" - using elim by (intro eval_f') (auto simp: R_def) - also from elim have "eval_fps f z \ 0" - by (intro assms) (auto simp: R_def) - hence "eval_fps f z * inverse (eval_fps f z) = eval_fps 1 z" - by simp - finally show "eval_fps (f * f') z = eval_fps 1 z" . - qed - qed simp_all - hence "f' = inverse f" - by (intro fps_inverse_unique [symmetric]) (simp_all add: mult_ac) - with eval_f' and radius show ?thesis by simp - next - case False - hence *: "eball 0 R = {}" - by (intro eball_empty) (auto simp: R_def min_def split: if_splits) - show ?thesis - proof safe - from False have "min r (fps_conv_radius f) \ 0" - by (simp add: min_def) - also have "0 \ fps_conv_radius (inverse f)" - by (simp add: fps_conv_radius_def conv_radius_nonneg) - finally show "min r (fps_conv_radius f) \ \" . - qed (unfold * [unfolded R_def], auto) - qed - - from * show "fps_conv_radius (inverse f) \ min r (fps_conv_radius f)" by blast - from * show "eval_fps (inverse f) z = inverse (eval_fps f z)" - if "ereal (norm z) < fps_conv_radius f" "ereal (norm z) < r" for z - using that by auto -qed - -lemma - fixes f g :: "complex fps" and r :: ereal - defines "R \ Min {r, fps_conv_radius f, fps_conv_radius g}" - assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0" "r > 0" - assumes nz: "\z. z \ eball 0 r \ eval_fps g z \ 0" - shows fps_conv_radius_divide': "fps_conv_radius (f / g) \ R" - and eval_fps_divide': - "ereal (norm z) < R \ eval_fps (f / g) z = eval_fps f z / eval_fps g z" -proof - - from nz[of 0] and \r > 0\ have nz': "fps_nth g 0 \ 0" - by (auto simp: eval_fps_at_0 zero_ereal_def) - have "R \ min r (fps_conv_radius g)" - by (auto simp: R_def intro: min.coboundedI2) - also have "min r (fps_conv_radius g) \ fps_conv_radius (inverse g)" - by (intro fps_conv_radius_inverse assms) (auto simp: zero_ereal_def) - finally have radius: "fps_conv_radius (inverse g) \ R" . - have "R \ min (fps_conv_radius f) (fps_conv_radius (inverse g))" - by (intro radius min.boundedI) (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) - also have "\ \ fps_conv_radius (f * inverse g)" - by (rule fps_conv_radius_mult) - also have "f * inverse g = f / g" - by (intro fps_divide_unit [symmetric] nz') - finally show "fps_conv_radius (f / g) \ R" . - - assume z: "ereal (norm z) < R" - have "eval_fps (f * inverse g) z = eval_fps f z * eval_fps (inverse g) z" - using radius by (intro eval_fps_mult less_le_trans[OF z]) - (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) - also have "eval_fps (inverse g) z = inverse (eval_fps g z)" using \r > 0\ - by (intro eval_fps_inverse[where r = r] less_le_trans[OF z] nz) - (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) - also have "f * inverse g = f / g" by fact - finally show "eval_fps (f / g) z = eval_fps f z / eval_fps g z" by (simp add: field_split_simps) -qed - -lemma - fixes f g :: "complex fps" and r :: ereal - defines "R \ Min {r, fps_conv_radius f, fps_conv_radius g}" - assumes "subdegree g \ subdegree f" - assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0" "r > 0" - assumes "\z. z \ eball 0 r \ z \ 0 \ eval_fps g z \ 0" - shows fps_conv_radius_divide: "fps_conv_radius (f / g) \ R" - and eval_fps_divide: - "ereal (norm z) < R \ c = fps_nth f (subdegree g) / fps_nth g (subdegree g) \ - eval_fps (f / g) z = (if z = 0 then c else eval_fps f z / eval_fps g z)" -proof - - define f' g' where "f' = fps_shift (subdegree g) f" and "g' = fps_shift (subdegree g) g" - have f_eq: "f = f' * fps_X ^ subdegree g" and g_eq: "g = g' * fps_X ^ subdegree g" - unfolding f'_def g'_def by (rule subdegree_decompose' le_refl | fact)+ - have subdegree: "subdegree f' = subdegree f - subdegree g" "subdegree g' = 0" - using assms(2) by (simp_all add: f'_def g'_def) - have [simp]: "fps_conv_radius f' = fps_conv_radius f" "fps_conv_radius g' = fps_conv_radius g" - by (simp_all add: f'_def g'_def) - have [simp]: "fps_nth f' 0 = fps_nth f (subdegree g)" - "fps_nth g' 0 = fps_nth g (subdegree g)" by (simp_all add: f'_def g'_def) - have g_nz: "g \ 0" - proof - - define z :: complex where "z = (if r = \ then 1 else of_real (real_of_ereal r / 2))" - from \r > 0\ have "z \ eball 0 r" - by (cases r) (auto simp: z_def eball_def) - moreover have "z \ 0" using \r > 0\ - by (cases r) (auto simp: z_def) - ultimately have "eval_fps g z \ 0" by (rule assms(6)) - thus "g \ 0" by auto - qed - have fg: "f / g = f' * inverse g'" - by (subst f_eq, subst (2) g_eq) (insert g_nz, simp add: fps_divide_unit) - - have g'_nz: "eval_fps g' z \ 0" if z: "norm z < min r (fps_conv_radius g)" for z - proof (cases "z = 0") - case False - with assms and z have "eval_fps g z \ 0" by auto - also from z have "eval_fps g z = eval_fps g' z * z ^ subdegree g" - by (subst g_eq) (auto simp: eval_fps_mult) - finally show ?thesis by auto - qed (insert \g \ 0\, auto simp: g'_def eval_fps_at_0) - - have "R \ min (min r (fps_conv_radius g)) (fps_conv_radius g')" - by (auto simp: R_def min.coboundedI1 min.coboundedI2) - also have "\ \ fps_conv_radius (inverse g')" - using g'_nz by (rule fps_conv_radius_inverse) - finally have conv_radius_inv: "R \ fps_conv_radius (inverse g')" . - hence "R \ fps_conv_radius (f' * inverse g')" - by (intro order.trans[OF _ fps_conv_radius_mult]) - (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) - thus "fps_conv_radius (f / g) \ R" by (simp add: fg) - - fix z c :: complex assume z: "ereal (norm z) < R" - assume c: "c = fps_nth f (subdegree g) / fps_nth g (subdegree g)" - show "eval_fps (f / g) z = (if z = 0 then c else eval_fps f z / eval_fps g z)" - proof (cases "z = 0") - case False - from z and conv_radius_inv have "ereal (norm z) < fps_conv_radius (inverse g')" - by simp - with z have "eval_fps (f / g) z = eval_fps f' z * eval_fps (inverse g') z" - unfolding fg by (subst eval_fps_mult) (auto simp: R_def) - also have "eval_fps (inverse g') z = inverse (eval_fps g' z)" - using z by (intro eval_fps_inverse[of "min r (fps_conv_radius g')"] g'_nz) (auto simp: R_def) - also have "eval_fps f' z * \ = eval_fps f z / eval_fps g z" - using z False assms(2) by (simp add: f'_def g'_def eval_fps_shift R_def) - finally show ?thesis using False by simp - qed (simp_all add: eval_fps_at_0 fg field_simps c) -qed +text \ + The case of division is more complicated and will therefore not be handled here. + Handling division becomes much more easy using complex analysis, and we will do so once + that is available. +\ -subsection \Power series expansion of complex functions\ +subsection \Power series expansions of analytic functions\ text \ This predicate contains the notion that the given formal power series converges in some disc of positive radius around the origin and is equal to the given complex function there. This relationship is unique in the sense that no complex function can have more than one formal power series to which it expands, and if two holomorphic functions that are holomorphic on a connected open set around the origin and have the same power series expansion, they must be equal on that set. More concrete statements about the radius of convergence can usually be made, but for many purposes, the statment that the series converges to the function in some neighbourhood of the origin is enough, and that can be shown almost fully automatically in most cases, as there are straightforward introduction rules to show this. In particular, when one wants to relate the coefficients of the power series to the values of the derivatives of the function at the origin, or if one wants to approximate the coefficients of the series with the singularities of the function, this predicate is enough. \ definition\<^marker>\tag important\ has_fps_expansion :: "('a :: {banach,real_normed_div_algebra} \ 'a) \ 'a fps \ bool" (infixl "has'_fps'_expansion" 60) where "(f has_fps_expansion F) \ fps_conv_radius F > 0 \ eventually (\z. eval_fps F z = f z) (nhds 0)" named_theorems fps_expansion_intros lemma fps_nth_fps_expansion: fixes f :: "complex \ complex" assumes "f has_fps_expansion F" shows "fps_nth F n = (deriv ^^ n) f 0 / fact n" proof - have "fps_nth F n = (deriv ^^ n) (eval_fps F) 0 / fact n" using assms by (intro fps_nth_conv_deriv) (auto simp: has_fps_expansion_def) also have "(deriv ^^ n) (eval_fps F) 0 = (deriv ^^ n) f 0" using assms by (intro higher_deriv_cong_ev) (auto simp: has_fps_expansion_def) finally show ?thesis . qed -lemma has_fps_expansion_fps_expansion [intro]: - assumes "open A" "0 \ A" "f holomorphic_on A" - shows "f has_fps_expansion fps_expansion f 0" -proof - - from assms(1,2) obtain r where r: "r > 0 " "ball 0 r \ A" - by (auto simp: open_contains_ball) - have holo: "f holomorphic_on eball 0 (ereal r)" - using r(2) and assms(3) by auto - from r(1) have "0 < ereal r" by simp - also have "r \ fps_conv_radius (fps_expansion f 0)" - using holo by (intro conv_radius_fps_expansion) auto - finally have "\ > 0" . - moreover have "eventually (\z. z \ ball 0 r) (nhds 0)" - using r(1) by (intro eventually_nhds_in_open) auto - hence "eventually (\z. eval_fps (fps_expansion f 0) z = f z) (nhds 0)" - by eventually_elim (subst eval_fps_expansion'[OF holo], auto) - ultimately show ?thesis using r(1) by (auto simp: has_fps_expansion_def) -qed - lemma has_fps_expansion_imp_continuous: fixes F :: "'a::{real_normed_field,banach} fps" assumes "f has_fps_expansion F" shows "continuous (at 0 within A) f" proof - from assms have "isCont (eval_fps F) 0" by (intro continuous_eval_fps) (auto simp: has_fps_expansion_def zero_ereal_def) also have "?this \ isCont f 0" using assms by (intro isCont_cong) (auto simp: has_fps_expansion_def) finally have "isCont f 0" . thus "continuous (at 0 within A) f" by (simp add: continuous_at_imp_continuous_within) qed lemma has_fps_expansion_const [simp, intro, fps_expansion_intros]: "(\_. c) has_fps_expansion fps_const c" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_0 [simp, intro, fps_expansion_intros]: "(\_. 0) has_fps_expansion 0" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_1 [simp, intro, fps_expansion_intros]: "(\_. 1) has_fps_expansion 1" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_numeral [simp, intro, fps_expansion_intros]: "(\_. numeral n) has_fps_expansion numeral n" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_fps_X_power [fps_expansion_intros]: "(\x. x ^ n) has_fps_expansion (fps_X ^ n)" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_fps_X [fps_expansion_intros]: "(\x. x) has_fps_expansion fps_X" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_cmult_left [fps_expansion_intros]: fixes c :: "'a :: {banach, real_normed_div_algebra, comm_ring_1}" assumes "f has_fps_expansion F" shows "(\x. c * f x) has_fps_expansion fps_const c * F" proof (cases "c = 0") case False from assms have "eventually (\z. z \ eball 0 (fps_conv_radius F)) (nhds 0)" by (intro eventually_nhds_in_open) (auto simp: has_fps_expansion_def zero_ereal_def) moreover from assms have "eventually (\z. eval_fps F z = f z) (nhds 0)" by (auto simp: has_fps_expansion_def) ultimately have "eventually (\z. eval_fps (fps_const c * F) z = c * f z) (nhds 0)" by eventually_elim (simp_all add: eval_fps_mult) with assms and False show ?thesis by (auto simp: has_fps_expansion_def fps_conv_radius_cmult_left) qed auto lemma has_fps_expansion_cmult_right [fps_expansion_intros]: fixes c :: "'a :: {banach, real_normed_div_algebra, comm_ring_1}" assumes "f has_fps_expansion F" shows "(\x. f x * c) has_fps_expansion F * fps_const c" proof - have "F * fps_const c = fps_const c * F" by (intro fps_ext) (auto simp: mult.commute) with has_fps_expansion_cmult_left [OF assms] show ?thesis by (simp add: mult.commute) qed lemma has_fps_expansion_minus [fps_expansion_intros]: assumes "f has_fps_expansion F" shows "(\x. - f x) has_fps_expansion -F" proof - from assms have "eventually (\x. x \ eball 0 (fps_conv_radius F)) (nhds 0)" by (intro eventually_nhds_in_open) (auto simp: has_fps_expansion_def zero_ereal_def) moreover from assms have "eventually (\x. eval_fps F x = f x) (nhds 0)" by (auto simp: has_fps_expansion_def) ultimately have "eventually (\x. eval_fps (-F) x = -f x) (nhds 0)" by eventually_elim (auto simp: eval_fps_minus) thus ?thesis using assms by (auto simp: has_fps_expansion_def) qed lemma has_fps_expansion_add [fps_expansion_intros]: assumes "f has_fps_expansion F" "g has_fps_expansion G" shows "(\x. f x + g x) has_fps_expansion F + G" proof - from assms have "0 < min (fps_conv_radius F) (fps_conv_radius G)" by (auto simp: has_fps_expansion_def) also have "\ \ fps_conv_radius (F + G)" by (rule fps_conv_radius_add) finally have radius: "\ > 0" . from assms have "eventually (\x. x \ eball 0 (fps_conv_radius F)) (nhds 0)" "eventually (\x. x \ eball 0 (fps_conv_radius G)) (nhds 0)" by (intro eventually_nhds_in_open; force simp: has_fps_expansion_def zero_ereal_def)+ moreover have "eventually (\x. eval_fps F x = f x) (nhds 0)" and "eventually (\x. eval_fps G x = g x) (nhds 0)" using assms by (auto simp: has_fps_expansion_def) ultimately have "eventually (\x. eval_fps (F + G) x = f x + g x) (nhds 0)" by eventually_elim (auto simp: eval_fps_add) with radius show ?thesis by (auto simp: has_fps_expansion_def) qed lemma has_fps_expansion_diff [fps_expansion_intros]: assumes "f has_fps_expansion F" "g has_fps_expansion G" shows "(\x. f x - g x) has_fps_expansion F - G" using has_fps_expansion_add[of f F "\x. - g x" "-G"] assms by (simp add: has_fps_expansion_minus) lemma has_fps_expansion_mult [fps_expansion_intros]: fixes F G :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps" assumes "f has_fps_expansion F" "g has_fps_expansion G" shows "(\x. f x * g x) has_fps_expansion F * G" proof - from assms have "0 < min (fps_conv_radius F) (fps_conv_radius G)" by (auto simp: has_fps_expansion_def) also have "\ \ fps_conv_radius (F * G)" by (rule fps_conv_radius_mult) finally have radius: "\ > 0" . from assms have "eventually (\x. x \ eball 0 (fps_conv_radius F)) (nhds 0)" "eventually (\x. x \ eball 0 (fps_conv_radius G)) (nhds 0)" by (intro eventually_nhds_in_open; force simp: has_fps_expansion_def zero_ereal_def)+ moreover have "eventually (\x. eval_fps F x = f x) (nhds 0)" and "eventually (\x. eval_fps G x = g x) (nhds 0)" using assms by (auto simp: has_fps_expansion_def) ultimately have "eventually (\x. eval_fps (F * G) x = f x * g x) (nhds 0)" by eventually_elim (auto simp: eval_fps_mult) with radius show ?thesis by (auto simp: has_fps_expansion_def) qed lemma has_fps_expansion_inverse [fps_expansion_intros]: fixes F :: "'a :: {banach, real_normed_field} fps" assumes "f has_fps_expansion F" assumes "fps_nth F 0 \ 0" shows "(\x. inverse (f x)) has_fps_expansion inverse F" proof - have radius: "fps_conv_radius (inverse F) > 0" using assms unfolding has_fps_expansion_def by (intro fps_conv_radius_inverse_pos) auto let ?R = "min (fps_conv_radius F) (fps_conv_radius (inverse F))" from assms radius have "eventually (\x. x \ eball 0 (fps_conv_radius F)) (nhds 0)" "eventually (\x. x \ eball 0 (fps_conv_radius (inverse F))) (nhds 0)" by (intro eventually_nhds_in_open; force simp: has_fps_expansion_def zero_ereal_def)+ moreover have "eventually (\z. eval_fps F z = f z) (nhds 0)" using assms by (auto simp: has_fps_expansion_def) ultimately have "eventually (\z. eval_fps (inverse F) z = inverse (f z)) (nhds 0)" proof eventually_elim case (elim z) hence "eval_fps (inverse F * F) z = eval_fps (inverse F) z * f z" by (subst eval_fps_mult) auto also have "eval_fps (inverse F * F) z = 1" using assms by (simp add: inverse_mult_eq_1) finally show ?case by (auto simp: field_split_simps) qed with radius show ?thesis by (auto simp: has_fps_expansion_def) qed lemma has_fps_expansion_exp [fps_expansion_intros]: fixes c :: "'a :: {banach, real_normed_field}" shows "(\x. exp (c * x)) has_fps_expansion fps_exp c" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_exp1 [fps_expansion_intros]: "(\x::'a :: {banach, real_normed_field}. exp x) has_fps_expansion fps_exp 1" using has_fps_expansion_exp[of 1] by simp lemma has_fps_expansion_exp_neg1 [fps_expansion_intros]: "(\x::'a :: {banach, real_normed_field}. exp (-x)) has_fps_expansion fps_exp (-1)" using has_fps_expansion_exp[of "-1"] by simp lemma has_fps_expansion_deriv [fps_expansion_intros]: assumes "f has_fps_expansion F" shows "deriv f has_fps_expansion fps_deriv F" proof - have "eventually (\z. z \ eball 0 (fps_conv_radius F)) (nhds 0)" using assms by (intro eventually_nhds_in_open) (auto simp: has_fps_expansion_def zero_ereal_def) moreover from assms have "eventually (\z. eval_fps F z = f z) (nhds 0)" by (auto simp: has_fps_expansion_def) then obtain s where "open s" "0 \ s" and s: "\w. w \ s \ eval_fps F w = f w" by (auto simp: eventually_nhds) hence "eventually (\w. w \ s) (nhds 0)" by (intro eventually_nhds_in_open) auto ultimately have "eventually (\z. eval_fps (fps_deriv F) z = deriv f z) (nhds 0)" proof eventually_elim case (elim z) hence "eval_fps (fps_deriv F) z = deriv (eval_fps F) z" by (simp add: eval_fps_deriv) also have "eventually (\w. w \ s) (nhds z)" using elim and \open s\ by (intro eventually_nhds_in_open) auto hence "eventually (\w. eval_fps F w = f w) (nhds z)" by eventually_elim (simp add: s) hence "deriv (eval_fps F) z = deriv f z" by (intro deriv_cong_ev refl) finally show ?case . qed with assms and fps_conv_radius_deriv[of F] show ?thesis by (auto simp: has_fps_expansion_def) qed lemma fps_conv_radius_binomial: fixes c :: "'a :: {real_normed_field,banach}" shows "fps_conv_radius (fps_binomial c) = (if c \ \ then \ else 1)" unfolding fps_conv_radius_def by (simp add: conv_radius_gchoose) lemma fps_conv_radius_ln: fixes c :: "'a :: {banach, real_normed_field, field_char_0}" shows "fps_conv_radius (fps_ln c) = (if c = 0 then \ else 1)" proof (cases "c = 0") case False have "conv_radius (\n. 1 / of_nat n :: 'a) = 1" proof (rule conv_radius_ratio_limit_nonzero) show "(\n. norm (1 / of_nat n :: 'a) / norm (1 / of_nat (Suc n) :: 'a)) \ 1" using LIMSEQ_Suc_n_over_n by (simp add: norm_divide del: of_nat_Suc) qed auto also have "conv_radius (\n. 1 / of_nat n :: 'a) = conv_radius (\n. if n = 0 then 0 else (- 1) ^ (n - 1) / of_nat n :: 'a)" by (intro conv_radius_cong eventually_mono[OF eventually_gt_at_top[of 0]]) (simp add: norm_mult norm_divide norm_power) finally show ?thesis using False unfolding fps_ln_def by (subst fps_conv_radius_cmult_left) (simp_all add: fps_conv_radius_def) qed (auto simp: fps_ln_def) lemma fps_conv_radius_ln_nonzero [simp]: assumes "c \ (0 :: 'a :: {banach,real_normed_field,field_char_0})" shows "fps_conv_radius (fps_ln c) = 1" using assms by (simp add: fps_conv_radius_ln) lemma fps_conv_radius_sin [simp]: fixes c :: "'a :: {banach, real_normed_field, field_char_0}" shows "fps_conv_radius (fps_sin c) = \" proof (cases "c = 0") case False have "\ = conv_radius (\n. of_real (sin_coeff n) :: 'a)" proof (rule sym, rule conv_radius_inftyI'', rule summable_norm_cancel, goal_cases) case (1 z) show ?case using summable_norm_sin[of z] by (simp add: norm_mult) qed also have "\ / norm c = conv_radius (\n. c ^ n * of_real (sin_coeff n) :: 'a)" using False by (subst conv_radius_mult_power) auto also have "\ = fps_conv_radius (fps_sin c)" unfolding fps_conv_radius_def by (rule conv_radius_cong_weak) (auto simp add: fps_sin_def sin_coeff_def) finally show ?thesis by simp qed simp_all lemma fps_conv_radius_cos [simp]: fixes c :: "'a :: {banach, real_normed_field, field_char_0}" shows "fps_conv_radius (fps_cos c) = \" proof (cases "c = 0") case False have "\ = conv_radius (\n. of_real (cos_coeff n) :: 'a)" proof (rule sym, rule conv_radius_inftyI'', rule summable_norm_cancel, goal_cases) case (1 z) show ?case using summable_norm_cos[of z] by (simp add: norm_mult) qed also have "\ / norm c = conv_radius (\n. c ^ n * of_real (cos_coeff n) :: 'a)" using False by (subst conv_radius_mult_power) auto also have "\ = fps_conv_radius (fps_cos c)" unfolding fps_conv_radius_def by (rule conv_radius_cong_weak) (auto simp add: fps_cos_def cos_coeff_def) finally show ?thesis by simp qed simp_all lemma eval_fps_sin [simp]: fixes z :: "'a :: {banach, real_normed_field, field_char_0}" shows "eval_fps (fps_sin c) z = sin (c * z)" proof - have "(\n. sin_coeff n *\<^sub>R (c * z) ^ n) sums sin (c * z)" by (rule sin_converges) also have "(\n. sin_coeff n *\<^sub>R (c * z) ^ n) = (\n. fps_nth (fps_sin c) n * z ^ n)" by (rule ext) (auto simp: sin_coeff_def fps_sin_def power_mult_distrib scaleR_conv_of_real) finally show ?thesis by (simp add: sums_iff eval_fps_def) qed lemma eval_fps_cos [simp]: fixes z :: "'a :: {banach, real_normed_field, field_char_0}" shows "eval_fps (fps_cos c) z = cos (c * z)" proof - have "(\n. cos_coeff n *\<^sub>R (c * z) ^ n) sums cos (c * z)" by (rule cos_converges) also have "(\n. cos_coeff n *\<^sub>R (c * z) ^ n) = (\n. fps_nth (fps_cos c) n * z ^ n)" by (rule ext) (auto simp: cos_coeff_def fps_cos_def power_mult_distrib scaleR_conv_of_real) finally show ?thesis by (simp add: sums_iff eval_fps_def) qed lemma cos_eq_zero_imp_norm_ge: assumes "cos (z :: complex) = 0" shows "norm z \ pi / 2" proof - from assms obtain n where "z = complex_of_real ((of_int n + 1 / 2) * pi)" by (auto simp: cos_eq_0 algebra_simps) also have "norm \ = \real_of_int n + 1 / 2\ * pi" by (subst norm_of_real) (simp_all add: abs_mult) also have "real_of_int n + 1 / 2 = of_int (2 * n + 1) / 2" by simp also have "\\\ = of_int \2 * n + 1\ / 2" by (subst abs_divide) simp_all also have "\ * pi = of_int \2 * n + 1\ * (pi / 2)" by simp also have "\ \ of_int 1 * (pi / 2)" by (intro mult_right_mono, subst of_int_le_iff) (auto simp: abs_if) finally show ?thesis by simp qed -lemma fps_conv_radius_tan: - fixes c :: complex - assumes "c \ 0" - shows "fps_conv_radius (fps_tan c) \ pi / (2 * norm c)" -proof - - have "fps_conv_radius (fps_tan c) \ - Min {pi / (2 * norm c), fps_conv_radius (fps_sin c), fps_conv_radius (fps_cos c)}" - unfolding fps_tan_def - proof (rule fps_conv_radius_divide) - fix z :: complex assume "z \ eball 0 (pi / (2 * norm c))" - with cos_eq_zero_imp_norm_ge[of "c*z"] assms - show "eval_fps (fps_cos c) z \ 0" by (auto simp: norm_mult field_simps) - qed (insert assms, auto) - thus ?thesis by (simp add: min_def) -qed -lemma eval_fps_tan: - fixes c :: complex - assumes "norm z < pi / (2 * norm c)" - shows "eval_fps (fps_tan c) z = tan (c * z)" -proof (cases "c = 0") - case False - show ?thesis unfolding fps_tan_def - proof (subst eval_fps_divide'[where r = "pi / (2 * norm c)"]) - fix z :: complex assume "z \ eball 0 (pi / (2 * norm c))" - with cos_eq_zero_imp_norm_ge[of "c*z"] assms - show "eval_fps (fps_cos c) z \ 0" using False by (auto simp: norm_mult field_simps) - qed (insert False assms, auto simp: field_simps tan_def) -qed simp_all lemma eval_fps_binomial: fixes c :: complex assumes "norm z < 1" shows "eval_fps (fps_binomial c) z = (1 + z) powr c" using gen_binomial_complex[OF assms] by (simp add: sums_iff eval_fps_def) lemma has_fps_expansion_binomial_complex [fps_expansion_intros]: fixes c :: complex shows "(\x. (1 + x) powr c) has_fps_expansion fps_binomial c" proof - have *: "eventually (\z::complex. z \ eball 0 1) (nhds 0)" by (intro eventually_nhds_in_open) auto thus ?thesis by (auto simp: has_fps_expansion_def eval_fps_binomial fps_conv_radius_binomial intro!: eventually_mono [OF *]) qed lemma has_fps_expansion_sin [fps_expansion_intros]: fixes c :: "'a :: {banach, real_normed_field, field_char_0}" shows "(\x. sin (c * x)) has_fps_expansion fps_sin c" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_sin' [fps_expansion_intros]: "(\x::'a :: {banach, real_normed_field}. sin x) has_fps_expansion fps_sin 1" using has_fps_expansion_sin[of 1] by simp lemma has_fps_expansion_cos [fps_expansion_intros]: fixes c :: "'a :: {banach, real_normed_field, field_char_0}" shows "(\x. cos (c * x)) has_fps_expansion fps_cos c" by (auto simp: has_fps_expansion_def) lemma has_fps_expansion_cos' [fps_expansion_intros]: "(\x::'a :: {banach, real_normed_field}. cos x) has_fps_expansion fps_cos 1" using has_fps_expansion_cos[of 1] by simp lemma has_fps_expansion_shift [fps_expansion_intros]: fixes F :: "'a :: {banach, real_normed_field} fps" assumes "f has_fps_expansion F" and "n \ subdegree F" assumes "c = fps_nth F n" shows "(\x. if x = 0 then c else f x / x ^ n) has_fps_expansion (fps_shift n F)" proof - have "eventually (\x. x \ eball 0 (fps_conv_radius F)) (nhds 0)" using assms by (intro eventually_nhds_in_open) (auto simp: has_fps_expansion_def zero_ereal_def) moreover have "eventually (\x. eval_fps F x = f x) (nhds 0)" using assms by (auto simp: has_fps_expansion_def) ultimately have "eventually (\x. eval_fps (fps_shift n F) x = (if x = 0 then c else f x / x ^ n)) (nhds 0)" by eventually_elim (auto simp: eval_fps_shift assms) with assms show ?thesis by (auto simp: has_fps_expansion_def) qed lemma has_fps_expansion_divide [fps_expansion_intros]: fixes F G :: "'a :: {banach, real_normed_field} fps" assumes "f has_fps_expansion F" and "g has_fps_expansion G" and "subdegree G \ subdegree F" "G \ 0" "c = fps_nth F (subdegree G) / fps_nth G (subdegree G)" shows "(\x. if x = 0 then c else f x / g x) has_fps_expansion (F / G)" proof - define n where "n = subdegree G" define F' and G' where "F' = fps_shift n F" and "G' = fps_shift n G" have "F = F' * fps_X ^ n" "G = G' * fps_X ^ n" unfolding F'_def G'_def n_def by (rule fps_shift_times_fps_X_power [symmetric] le_refl | fact)+ moreover from assms have "fps_nth G' 0 \ 0" by (simp add: G'_def n_def) ultimately have FG: "F / G = F' * inverse G'" by (simp add: fps_divide_unit) have "(\x. (if x = 0 then fps_nth F n else f x / x ^ n) * inverse (if x = 0 then fps_nth G n else g x / x ^ n)) has_fps_expansion F / G" (is "?h has_fps_expansion _") unfolding FG F'_def G'_def n_def using \G \ 0\ by (intro has_fps_expansion_mult has_fps_expansion_inverse has_fps_expansion_shift assms) auto also have "?h = (\x. if x = 0 then c else f x / g x)" using assms(5) unfolding n_def by (intro ext) (auto split: if_splits simp: field_simps) finally show ?thesis . qed lemma has_fps_expansion_divide' [fps_expansion_intros]: fixes F G :: "'a :: {banach, real_normed_field} fps" assumes "f has_fps_expansion F" and "g has_fps_expansion G" and "fps_nth G 0 \ 0" shows "(\x. f x / g x) has_fps_expansion (F / G)" proof - have "(\x. if x = 0 then fps_nth F 0 / fps_nth G 0 else f x / g x) has_fps_expansion (F / G)" (is "?h has_fps_expansion _") using assms(3) by (intro has_fps_expansion_divide assms) auto also from assms have "fps_nth F 0 = f 0" "fps_nth G 0 = g 0" by (auto simp: has_fps_expansion_def eval_fps_at_0 dest: eventually_nhds_x_imp_x) hence "?h = (\x. f x / g x)" by auto finally show ?thesis . qed lemma has_fps_expansion_tan [fps_expansion_intros]: fixes c :: "'a :: {banach, real_normed_field, field_char_0}" shows "(\x. tan (c * x)) has_fps_expansion fps_tan c" proof - have "(\x. sin (c * x) / cos (c * x)) has_fps_expansion fps_sin c / fps_cos c" by (intro fps_expansion_intros) auto thus ?thesis by (simp add: tan_def fps_tan_def) qed lemma has_fps_expansion_tan' [fps_expansion_intros]: "tan has_fps_expansion fps_tan (1 :: 'a :: {banach, real_normed_field, field_char_0})" using has_fps_expansion_tan[of 1] by simp lemma has_fps_expansion_imp_holomorphic: assumes "f has_fps_expansion F" obtains s where "open s" "0 \ s" "f holomorphic_on s" "\z. z \ s \ f z = eval_fps F z" proof - from assms obtain s where s: "open s" "0 \ s" "\z. z \ s \ eval_fps F z = f z" unfolding has_fps_expansion_def eventually_nhds by blast let ?s' = "eball 0 (fps_conv_radius F) \ s" have "eval_fps F holomorphic_on ?s'" by (intro holomorphic_intros) auto also have "?this \ f holomorphic_on ?s'" using s by (intro holomorphic_cong) auto finally show ?thesis using s assms by (intro that[of ?s']) (auto simp: has_fps_expansion_def zero_ereal_def) qed -theorem residue_fps_expansion_over_power_at_0: - assumes "f has_fps_expansion F" - shows "residue (\z. f z / z ^ Suc n) 0 = fps_nth F n" -proof - - from has_fps_expansion_imp_holomorphic[OF assms] guess s . note s = this - have "residue (\z. f z / (z - 0) ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" - using assms s unfolding has_fps_expansion_def - by (intro residue_holomorphic_over_power[of s]) (auto simp: zero_ereal_def) - also from assms have "\ = fps_nth F n" - by (subst fps_nth_fps_expansion) auto - finally show ?thesis by simp -qed - -end +end \ No newline at end of file diff --git a/src/HOL/Analysis/Gamma_Function.thy b/src/HOL/Analysis/Gamma_Function.thy --- a/src/HOL/Analysis/Gamma_Function.thy +++ b/src/HOL/Analysis/Gamma_Function.thy @@ -1,3330 +1,3505 @@ (* Title: HOL/Analysis/Gamma_Function.thy Author: Manuel Eberl, TU München *) section \The Gamma Function\ theory Gamma_Function imports - Conformal_Mappings Equivalence_Lebesgue_Henstock_Integration Summation_Tests Harmonic_Numbers "HOL-Library.Nonpos_Ints" "HOL-Library.Periodic_Fun" begin text \ Several equivalent definitions of the Gamma function and its most important properties. Also contains the definition and some properties of the log-Gamma function and the Digamma function and the other Polygamma functions. Based on the Gamma function, we also prove the Weierstraß product form of the sin function and, based on this, the solution of the Basel problem (the sum over all \<^term>\1 / (n::nat)^2\. \ lemma pochhammer_eq_0_imp_nonpos_Int: "pochhammer (x::'a::field_char_0) n = 0 \ x \ \\<^sub>\\<^sub>0" by (auto simp: pochhammer_eq_0_iff) lemma closed_nonpos_Ints [simp]: "closed (\\<^sub>\\<^sub>0 :: 'a :: real_normed_algebra_1 set)" proof - have "\\<^sub>\\<^sub>0 = (of_int ` {n. n \ 0} :: 'a set)" by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int) also have "closed \" by (rule closed_of_int_image) finally show ?thesis . qed lemma plus_one_in_nonpos_Ints_imp: "z + 1 \ \\<^sub>\\<^sub>0 \ z \ \\<^sub>\\<^sub>0" using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all lemma of_int_in_nonpos_Ints_iff: "(of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ 0" by (auto simp: nonpos_Ints_def) lemma one_plus_of_int_in_nonpos_Ints_iff: "(1 + of_int n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n \ -1" proof - have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp also have "\ \ \\<^sub>\\<^sub>0 \ n + 1 \ 0" by (subst of_int_in_nonpos_Ints_iff) simp_all also have "\ \ n \ -1" by presburger finally show ?thesis . qed lemma one_minus_of_nat_in_nonpos_Ints_iff: "(1 - of_nat n :: 'a :: ring_char_0) \ \\<^sub>\\<^sub>0 \ n > 0" proof - have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp also have "\ \ \\<^sub>\\<^sub>0 \ n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger finally show ?thesis . qed lemma fraction_not_in_ints: assumes "\(n dvd m)" "n \ 0" shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" proof assume "of_int m / (of_int n :: 'a) \ \" then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases) with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: field_split_simps) hence "m = k * n" by (subst (asm) of_int_eq_iff) hence "n dvd m" by simp with assms(1) show False by contradiction qed lemma fraction_not_in_nats: assumes "\n dvd m" "n \ 0" shows "of_int m / of_int n \ (\ :: 'a :: {division_ring,ring_char_0} set)" proof assume "of_int m / of_int n \ (\ :: 'a set)" also note Nats_subset_Ints finally have "of_int m / of_int n \ (\ :: 'a set)" . moreover have "of_int m / of_int n \ (\ :: 'a set)" using assms by (intro fraction_not_in_ints) ultimately show False by contradiction qed lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \ \ \ z \ \\<^sub>\\<^sub>0" by (auto simp: Ints_def nonpos_Ints_def) lemma double_in_nonpos_Ints_imp: assumes "2 * (z :: 'a :: field_char_0) \ \\<^sub>\\<^sub>0" shows "z \ \\<^sub>\\<^sub>0 \ z + 1/2 \ \\<^sub>\\<^sub>0" proof- from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases') thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps) qed lemma sin_series: "(\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" proof - from sin_converges[of z] have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z" . also have "(\n. sin_coeff n *\<^sub>R z^n) sums sin z \ (\n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z" by (subst sums_mono_reindex[of "\n. 2*n+1", symmetric]) (auto simp: sin_coeff_def strict_mono_def ac_simps elim!: oddE) finally show ?thesis . qed lemma cos_series: "(\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" proof - from cos_converges[of z] have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z" . also have "(\n. cos_coeff n *\<^sub>R z^n) sums cos z \ (\n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z" by (subst sums_mono_reindex[of "\n. 2*n", symmetric]) (auto simp: cos_coeff_def strict_mono_def ac_simps elim!: evenE) finally show ?thesis . qed lemma sin_z_over_z_series: fixes z :: "'a :: {real_normed_field,banach}" assumes "z \ 0" shows "(\n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)" proof - from sin_series[of z] have "(\n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z" by (simp add: field_simps scaleR_conv_of_real) from sums_mult[OF this, of "inverse z"] and assms show ?thesis by (simp add: field_simps) qed lemma sin_z_over_z_series': fixes z :: "'a :: {real_normed_field,banach}" assumes "z \ 0" shows "(\n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)" proof - from sums_split_initial_segment[OF sin_converges[of z], of 1] have "(\n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps) qed lemma has_field_derivative_sin_z_over_z: fixes A :: "'a :: {real_normed_field,banach} set" shows "((\z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)" (is "(?f has_field_derivative ?f') _") proof (rule has_field_derivative_at_within) have "((\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) has_field_derivative (\n. diffs (\n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)" proof (rule termdiffs_strong) from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1] show "summable (\n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def) qed simp also have "(\z::'a. \n. of_real (sin_coeff (n+1)) * z^n) = ?f" proof fix z show "(\n. of_real (sin_coeff (n+1)) * z^n) = ?f z" by (cases "z = 0") (insert sin_z_over_z_series'[of z], simp_all add: scaleR_conv_of_real sums_iff sin_coeff_def) qed also have "(\n. diffs (\n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) = diffs (\n. of_real (sin_coeff (Suc n))) 0" by simp also have "\ = 0" by (simp add: sin_coeff_def diffs_def) finally show "((\z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" . qed lemma round_Re_minimises_norm: "norm ((z::complex) - of_int m) \ norm (z - of_int (round (Re z)))" proof - let ?n = "round (Re z)" have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)" by (simp add: cmod_def) also have "\Re z - of_int ?n\ \ \Re z - of_int m\" by (rule round_diff_minimal) hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \ sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)" by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff) also have "\ = norm (z - of_int m)" by (simp add: cmod_def) finally show ?thesis . qed lemma Re_pos_in_ball: assumes "Re z > 0" "t \ ball z (Re z/2)" shows "Re t > 0" proof - have "Re (z - t) \ norm (z - t)" by (rule complex_Re_le_cmod) also from assms have "\ < Re z / 2" by (simp add: dist_complex_def) finally show "Re t > 0" using assms by simp qed lemma no_nonpos_Int_in_ball_complex: assumes "Re z > 0" "t \ ball z (Re z/2)" shows "t \ \\<^sub>\\<^sub>0" using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases) lemma no_nonpos_Int_in_ball: assumes "t \ ball z (dist z (round (Re z)))" shows "t \ \\<^sub>\\<^sub>0" proof assume "t \ \\<^sub>\\<^sub>0" then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases) have "dist z (of_int n) \ dist z t + dist t (of_int n)" by (rule dist_triangle) also from assms have "dist z t < dist z (round (Re z))" by simp also have "\ \ dist z (of_int n)" using round_Re_minimises_norm[of z] by (simp add: dist_complex_def) finally have "dist t (of_int n) > 0" by simp with \t = of_int n\ show False by simp qed lemma no_nonpos_Int_in_ball': assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \ \\<^sub>\\<^sub>0" obtains d where "d > 0" "\t. t \ ball z d \ t \ \\<^sub>\\<^sub>0" proof (rule that) from assms show "setdist {z} \\<^sub>\\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto next fix t assume "t \ ball z (setdist {z} \\<^sub>\\<^sub>0)" thus "t \ \\<^sub>\\<^sub>0" using setdist_le_dist[of z "{z}" t "\\<^sub>\\<^sub>0"] by force qed lemma no_nonpos_Real_in_ball: assumes z: "z \ \\<^sub>\\<^sub>0" and t: "t \ ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" shows "t \ \\<^sub>\\<^sub>0" using z proof (cases "Im z = 0") assume A: "Im z = 0" with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff) with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff) next assume A: "Im z \ 0" have "abs (Im z) - abs (Im t) \ abs (Im z - Im t)" by linarith also have "\ = abs (Im (z - t))" by simp also have "\ \ norm (z - t)" by (rule abs_Im_le_cmod) also from A t have "\ \ abs (Im z) / 2" by (simp add: dist_complex_def) finally have "abs (Im t) > 0" using A by simp thus ?thesis by (force simp add: complex_nonpos_Reals_iff) qed subsection \The Euler form and the logarithmic Gamma function\ text \ We define the Gamma function by first defining its multiplicative inverse \rGamma\. This is more convenient because \rGamma\ is entire, which makes proofs of its properties more convenient because one does not have to watch out for discontinuities. (e.g. \rGamma\ fulfils \rGamma z = z * rGamma (z + 1)\ everywhere, whereas the \\\ function does not fulfil the analogous equation on the non-positive integers) We define the \\\ function (resp.\ its reciprocale) in the Euler form. This form has the advantage that it is a relatively simple limit that converges everywhere. The limit at the poles is 0 (due to division by 0). The functional equation \Gamma (z + 1) = z * Gamma z\ follows immediately from the definition. \ definition\<^marker>\tag important\ Gamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)" definition Gamma_series' :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n" definition rGamma_series :: "('a :: {banach,real_normed_field}) \ nat \ 'a" where "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))" lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)" and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)" unfolding Gamma_series_def rGamma_series_def by simp_all lemma rGamma_series_minus_of_nat: "eventually (\n. rGamma_series (- of_nat k) n = 0) sequentially" using eventually_ge_at_top[of k] by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff) lemma Gamma_series_minus_of_nat: "eventually (\n. Gamma_series (- of_nat k) n = 0) sequentially" using eventually_ge_at_top[of k] by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff) lemma Gamma_series'_minus_of_nat: "eventually (\n. Gamma_series' (- of_nat k) n = 0) sequentially" using eventually_gt_at_top[of k] by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff) lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ rGamma_series z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp) lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp) lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \ \\<^sub>\\<^sub>0 \ Gamma_series' z \ 0" by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp) lemma Gamma_series_Gamma_series': assumes z: "z \ \\<^sub>\\<^sub>0" shows "(\n. Gamma_series' z n / Gamma_series z n) \ 1" proof (rule Lim_transform_eventually) from eventually_gt_at_top[of "0::nat"] show "eventually (\n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n" by (cases n, simp) (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec' dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp) also from n have "\ = z / of_nat n + 1" by (simp add: field_split_simps) finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" .. qed have "(\x. z / of_nat x) \ 0" by (rule tendsto_norm_zero_cancel) (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n], simp add: norm_divide inverse_eq_divide) from tendsto_add[OF this tendsto_const[of 1]] show "(\n. z / of_nat n + 1) \ 1" by simp qed text \ We now show that the series that defines the \\\ function in the Euler form converges and that the function defined by it is continuous on the complex halfspace with positive real part. We do this by showing that the logarithm of the Euler series is continuous and converges locally uniformly, which means that the log-Gamma function defined by its limit is also continuous. This will later allow us to lift holomorphicity and continuity from the log-Gamma function to the inverse of the Gamma function, and from that to the Gamma function itself. \ definition\<^marker>\tag important\ ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\k=1..n. ln (z / of_nat k + 1))" definition\<^marker>\tag unimportant\ ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \ nat \ 'a" where "ln_Gamma_series' z n = - euler_mascheroni*z - ln z + (\k=1..n. z / of_nat n - ln (z / of_nat k + 1))" definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \ 'a" where "ln_Gamma z = lim (ln_Gamma_series z)" text \ We now show that the log-Gamma series converges locally uniformly for all complex numbers except the non-positive integers. We do this by proving that the series is locally Cauchy. \ context begin private lemma ln_Gamma_series_complex_converges_aux: fixes z :: complex and k :: nat assumes z: "z \ 0" and k: "of_nat k \ 2*norm z" "k \ 2" shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \ 2*(norm z + norm z^2) / of_nat k^2" proof - let ?k = "of_nat k :: complex" and ?z = "norm z" have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)" by (simp add: algebra_simps) also have "norm ... \ ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)" by (subst norm_mult [symmetric], rule norm_triangle_ineq) also have "norm (Ln (1 + -1/?k) - (-1/?k)) \ (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))" using k by (intro Ln_approx_linear) (simp add: norm_divide) hence "?z * norm (ln (1-1/?k) + 1/?k) \ ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))" by (intro mult_left_mono) simp_all also have "... \ (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k by (simp add: field_simps power2_eq_square norm_divide) also have "... \ (?z * 2) / of_nat k^2" using k by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) also have "norm (ln (1+z/?k) - z/?k) \ norm (z/?k)^2 / (1 - norm (z/?k))" using k by (intro Ln_approx_linear) (simp add: norm_divide) hence "norm (ln (1+z/?k) - z/?k) \ ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)" by (simp add: field_simps norm_divide) also have "... \ (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k by (simp add: field_simps power2_eq_square) also have "... \ (?z^2 * 2) / of_nat k^2" using k by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps) also note add_divide_distrib [symmetric] finally show ?thesis by (simp only: distrib_left mult.commute) qed lemma ln_Gamma_series_complex_converges: assumes z: "z \ \\<^sub>\\<^sub>0" assumes d: "d > 0" "\n. n \ \\<^sub>\\<^sub>0 \ norm (z - of_int n) > d" shows "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n :: complex)" proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI') fix e :: real assume e: "e > 0" define e'' where "e'' = (SUP t\ball z d. norm t + norm t^2)" define e' where "e' = e / (2*e'')" have "bounded ((\t. norm t + norm t^2) ` cball z d)" by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros) hence "bounded ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto hence bdd: "bdd_above ((\t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above) with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z]) have e'': "norm t + norm t^2 \ e''" if "t \ ball z d" for t unfolding e''_def using that by (rule cSUP_upper[OF _ bdd]) from e z e''_pos have e': "e' > 0" unfolding e'_def by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps) have "summable (\k. inverse ((real_of_nat k)^2))" by (rule inverse_power_summable) simp from summable_partial_sum_bound[OF this e'] guess M . note M = this define N where "N = max 2 (max (nat \2 * (norm z + d)\) M)" { from d have "\2 * (cmod z + d)\ \ \0::real\" by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all hence "2 * (norm z + d) \ of_nat (nat \2 * (norm z + d)\)" unfolding N_def by (simp_all add: le_of_int_ceiling) also have "... \ of_nat N" unfolding N_def by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1) finally have "of_nat N \ 2 * (norm z + d)" . moreover have "N \ 2" "N \ M" unfolding N_def by simp_all moreover have "(\k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \ N" for m n using M[OF order.trans[OF \N \ M\ that]] unfolding real_norm_def by (subst (asm) abs_of_nonneg) (auto intro: sum_nonneg simp: field_split_simps) moreover note calculation } note N = this show "\M. \t\ball z d. \m\M. \n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e" unfolding dist_complex_def proof (intro exI[of _ N] ballI allI impI) fix t m n assume t: "t \ ball z d" and mn: "m \ N" "n > m" from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def) also have "dist z 0 - dist z t \ dist 0 t" using dist_triangle[of 0 z t] by (simp add: dist_commute) finally have t_nz: "t \ 0" by auto have "norm t \ norm z + norm (t - z)" by (rule norm_triangle_sub) also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute) also have "2 * (norm z + d) \ of_nat N" by (rule N) also have "N \ m" by (rule mn) finally have norm_t: "2 * norm t < of_nat m" by simp have "ln_Gamma_series t m - ln_Gamma_series t n = (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) + ((\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)))" by (simp add: ln_Gamma_series_def algebra_simps) also have "(\k=1..n. Ln (t / of_nat k + 1)) - (\k=1..m. Ln (t / of_nat k + 1)) = (\k\{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn by (simp add: sum_diff) also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) = (\k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn by (subst sum_telescope'' [symmetric]) simp_all also have "... = (\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N by (intro sum_cong_Suc) (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat) also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \ {Suc m..n}" for k using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: field_split_simps) hence "(\k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) = (\k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N by (intro sum.cong) simp_all also note sum.distrib [symmetric] also have "norm (\k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \ (\k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t by (intro order.trans[OF norm_sum sum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all also have "... \ 2 * (norm t + norm t^2) * (\k=Suc m..n. 1 / (of_nat k)\<^sup>2)" by (simp add: sum_distrib_left) also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')" by (simp add: e'_def field_simps power2_eq_square) also from e''[OF t] e''_pos e have "\ \ e * 1" by (intro mult_left_mono) (simp_all add: field_simps) finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp qed qed end lemma ln_Gamma_series_complex_converges': assumes z: "(z :: complex) \ \\<^sub>\\<^sub>0" shows "\d>0. uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" proof - define d' where "d' = Re z" define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)" have "of_int (round d') \ \\<^sub>\\<^sub>0" if "d' \ 0" using that by (intro nonpos_Ints_of_int) (simp_all add: round_def) with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less) have "d < cmod (z - of_int n)" if "n \ \\<^sub>\\<^sub>0" for n proof (cases "Re z > 0") case True from nonpos_Ints_nonpos[OF that] have n: "n \ 0" by simp from True have "d = Re z/2" by (simp add: d_def d'_def) also from n True have "\ < Re (z - of_int n)" by simp also have "\ \ norm (z - of_int n)" by (rule complex_Re_le_cmod) finally show ?thesis . next case False with assms nonpos_Ints_of_int[of "round (Re z)"] have "z \ of_int (round d')" by (auto simp: not_less) with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def) also have "\ \ norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm) finally show ?thesis . qed hence conv: "uniformly_convergent_on (ball z d) (\n z. ln_Gamma_series z n)" by (intro ln_Gamma_series_complex_converges d_pos z) simp_all from d_pos conv show ?thesis by blast qed lemma ln_Gamma_series_complex_converges'': "(z :: complex) \ \\<^sub>\\<^sub>0 \ convergent (ln_Gamma_series z)" by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent) theorem ln_Gamma_complex_LIMSEQ: "(z :: complex) \ \\<^sub>\\<^sub>0 \ ln_Gamma_series z \ ln_Gamma z" using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def) lemma exp_ln_Gamma_series_complex: assumes "n > 0" "z \ \\<^sub>\\<^sub>0" shows "exp (ln_Gamma_series z n :: complex) = Gamma_series z n" proof - from assms obtain m where m: "n = Suc m" by (cases n) blast from assms have "z \ 0" by (intro notI) auto with assms have "exp (ln_Gamma_series z n) = (of_nat n) powr z / (z * (\k=1..n. exp (Ln (z / of_nat k + 1))))" unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_sum) also from assms have "(\k=1..n. exp (Ln (z / of_nat k + 1))) = (\k=1..n. z / of_nat k + 1)" by (intro prod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp) also have "... = (\k=1..n. z + k) / fact n" by (simp add: fact_prod) (subst prod_dividef [symmetric], simp_all add: field_simps) also from m have "z * ... = (\k=0..n. z + k) / fact n" by (simp add: prod.atLeast0_atMost_Suc_shift prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) also have "(\k=0..n. z + k) = pochhammer z (Suc n)" unfolding pochhammer_prod by (simp add: prod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost) also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n" unfolding Gamma_series_def using assms by (simp add: field_split_simps powr_def) finally show ?thesis . qed lemma ln_Gamma_series'_aux: assumes "(z::complex) \ \\<^sub>\\<^sub>0" shows "(\k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s") unfolding sums_def proof (rule Lim_transform) show "(\n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \ ?s" (is "?g \ _") by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms) have A: "eventually (\n. (\k 0" have "(\kk=1..n. z / of_nat k - ln (1 + z / of_nat k))" by (subst atLeast0LessThan [symmetric], subst sum.shift_bounds_Suc_ivl [symmetric], subst atLeastLessThanSuc_atLeastAtMost) simp_all also have "\ = z * of_real (harm n) - (\k=1..n. ln (1 + z / of_nat k))" by (simp add: harm_def sum_subtractf sum_distrib_left divide_inverse) also from n have "\ - ?g n = 0" by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps Ln_of_nat) finally show "(\kn. (\k 0" by (subst tendsto_cong[OF A]) simp_all qed lemma uniformly_summable_deriv_ln_Gamma: assumes z: "(z :: 'a :: {real_normed_field,banach}) \ 0" and d: "d > 0" "d \ norm z/2" shows "uniformly_convergent_on (ball z d) (\k z. \ik z. \i ball z d" have "norm z = norm (t + (z - t))" by simp have "norm (t + (z - t)) \ norm t + norm (z - t)" by (rule norm_triangle_ineq) also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm) finally have A: "norm t > norm z / 2" using z by (simp add: field_simps) have "norm t = norm (z + (t - z))" by simp also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) also from t d have "norm (t - z) \ norm z / 2" by (simp add: dist_norm norm_minus_commute) also from z have "\ < norm z" by simp finally have B: "norm t < 2 * norm z" by simp note A B } note ball = this show "eventually (\n. \t\ball z d. norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially" using eventually_gt_at_top apply eventually_elim proof safe fix t :: 'a assume t: "t \ ball z d" from z ball[OF t] have t_nz: "t \ 0" by auto fix n :: nat assume n: "n > nat \4 * norm z\" from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp also from n have "\ < of_nat n" by linarith finally have n: "of_nat n > 2 * norm t" . hence "of_nat n > norm t" by simp hence t': "t \ -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc) with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))" by (simp add: field_split_simps eq_neg_iff_add_eq_0 del: of_nat_Suc) also have "norm \ = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))" by (simp add: norm_divide norm_mult field_split_simps add_ac del: of_nat_Suc) also { from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \ of_nat (Suc n) / (2 * norm t)" by (intro divide_left_mono mult_pos_pos) simp_all also have "\ < norm (of_nat (Suc n) / t) - norm (1 :: 'a)" using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc) also have "\ \ norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq) finally have "inverse (norm (of_nat (Suc n)/t + 1)) \ 4 * norm z / of_nat (Suc n)" using z by (simp add: field_split_simps norm_divide mult_ac del: of_nat_Suc) } also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) = 4 * norm z * inverse (of_nat (Suc n)^2)" by (simp add: field_split_simps power2_eq_square del: of_nat_Suc) finally show "norm (?f n t) \ 4 * norm z * inverse (of_nat (Suc n)^2)" by (simp del: of_nat_Suc) qed next show "summable (\n. 4 * norm z * inverse ((of_nat (Suc n))^2))" by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable) qed subsection \The Polygamma functions\ lemma summable_deriv_ln_Gamma: "z \ (0 :: 'a :: {real_normed_field,banach}) \ summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))" unfolding summable_iff_convergent by (rule uniformly_convergent_imp_convergent, rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all definition\<^marker>\tag important\ Polygamma :: "nat \ ('a :: {real_normed_field,banach}) \ 'a" where "Polygamma n z = (if n = 0 then (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else (-1)^Suc n * fact n * (\k. inverse ((z + of_nat k)^Suc n)))" abbreviation\<^marker>\tag important\ Digamma :: "('a :: {real_normed_field,banach}) \ 'a" where "Digamma \ Polygamma 0" lemma Digamma_def: "Digamma z = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni" by (simp add: Polygamma_def) lemma summable_Digamma: assumes "(z :: 'a :: {real_normed_field,banach}) \ 0" shows "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" proof - have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums (0 - inverse (z + of_nat 0))" by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]] show "summable (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp qed lemma summable_offset: assumes "summable (\n. f (n + k) :: 'a :: real_normed_vector)" shows "summable f" proof - from assms have "convergent (\m. \nm. (\nnm. (\nnm. \n {k..n\\. f n) = (\nn=k..n=k..n=0..nnna. sum f {.. lim (\m. sum f {.. 0" and n: "n \ 2" shows "uniformly_convergent_on (ball z d) (\k z. \inorm z * e\" { fix t assume t: "t \ ball z d" have "norm t = norm (z + (t - z))" by simp also have "\ \ norm z + norm (t - z)" by (rule norm_triangle_ineq) also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute) finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def) } note ball = this show "eventually (\k. \t\ball z d. norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)) sequentially" using eventually_gt_at_top[of m] apply eventually_elim proof (intro ballI) fix k :: nat and t :: 'a assume k: "k > m" and t: "t \ ball z d" from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff) also have "\ \ norm (of_nat k :: 'a) - norm z * e" unfolding m_def by (subst norm_of_nat) linarith also from ball[OF t] have "\ \ norm (of_nat k :: 'a) - norm t" by simp also have "\ \ norm (of_nat k + t)" by (rule norm_diff_ineq) finally have "inverse ((norm (t + of_nat k))^n) \ inverse (real_of_nat (k - m)^n)" using k n by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc) thus "norm (inverse ((t + of_nat k)^n)) \ inverse (of_nat (k - m)^n)" by (simp add: norm_inverse norm_power power_inverse) qed have "summable (\k. inverse ((real_of_nat k)^n))" using inverse_power_summable[of n] n by simp hence "summable (\k. inverse ((real_of_nat (k + m - m))^n))" by simp thus "summable (\k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset) qed lemma Polygamma_converges': fixes z :: "'a :: {real_normed_field,banach}" assumes z: "z \ 0" and n: "n \ 2" shows "summable (\k. inverse ((z + of_nat k)^n))" using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z] by (simp add: summable_iff_convergent) theorem Digamma_LIMSEQ: fixes z :: "'a :: {banach,real_normed_field}" assumes z: "z \ 0" shows "(\m. of_real (ln (real m)) - (\n Digamma z" proof - have "(\n. of_real (ln (real n / (real (Suc n))))) \ (of_real (ln 1) :: 'a)" by (intro tendsto_intros LIMSEQ_n_over_Suc_n) simp_all hence "(\n. of_real (ln (real n / (real n + 1)))) \ (0 :: 'a)" by (simp add: add_ac) hence lim: "(\n. of_real (ln (real n)) - of_real (ln (real n + 1))) \ (0::'a)" proof (rule Lim_transform_eventually) show "eventually (\n. of_real (ln (real n / (real n + 1))) = of_real (ln (real n)) - (of_real (ln (real n + 1)) :: 'a)) at_top" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_div) qed from summable_Digamma[OF z] have "(\n. inverse (of_nat (n+1)) - inverse (z + of_nat n)) sums (Digamma z + euler_mascheroni)" by (simp add: Digamma_def summable_sums) from sums_diff[OF this euler_mascheroni_sum] have "(\n. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)) - inverse (z + of_nat n)) sums Digamma z" by (simp add: add_ac) hence "(\m. (\nn Digamma z" by (simp add: sums_def sum_subtractf) also have "(\m. (\nm. of_real (ln (m + 1)) :: 'a)" by (subst sum_lessThan_telescope) simp_all finally show ?thesis by (rule Lim_transform) (insert lim, simp) qed theorem Polygamma_LIMSEQ: fixes z :: "'a :: {banach,real_normed_field}" assumes "z \ 0" and "n > 0" shows "(\k. inverse ((z + of_nat k)^Suc n)) sums ((-1) ^ Suc n * Polygamma n z / fact n)" using Polygamma_converges'[OF assms(1), of "Suc n"] assms(2) by (simp add: sums_iff Polygamma_def) theorem has_field_derivative_ln_Gamma_complex [derivative_intros]: fixes z :: complex assumes z: "z \ \\<^sub>\\<^sub>0" shows "(ln_Gamma has_field_derivative Digamma z) (at z)" proof - have not_nonpos_Int [simp]: "t \ \\<^sub>\\<^sub>0" if "Re t > 0" for t using that by (auto elim!: nonpos_Ints_cases') from z have z': "z \ \\<^sub>\\<^sub>0" and z'': "z \ 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I by blast+ let ?f' = "\z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))" let ?f = "\z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\z. \n. ?f' z n" define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)" from z have d: "d > 0" "norm z/2 \ d" by (auto simp add: complex_nonpos_Reals_iff d_def) have ball: "Im t = 0 \ Re t > 0" if "dist z t < d" for t using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff) have sums: "(\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums (0 - inverse (z + of_nat 0))" by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) have "((\z. \n. ?f z n) has_field_derivative ?F' z) (at z)" using d z ln_Gamma_series'_aux[OF z'] apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma) apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff simp del: of_nat_Suc) apply (auto simp add: complex_nonpos_Reals_iff) done with z have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative ?F' z - euler_mascheroni - inverse z) (at z)" by (force intro!: derivative_eq_intros simp: Digamma_def) also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp also from sums have "-inverse z = (\n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))" by (simp add: sums_iff) also from sums summable_deriv_ln_Gamma[OF z''] have "?F' z + \ = (\n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by (subst suminf_add) (simp_all add: add_ac sums_iff) also have "\ - euler_mascheroni = Digamma z" by (simp add: Digamma_def) finally have "((\z. (\k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative Digamma z) (at z)" . moreover from eventually_nhds_ball[OF d(1), of z] have "eventually (\z. ln_Gamma z = (\k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)" proof eventually_elim fix t assume "t \ ball z d" hence "t \ \\<^sub>\\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases) from ln_Gamma_series'_aux[OF this] show "ln_Gamma t = (\k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff) qed ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) qed declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros] lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni" by (simp add: Digamma_def) lemma Digamma_plus1: assumes "z \ 0" shows "Digamma (z+1) = Digamma z + 1/z" proof - have sums: "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) sums (inverse (z + of_nat 0) - 0)" by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]] tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat) have "Digamma (z+1) = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) - euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac) also have "suminf ?f = (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) + (\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))" using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff) also have "(\k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z" using sums by (simp add: sums_iff inverse_eq_divide) finally show ?thesis by (simp add: Digamma_def[of z]) qed theorem Polygamma_plus1: assumes "z \ 0" shows "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" proof (cases "n = 0") assume n: "n \ 0" let ?f = "\k. inverse ((z + of_nat k) ^ Suc n)" have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\k. ?f (k+1))" using n by (simp add: Polygamma_def add_ac) also have "(\k. ?f (k+1)) + (\k<1. ?f k) = (\k. ?f k)" using Polygamma_converges'[OF assms, of "Suc n"] n by (subst suminf_split_initial_segment [symmetric]) simp_all hence "(\k. ?f (k+1)) = (\k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps) also have "(-1) ^ Suc n * fact n * ((\k. ?f k) - inverse (z ^ Suc n)) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n by (simp add: inverse_eq_divide algebra_simps Polygamma_def) finally show ?thesis . qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide) theorem Digamma_of_nat: "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni" proof (induction n) case (Suc n) have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp also have "\ = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))" by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc) also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc) also have "\ + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni" by (simp add: harm_Suc) finally show ?case . qed (simp add: harm_def) lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni" by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl) lemma Polygamma_of_real: "x \ 0 \ Polygamma n (of_real x) = of_real (Polygamma n x)" unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"] by (simp_all add: suminf_of_real) lemma Polygamma_Real: "z \ \ \ z \ 0 \ Polygamma n z \ \" by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all lemma Digamma_half_integer: "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) = (\kk. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni" by (simp add: Digamma_def add_ac) also have "(\k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) = (\k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))" by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide) also have "\ = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums'] by (subst suminf_mult) (simp_all add: algebra_simps sums_iff) finally show ?case by simp next case (Suc n) have nz: "2 * of_nat n + (1:: 'a) \ 0" using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac) hence nz': "of_nat n + (1/2::'a) \ 0" by (simp add: field_simps) have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp also from nz' have "\ = Digamma (of_nat n + 1/2) + 1 / (of_nat n + 1/2)" by (rule Digamma_plus1) also from nz nz' have "1 / (of_nat n + 1/2 :: 'a) = 2 / (2 * of_nat n + 1)" by (subst divide_eq_eq) simp_all also note Suc finally show ?case by (simp add: add_ac) qed lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)" using Digamma_half_integer[of 0] by simp lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0" proof - have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp also have "\ = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp also note euler_mascheroni_less_13_over_22 also note ln2_le_25_over_36 finally show ?thesis by simp qed theorem has_field_derivative_Polygamma [derivative_intros]: fixes z :: "'a :: {real_normed_field,euclidean_space}" assumes z: "z \ \\<^sub>\\<^sub>0" shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)" proof (rule has_field_derivative_at_within, cases "n = 0") assume n: "n = 0" let ?f = "\k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)" let ?F = "\z. \k. ?f k z" and ?f' = "\k z. inverse ((z + of_nat k)\<^sup>2)" from no_nonpos_Int_in_ball'[OF z] guess d . note d = this from z have summable: "summable (\k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))" by (intro summable_Digamma) force from z have conv: "uniformly_convergent_on (ball z d) (\k z. \i2))" by (intro Polygamma_converges) auto with d have "summable (\k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent ) have "(?F has_field_derivative (\k. ?f' k z)) (at z)" proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) fix k :: nat and t :: 'a assume t: "t \ ball z d" from t d(2)[of t] show "((\z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)" by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases) qed (insert d(1) summable conv, (assumption|simp)+) with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n by (force simp: power2_eq_square intro!: derivative_eq_intros) next assume n: "n \ 0" from z have z': "z \ 0" by auto from no_nonpos_Int_in_ball'[OF z] guess d . note d = this define n' where "n' = Suc n" from n have n': "n' \ 2" by (simp add: n'_def) have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative (\k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)" proof (rule has_field_derivative_series'[of "ball z d" _ _ z]) fix k :: nat and t :: 'a assume t: "t \ ball z d" with d have t': "t \ \\<^sub>\\<^sub>0" "t \ 0" by auto show "((\a. inverse ((a + of_nat k) ^ n')) has_field_derivative - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t' by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp) next have "uniformly_convergent_on (ball z d) (\k z. (- of_nat n' :: 'a) * (\ik z. \ik. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) = (- of_nat n') * (\k. inverse ((z + of_nat k) ^ (n' + 1)))" using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all finally have "((\z. \k. inverse ((z + of_nat k) ^ n')) has_field_derivative - of_nat n' * (\k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" . from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"] show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)" unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps) qed declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros] lemma isCont_Polygamma [continuous_intros]: fixes f :: "_ \ 'a :: {real_normed_field,euclidean_space}" shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Polygamma n (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Polygamma]]) lemma continuous_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A (Polygamma n :: _ \ 'a :: {real_normed_field,euclidean_space})" by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast lemma isCont_ln_Gamma_complex [continuous_intros]: fixes f :: "'a::t2_space \ complex" shows "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\z. ln_Gamma (f z)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]]) lemma continuous_on_ln_Gamma_complex [continuous_intros]: fixes A :: "complex set" shows "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A ln_Gamma" by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident]) fastforce lemma deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "deriv (Polygamma m) z = Polygamma (Suc m) (z :: 'a :: {real_normed_field,euclidean_space})" by (intro DERIV_imp_deriv has_field_derivative_Polygamma assms) thm has_field_derivative_Polygamma lemma higher_deriv_Polygamma: assumes "z \ \\<^sub>\\<^sub>0" shows "(deriv ^^ n) (Polygamma m) z = Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})" proof - have "eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)" proof (induction n) case (Suc n) from Suc.IH have "eventually (\z. eventually (\u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)" by (simp add: eventually_eventually) hence "eventually (\z. deriv ((deriv ^^ n) (Polygamma m)) z = deriv (Polygamma (m + n)) z) (nhds z)" by eventually_elim (intro deriv_cong_ev refl) moreover have "eventually (\z. z \ UNIV - \\<^sub>\\<^sub>0) (nhds z)" using assms by (intro eventually_nhds_in_open open_Diff open_UNIV) auto ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma) qed simp_all thus ?thesis by (rule eventually_nhds_x_imp_x) qed lemma deriv_ln_Gamma_complex: assumes "z \ \\<^sub>\\<^sub>0" shows "deriv ln_Gamma z = Digamma (z :: complex)" by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_complex assms) text \ We define a type class that captures all the fundamental properties of the inverse of the Gamma function and defines the Gamma function upon that. This allows us to instantiate the type class both for the reals and for the complex numbers with a minimal amount of proof duplication. \ class\<^marker>\tag unimportant\ Gamma = real_normed_field + complete_space + fixes rGamma :: "'a \ 'a" assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \ (\n. z = - of_nat n)" assumes differentiable_rGamma_aux1: "(\n. z \ - of_nat n) \ let d = (THE d. (\n. \k d) - scaleR euler_mascheroni 1 in filterlim (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R norm (y - z)) (nhds 0) (at z)" assumes differentiable_rGamma_aux2: "let z = - of_nat n in filterlim (\y. (rGamma y - rGamma z - (-1)^n * (prod of_nat {1..n}) * (y - z)) /\<^sub>R norm (y - z)) (nhds 0) (at z)" assumes rGamma_series_aux: "(\n. z \ - of_nat n) \ let fact' = (\n. prod of_nat {1..n}); exp = (\x. THE e. (\n. \kR fact k) \ e); pochhammer' = (\a n. (\n = 0..n. a + of_nat n)) in filterlim (\n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1)))) (nhds (rGamma z)) sequentially" begin subclass banach .. end definition "Gamma z = inverse (rGamma z)" subsection \Basic properties\ lemma Gamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ Gamma z = 0" and rGamma_nonpos_Int: "z \ \\<^sub>\\<^sub>0 \ rGamma z = 0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma Gamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ Gamma z \ 0" and rGamma_nonzero: "z \ \\<^sub>\\<^sub>0 \ rGamma z \ 0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma Gamma_eq_zero_iff: "Gamma z = 0 \ z \ \\<^sub>\\<^sub>0" and rGamma_eq_zero_iff: "rGamma z = 0 \ z \ \\<^sub>\\<^sub>0" using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases') lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)" unfolding Gamma_def by simp lemma rGamma_series_LIMSEQ [tendsto_intros]: "rGamma_series z \ rGamma z" proof (cases "z \ \\<^sub>\\<^sub>0") case False hence "z \ - of_nat n" for n by auto from rGamma_series_aux[OF this] show ?thesis by (simp add: rGamma_series_def[abs_def] fact_prod pochhammer_Suc_prod exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost) qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ) theorem Gamma_series_LIMSEQ [tendsto_intros]: "Gamma_series z \ Gamma z" proof (cases "z \ \\<^sub>\\<^sub>0") case False hence "(\n. inverse (rGamma_series z n)) \ inverse (rGamma z)" by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff) also have "(\n. inverse (rGamma_series z n)) = Gamma_series z" by (simp add: rGamma_series_def Gamma_series_def[abs_def]) finally show ?thesis by (simp add: Gamma_def) qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ) lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)" using Gamma_series_LIMSEQ[of z] by (simp add: limI) lemma rGamma_1 [simp]: "rGamma 1 = 1" proof - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact field_split_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) moreover have "rGamma_series 1 \ rGamma 1" by (rule tendsto_intros) ultimately show ?thesis by (intro LIMSEQ_unique) qed lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z" proof - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) also from n have "\ = ?f n * rGamma_series z n" by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def add_ac) finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. qed moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" by (intro tendsto_intros lim_inverse_n) hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" by (blast intro: Lim_transform_eventually) moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" by (intro tendsto_intros) ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast qed lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)" proof (induction n arbitrary: z) case (Suc n z) have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH) also note rGamma_plus1 [symmetric] finally show ?case by (simp add: add_ac pochhammer_rec') qed simp_all theorem Gamma_plus1: "z \ \\<^sub>\\<^sub>0 \ Gamma (z + 1) = z * Gamma z" using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff) theorem pochhammer_Gamma: "z \ \\<^sub>\\<^sub>0 \ pochhammer z n = Gamma (z + of_nat n) / Gamma z" using pochhammer_rGamma[of z] by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps) lemma Gamma_0 [simp]: "Gamma 0 = 0" and rGamma_0 [simp]: "rGamma 0 = 0" and Gamma_neg_1 [simp]: "Gamma (- 1) = 0" and rGamma_neg_1 [simp]: "rGamma (- 1) = 0" and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0" and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0" and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0" and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0" by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff) lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp theorem Gamma_fact: "Gamma (1 + of_nat n) = fact n" by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff flip: of_nat_Suc) lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)" by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst of_nat_Suc, subst Gamma_fact) (rule refl) lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)" proof (cases "n > 0") case True hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int) lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)" by (simp add: Gamma_of_int rGamma_inverse_Gamma) lemma Gamma_seriesI: assumes "(\n. g n / Gamma_series z n) \ 1" shows "g \ Gamma z" proof (rule Lim_transform_eventually) have "1/2 > (0::real)" by simp from tendstoD[OF assms, OF this] show "eventually (\n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially" by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) from assms have "(\n. g n / Gamma_series z n * Gamma_series z n) \ 1 * Gamma z" by (intro tendsto_intros) thus "(\n. g n / Gamma_series z n * Gamma_series z n) \ Gamma z" by simp qed lemma Gamma_seriesI': assumes "f \ rGamma z" assumes "(\n. g n * f n) \ 1" assumes "z \ \\<^sub>\\<^sub>0" shows "g \ Gamma z" proof (rule Lim_transform_eventually) have "1/2 > (0::real)" by simp from tendstoD[OF assms(2), OF this] show "eventually (\n. g n * f n / f n = g n) sequentially" by (force elim!: eventually_mono simp: dist_real_def dist_0_norm) from assms have "(\n. g n * f n / f n) \ 1 / rGamma z" by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff) thus "(\n. g n * f n / f n) \ Gamma z" by (simp add: Gamma_def divide_inverse) qed lemma Gamma_series'_LIMSEQ: "Gamma_series' z \ Gamma z" by (cases "z \ \\<^sub>\\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series'] Gamma_series'_nonpos_Ints_LIMSEQ[of z]) subsection \Differentiability\ lemma has_field_derivative_rGamma_no_nonpos_int: assumes "z \ \\<^sub>\\<^sub>0" shows "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)" proof (rule has_field_derivative_at_within) from assms have "z \ - of_nat n" for n by auto from differentiable_rGamma_aux1[OF this] show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)" unfolding Digamma_def suminf_def sums_def[abs_def] has_field_derivative_def has_derivative_def netlimit_at by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric]) qed lemma has_field_derivative_rGamma_nonpos_int: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)" apply (rule has_field_derivative_at_within) using differentiable_rGamma_aux2[of n] unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_prod) simp lemma has_field_derivative_rGamma [derivative_intros]: "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \norm z\) * fact (nat \norm z\) else -rGamma z * Digamma z)) (at z within A)" using has_field_derivative_rGamma_no_nonpos_int[of z A] has_field_derivative_rGamma_nonpos_int[of "nat \norm z\" A] by (auto elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros] declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros] declare has_field_derivative_rGamma_nonpos_int [derivative_intros] declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros] declare has_field_derivative_rGamma [derivative_intros] theorem has_field_derivative_Gamma [derivative_intros]: "z \ \\<^sub>\\<^sub>0 \ (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)" unfolding Gamma_def [abs_def] by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff) declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros] (* TODO: Hide ugly facts properly *) hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2 differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma" by (rule DERIV_continuous_on has_field_derivative_rGamma)+ lemma continuous_on_Gamma [continuous_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ continuous_on A Gamma" by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast lemma isCont_rGamma [continuous_intros]: "isCont f z \ isCont (\x. rGamma (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_rGamma]]) lemma isCont_Gamma [continuous_intros]: "isCont f z \ f z \ \\<^sub>\\<^sub>0 \ isCont (\x. Gamma (f x)) z" by (rule isCont_o2[OF _ DERIV_isCont[OF has_field_derivative_Gamma]]) subsection\<^marker>\tag unimportant\ \The complex Gamma function\ instantiation\<^marker>\tag unimportant\ complex :: Gamma begin definition\<^marker>\tag unimportant\ rGamma_complex :: "complex \ complex" where "rGamma_complex z = lim (rGamma_series z)" lemma rGamma_series_complex_converges: "convergent (rGamma_series (z :: complex))" (is "?thesis1") and rGamma_complex_altdef: "rGamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2") proof - have "?thesis1 \ ?thesis2" proof (cases "z \ \\<^sub>\\<^sub>0") case False have "rGamma_series z \ exp (- ln_Gamma z)" proof (rule Lim_transform_eventually) from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE) from this(1) uniformly_convergent_imp_convergent[OF this(2), of z] have "ln_Gamma_series z \ lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff) thus "(\n. exp (-ln_Gamma_series z n)) \ exp (- ln_Gamma z)" unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus) from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False show "eventually (\n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially" by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def) qed with False show ?thesis by (auto simp: convergent_def rGamma_complex_def intro!: limI) next case True then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases') also have "rGamma_series \ \ 0" by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const) finally show ?thesis using True by (auto simp: rGamma_complex_def convergent_def intro!: limI) qed thus "?thesis1" "?thesis2" by blast+ qed context\<^marker>\tag unimportant\ begin (* TODO: duplication *) private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)" proof - let ?f = "\n. (z + 1) * inverse (of_nat n) + 1" have "eventually (\n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" hence "z * rGamma_series (z + 1) n = inverse (of_nat n) * pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))" by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real) also from n have "\ = ?f n * rGamma_series z n" by (subst pochhammer_rec') (simp_all add: field_split_simps rGamma_series_def add_ac) finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" .. qed moreover have "(\n. ?f n * rGamma_series z n) \ ((z+1) * 0 + 1) * rGamma z" using rGamma_series_complex_converges by (intro tendsto_intros lim_inverse_n) (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def) hence "(\n. ?f n * rGamma_series z n) \ rGamma z" by simp ultimately have "(\n. z * rGamma_series (z + 1) n) \ rGamma z" by (blast intro: Lim_transform_eventually) moreover have "(\n. z * rGamma_series (z + 1) n) \ z * rGamma (z + 1)" using rGamma_series_complex_converges by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff) ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast qed private lemma has_field_derivative_rGamma_complex_no_nonpos_Int: assumes "(z :: complex) \ \\<^sub>\\<^sub>0" shows "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" proof - have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z proof (subst DERIV_cong_ev[OF refl _ refl]) from that have "eventually (\t. t \ ball z (Re z/2)) (nhds z)" by (intro eventually_nhds_in_nhd) simp_all thus "eventually (\t. rGamma t = exp (- ln_Gamma t)) (nhds z)" using no_nonpos_Int_in_ball_complex[OF that] by (auto elim!: eventually_mono simp: rGamma_complex_altdef) next have "z \ \\<^sub>\\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff) with that show "((\t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)" by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef) qed from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" proof (induction "nat \1 - Re z\" arbitrary: z) case (Suc n z) from Suc.prems have z: "z \ 0" by auto from Suc.hyps have "n = nat \- Re z\" by linarith hence A: "n = nat \1 - Re (z + 1)\" by simp from Suc.prems have B: "z + 1 \ \\<^sub>\\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp) have "((\z. z * (rGamma \ (\z. z + 1)) z) has_field_derivative -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)" by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps) also have "(\z. z * (rGamma \ (\z. z + 1 :: complex)) z) = rGamma" by (simp add: rGamma_complex_plus1) also from z have "Digamma (z + 1) * z - 1 = z * Digamma z" by (subst Digamma_plus1) (simp_all add: field_simps) also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z" by (simp add: rGamma_complex_plus1[of z, symmetric]) finally show ?case . qed (intro diff, simp) qed private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1" proof - have A: "eventually (\n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially" using eventually_gt_at_top[of "0::nat"] by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact field_split_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int) have "rGamma_series 1 \ 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n) thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI) qed private lemma has_field_derivative_rGamma_complex_nonpos_Int: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))" proof (induction n) case 0 have A: "(0::complex) + 1 \ \\<^sub>\\<^sub>0" by simp have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)" by (rule derivative_eq_intros DERIV_chain refl has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1) thus ?case by (simp add: rGamma_complex_plus1) next case (Suc n) hence A: "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat (Suc n) + 1 :: complex))" by simp have "((\z. z * (rGamma \ (\z. z + 1 :: complex)) z) has_field_derivative (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))" by (rule derivative_eq_intros refl A DERIV_chain)+ (simp add: algebra_simps rGamma_complex_altdef) thus ?case by (simp add: rGamma_complex_plus1) qed instance proof fix z :: complex show "(rGamma z = 0) \ (\n. z = - of_nat n)" by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases') next fix z :: complex assume "\n. z \ - of_nat n" hence "z \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases') from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this] show "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] netlimit_at of_real_def[symmetric] suminf_def) next fix n :: nat from has_field_derivative_rGamma_complex_nonpos_Int[of n] show "let z = - of_nat n in (\y. (rGamma y - rGamma z - (- 1) ^ n * prod of_nat {1..n} * (y - z)) /\<^sub>R cmod (y - z)) \z\ 0" by (simp add: has_field_derivative_def has_derivative_def fact_prod netlimit_at Let_def) next fix z :: complex from rGamma_series_complex_converges[of z] have "rGamma_series z \ rGamma z" by (simp add: convergent_LIMSEQ_iff rGamma_complex_def) thus "let fact' = \n. prod of_nat {1..n}; exp = \x. THE e. (\n. \kR fact k) \ e; pochhammer' = \a n. \n = 0..n. a + of_nat n in (\n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma z" by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) qed end end lemma Gamma_complex_altdef: "Gamma z = (if z \ \\<^sub>\\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))" unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus) lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)" proof - have "rGamma_series (cnj z) = (\n. cnj (rGamma_series z n))" by (intro ext) (simp_all add: rGamma_series_def exp_cnj) also have "... \ cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros) finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI]) qed lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)" unfolding Gamma_def by (simp add: cnj_rGamma) lemma Gamma_complex_real: "z \ \ \ Gamma z \ (\ :: complex set)" and rGamma_complex_real: "z \ \ \ rGamma z \ \" by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma) lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)" using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast lemma holomorphic_rGamma [holomorphic_intros]: "rGamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma) lemma holomorphic_rGamma' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. rGamma (f x)) holomorphic_on A" proof - have "rGamma \ f holomorphic_on A" using assms by (intro holomorphic_on_compose assms holomorphic_rGamma) thus ?thesis by (simp only: o_def) qed lemma analytic_rGamma: "rGamma analytic_on A" unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_rGamma) lemma field_differentiable_Gamma: "z \ \\<^sub>\\<^sub>0 \ Gamma field_differentiable (at z within A)" using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto lemma holomorphic_Gamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma) lemma holomorphic_Gamma' [holomorphic_intros]: assumes "f holomorphic_on A" and "\x. x \ A \ f x \ \\<^sub>\\<^sub>0" shows "(\x. Gamma (f x)) holomorphic_on A" proof - have "Gamma \ f holomorphic_on A" using assms by (intro holomorphic_on_compose assms holomorphic_Gamma) auto thus ?thesis by (simp only: o_def) qed lemma analytic_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ Gamma analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_Gamma) lemma field_differentiable_ln_Gamma_complex: "z \ \\<^sub>\\<^sub>0 \ ln_Gamma field_differentiable (at (z::complex) within A)" by (rule field_differentiable_within_subset[of _ _ UNIV]) (force simp: field_differentiable_def intro!: derivative_intros)+ lemma holomorphic_ln_Gamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ ln_Gamma holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_ln_Gamma_complex) lemma analytic_ln_Gamma: "A \ \\<^sub>\\<^sub>0 = {} \ ln_Gamma analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_ln_Gamma) lemma has_field_derivative_rGamma_complex' [derivative_intros]: "(rGamma has_field_derivative (if z \ \\<^sub>\\<^sub>0 then (-1)^(nat \-Re z\) * fact (nat \-Re z\) else -rGamma z * Digamma z)) (at z within A)" using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros] lemma field_differentiable_Polygamma: fixes z :: complex shows "z \ \\<^sub>\\<^sub>0 \ Polygamma n field_differentiable (at z within A)" using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto lemma holomorphic_on_Polygamma [holomorphic_intros]: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n holomorphic_on A" unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma) lemma analytic_on_Polygamma: "A \ \\<^sub>\\<^sub>0 = {} \ Polygamma n analytic_on A" by (rule analytic_on_subset[of _ "UNIV - \\<^sub>\\<^sub>0"], subst analytic_on_open) (auto intro!: holomorphic_on_Polygamma) subsection\<^marker>\tag unimportant\ \The real Gamma function\ lemma rGamma_series_real: "eventually (\n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially" using eventually_gt_at_top[of "0 :: nat"] proof eventually_elim fix n :: nat assume n: "n > 0" have "Re (rGamma_series (of_real x) n) = Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))" using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real) also from n have "\ = Re (of_real ((pochhammer x (Suc n)) / (fact n * (exp (x * ln (real_of_nat n))))))" by (subst exp_of_real) simp also from n have "\ = rGamma_series x n" by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def) finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" .. qed instantiation\<^marker>\tag unimportant\ real :: Gamma begin definition "rGamma_real x = Re (rGamma (of_real x :: complex))" instance proof fix x :: real have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def) also have "of_real \ = rGamma (of_real x :: complex)" by (intro of_real_Re rGamma_complex_real) simp_all also have "\ = 0 \ x \ \\<^sub>\\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff) also have "\ \ (\n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases') finally show "(rGamma x) = 0 \ (\n. x = - real_of_nat n)" by simp next fix x :: real assume "\n. x \ - of_nat n" hence x: "complex_of_real x \ \\<^sub>\\<^sub>0" by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases') then have "x \ 0" by auto with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)" by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real rGamma_real_def [abs_def]) thus "let d = (THE d. (\n. \k d) - euler_mascheroni *\<^sub>R 1 in (\y. (rGamma y - rGamma x + rGamma x * d * (y - x)) /\<^sub>R norm (y - x)) \x\ 0" by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def] netlimit_at of_real_def[symmetric] suminf_def) next fix n :: nat have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))" by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real rGamma_real_def [abs_def]) thus "let x = - of_nat n in (\y. (rGamma y - rGamma x - (- 1) ^ n * prod of_nat {1..n} * (y - x)) /\<^sub>R norm (y - x)) \x::real\ 0" by (simp add: has_field_derivative_def has_derivative_def fact_prod netlimit_at Let_def) next fix x :: real have "rGamma_series x \ rGamma x" proof (rule Lim_transform_eventually) show "(\n. Re (rGamma_series (of_real x) n)) \ rGamma x" unfolding rGamma_real_def by (intro tendsto_intros) qed (insert rGamma_series_real, simp add: eq_commute) thus "let fact' = \n. prod of_nat {1..n}; exp = \x. THE e. (\n. \kR fact k) \ e; pochhammer' = \a n. \n = 0..n. a + of_nat n in (\n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \ rGamma x" by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost) qed end lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)" unfolding rGamma_real_def using rGamma_complex_real by simp lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)" unfolding Gamma_def by (simp add: rGamma_complex_of_real) lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))" by (rule sym, rule limI, rule tendsto_intros) lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))" by (rule sym, rule limI, rule tendsto_intros) lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))" using rGamma_complex_real[OF Reals_of_real[of x]] by (elim Reals_cases) (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real) lemma ln_Gamma_series_complex_of_real: "x > 0 \ n > 0 \ ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)" proof - assume xn: "x > 0" "n > 0" have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \ 1" for k using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps) with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real) qed lemma ln_Gamma_real_converges: assumes "(x::real) > 0" shows "convergent (ln_Gamma_series x)" proof - have "(\n. ln_Gamma_series (complex_of_real x) n) \ ln_Gamma (of_real x)" using assms by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff) moreover from eventually_gt_at_top[of "0::nat"] have "eventually (\n. complex_of_real (ln_Gamma_series x n) = ln_Gamma_series (complex_of_real x) n) sequentially" by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms) ultimately have "(\n. complex_of_real (ln_Gamma_series x n)) \ ln_Gamma (of_real x)" by (subst tendsto_cong) assumption+ from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def) qed lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \ ln_Gamma_series x \ ln_Gamma x" using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff) lemma ln_Gamma_complex_of_real: "x > 0 \ ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)" proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually) assume x: "x > 0" show "eventually (\n. of_real (ln_Gamma_series x n) = ln_Gamma_series (complex_of_real x) n) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_Gamma_series_complex_of_real x) qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def) lemma Gamma_real_pos_exp: "x > (0 :: real) \ Gamma x = exp (ln_Gamma x)" by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff ln_Gamma_complex_of_real exp_of_real) lemma ln_Gamma_real_pos: "x > 0 \ ln_Gamma x = ln (Gamma x :: real)" unfolding Gamma_real_pos_exp by simp lemma ln_Gamma_complex_conv_fact: "n > 0 \ ln_Gamma (of_nat n :: complex) = ln (fact (n - 1))" using ln_Gamma_complex_of_real[of "real n"] Gamma_fact[of "n - 1", where 'a = real] by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric]) lemma ln_Gamma_real_conv_fact: "n > 0 \ ln_Gamma (real n) = ln (fact (n - 1))" using Gamma_fact[of "n - 1", where 'a = real] by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric]) lemma Gamma_real_pos [simp, intro]: "x > (0::real) \ Gamma x > 0" by (simp add: Gamma_real_pos_exp) lemma Gamma_real_nonneg [simp, intro]: "x > (0::real) \ Gamma x \ 0" by (simp add: Gamma_real_pos_exp) lemma has_field_derivative_ln_Gamma_real [derivative_intros]: assumes x: "x > (0::real)" shows "(ln_Gamma has_field_derivative Digamma x) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms show "((Re \ ln_Gamma \ complex_of_real) has_field_derivative Digamma x) (at x)" by (auto intro!: derivative_eq_intros has_vector_derivative_real_field simp: Polygamma_of_real o_def) from eventually_nhds_in_nhd[of x "{0<..}"] assms show "eventually (\y. ln_Gamma y = (Re \ ln_Gamma \ of_real) y) (nhds x)" by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open) qed lemma field_differentiable_ln_Gamma_real: "x > 0 \ ln_Gamma field_differentiable (at (x::real) within A)" by (rule field_differentiable_within_subset[of _ _ UNIV]) (auto simp: field_differentiable_def intro!: derivative_intros)+ declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros] lemma deriv_ln_Gamma_real: assumes "z > 0" shows "deriv ln_Gamma z = Digamma (z :: real)" by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_real assms) lemma has_field_derivative_rGamma_real' [derivative_intros]: "(rGamma has_field_derivative (if x \ \\<^sub>\\<^sub>0 then (-1)^(nat \-x\) * fact (nat \-x\) else -rGamma x * Digamma x)) (at x within A)" using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases') declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros] lemma Polygamma_real_odd_pos: assumes "(x::real) \ \\<^sub>\\<^sub>0" "odd n" shows "Polygamma n x > 0" proof - from assms have "x \ 0" by auto with assms show ?thesis unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] by (auto simp: zero_less_power_eq simp del: power_Suc dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos) qed lemma Polygamma_real_even_neg: assumes "(x::real) > 0" "n > 0" "even n" shows "Polygamma n x < 0" using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"] by (auto intro!: mult_pos_pos suminf_pos) lemma Polygamma_real_strict_mono: assumes "x > 0" "x < (y::real)" "even n" shows "Polygamma n x < Polygamma n y" proof - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) then guess \ by (elim exE conjE) note \ = this note \(3) also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ > 0" by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases) finally show ?thesis by simp qed lemma Polygamma_real_strict_antimono: assumes "x > 0" "x < (y::real)" "odd n" shows "Polygamma n x > Polygamma n y" proof - have "\\. x < \ \ \ < y \ Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) then guess \ by (elim exE conjE) note \ = this note \(3) also from \(1,2) assms have "(y - x) * Polygamma (Suc n) \ < 0" by (intro mult_pos_neg Polygamma_real_even_neg) simp_all finally show ?thesis by simp qed lemma Polygamma_real_mono: assumes "x > 0" "x \ (y::real)" "even n" shows "Polygamma n x \ Polygamma n y" using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2) by (cases "x = y") simp_all lemma Digamma_real_strict_mono: "(0::real) < x \ x < y \ Digamma x < Digamma y" by (rule Polygamma_real_strict_mono) simp_all lemma Digamma_real_mono: "(0::real) < x \ x \ y \ Digamma x \ Digamma y" by (rule Polygamma_real_mono) simp_all lemma Digamma_real_ge_three_halves_pos: assumes "x \ 3/2" shows "Digamma (x :: real) > 0" proof - have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos) also from assms have "\ \ Digamma x" by (intro Polygamma_real_mono) simp_all finally show ?thesis . qed lemma ln_Gamma_real_strict_mono: assumes "x \ 3/2" "x < y" shows "ln_Gamma (x :: real) < ln_Gamma y" proof - have "\\. x < \ \ \ < y \ ln_Gamma y - ln_Gamma x = (y - x) * Digamma \" using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases) then guess \ by (elim exE conjE) note \ = this note \(3) also from \(1,2) assms have "(y - x) * Digamma \ > 0" by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all finally show ?thesis by simp qed lemma Gamma_real_strict_mono: assumes "x \ 3/2" "x < y" shows "Gamma (x :: real) < Gamma y" proof - from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp also have "\ < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms) also from Gamma_real_pos_exp[of y] assms have "\ = Gamma y" by simp finally show ?thesis . qed theorem log_convex_Gamma_real: "convex_on {0<..} (ln \ Gamma :: real \ real)" by (rule convex_on_realI[of _ _ Digamma]) (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases') subsection \The uniqueness of the real Gamma function\ text \ The following is a proof of the Bohr--Mollerup theorem, which states that any log-convex function \G\ on the positive reals that fulfils \G(1) = 1\ and satisfies the functional equation \G(x + 1) = x G(x)\ must be equal to the Gamma function. In principle, if \G\ is a holomorphic complex function, one could then extend this from the positive reals to the entire complex plane (minus the non-positive integers, where the Gamma function is not defined). \ context\<^marker>\tag unimportant\ fixes G :: "real \ real" assumes G_1: "G 1 = 1" assumes G_plus1: "x > 0 \ G (x + 1) = x * G x" assumes G_pos: "x > 0 \ G x > 0" assumes log_convex_G: "convex_on {0<..} (ln \ G)" begin private lemma G_fact: "G (of_nat n + 1) = fact n" using G_plus1[of "real n + 1" for n] by (induction n) (simp_all add: G_1 G_plus1) private definition S :: "real \ real \ real" where "S x y = (ln (G y) - ln (G x)) / (y - x)" private lemma S_eq: "n \ 2 \ S (of_nat n) (of_nat n + x) = (ln (G (real n + x)) - ln (fact (n - 1))) / x" by (subst G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff) private lemma G_lower: assumes x: "x > 0" and n: "n \ 1" shows "Gamma_series x n \ G x" proof - have "(ln \ G) (real (Suc n)) \ ((ln \ G) (real (Suc n) + x) - (ln \ G) (real (Suc n) - 1)) / (real (Suc n) + x - (real (Suc n) - 1)) * (real (Suc n) - (real (Suc n) - 1)) + (ln \ G) (real (Suc n) - 1)" using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto hence "S (of_nat n) (of_nat (Suc n)) \ S (of_nat (Suc n)) (of_nat (Suc n) + x)" unfolding S_def using x by (simp add: field_simps) also have "S (of_nat n) (of_nat (Suc n)) = ln (fact n) - ln (fact (n-1))" unfolding S_def using n by (subst (1 2) G_fact [symmetric]) (simp_all add: add_ac of_nat_diff) also have "\ = ln (fact n / fact (n-1))" by (subst ln_div) simp_all also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all finally have "x * ln (real n) + ln (fact n) \ ln (G (real (Suc n) + x))" using x n by (subst (asm) S_eq) (simp_all add: field_simps) also have "x * ln (real n) + ln (fact n) = ln (exp (x * ln (real n)) * fact n)" using x by (simp add: ln_mult) finally have "exp (x * ln (real n)) * fact n \ G (real (Suc n) + x)" using x by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos) also have "G (real (Suc n) + x) = pochhammer x (Suc n) * G x" using G_plus1[of "real (Suc n) + x" for n] G_plus1[of x] x by (induction n) (simp_all add: pochhammer_Suc add_ac) finally show "Gamma_series x n \ G x" using x by (simp add: field_simps pochhammer_pos Gamma_series_def) qed private lemma G_upper: assumes x: "x > 0" "x \ 1" and n: "n \ 2" shows "G x \ Gamma_series x n * (1 + x / real n)" proof - have "(ln \ G) (real n + x) \ ((ln \ G) (real n + 1) - (ln \ G) (real n)) / (real n + 1 - (real n)) * ((real n + x) - real n) + (ln \ G) (real n)" using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto hence "S (of_nat n) (of_nat n + x) \ S (of_nat n) (of_nat n + 1)" unfolding S_def using x by (simp add: field_simps) also from n have "S (of_nat n) (of_nat n + 1) = ln (fact n) - ln (fact (n-1))" by (subst (1 2) G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff) also have "\ = ln (fact n / (fact (n-1)))" using n by (subst ln_div) simp_all also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all finally have "ln (G (real n + x)) \ x * ln (real n) + ln (fact (n - 1))" using x n by (subst (asm) S_eq) (simp_all add: field_simps) also have "\ = ln (exp (x * ln (real n)) * fact (n - 1))" using x by (simp add: ln_mult) finally have "G (real n + x) \ exp (x * ln (real n)) * fact (n - 1)" using x by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos) also have "G (real n + x) = pochhammer x n * G x" using G_plus1[of "real n + x" for n] x by (induction n) (simp_all add: pochhammer_Suc add_ac) finally have "G x \ exp (x * ln (real n)) * fact (n- 1) / pochhammer x n" using x by (simp add: field_simps pochhammer_pos) also from n have "fact (n - 1) = fact n / n" by (cases n) simp_all also have "exp (x * ln (real n)) * \ / pochhammer x n = Gamma_series x n * (1 + x / real n)" using n x by (simp add: Gamma_series_def divide_simps pochhammer_Suc) finally show ?thesis . qed private lemma G_eq_Gamma_aux: assumes x: "x > 0" "x \ 1" shows "G x = Gamma x" proof (rule antisym) show "G x \ Gamma x" proof (rule tendsto_upperbound) from G_lower[of x] show "eventually (\n. Gamma_series x n \ G x) sequentially" using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "1::nat"]]) qed (simp_all add: Gamma_series_LIMSEQ) next show "G x \ Gamma x" proof (rule tendsto_lowerbound) have "(\n. Gamma_series x n * (1 + x / real n)) \ Gamma x * (1 + 0)" by (rule tendsto_intros real_tendsto_divide_at_top Gamma_series_LIMSEQ filterlim_real_sequentially)+ thus "(\n. Gamma_series x n * (1 + x / real n)) \ Gamma x" by simp next from G_upper[of x] show "eventually (\n. Gamma_series x n * (1 + x / real n) \ G x) sequentially" using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "2::nat"]]) qed simp_all qed theorem Gamma_pos_real_unique: assumes x: "x > 0" shows "G x = Gamma x" proof - have G_eq: "G (real n + x) = Gamma (real n + x)" if "x \ {0<..1}" for n x using that proof (induction n) case (Suc n) from Suc have "x + real n > 0" by simp hence "x + real n \ \\<^sub>\\<^sub>0" by auto with Suc show ?case using G_plus1[of "real n + x"] Gamma_plus1[of "real n + x"] by (auto simp: add_ac) qed (simp_all add: G_eq_Gamma_aux) show ?thesis proof (cases "frac x = 0") case True hence "x = of_int (floor x)" by (simp add: frac_def) with x have x_eq: "x = of_nat (nat (floor x) - 1) + 1" by simp show ?thesis by (subst (1 2) x_eq, rule G_eq) simp_all next case False from assms have x_eq: "x = of_nat (nat (floor x)) + frac x" by (simp add: frac_def) have frac_le_1: "frac x \ 1" unfolding frac_def by linarith show ?thesis by (subst (1 2) x_eq, rule G_eq, insert False frac_le_1) simp_all qed qed end subsection \The Beta function\ definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)" lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)" by (simp add: inverse_eq_divide Beta_def Gamma_def) lemma Beta_commute: "Beta a b = Beta b a" unfolding Beta_def by (simp add: ac_simps) lemma has_field_derivative_Beta1 [derivative_intros]: assumes "x \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" shows "((\x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y)))) (at x within A)" unfolding Beta_altdef by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps) lemma Beta_pole1: "x \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma Beta_pole2: "y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma Beta_zero: "x + y \ \\<^sub>\\<^sub>0 \ Beta x y = 0" by (auto simp add: Beta_def elim!: nonpos_Ints_cases') lemma has_field_derivative_Beta2 [derivative_intros]: assumes "y \ \\<^sub>\\<^sub>0" "x + y \ \\<^sub>\\<^sub>0" shows "((\y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y)))) (at y within A)" using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac) theorem Beta_plus1_plus1: assumes "x \ \\<^sub>\\<^sub>0" "y \ \\<^sub>\\<^sub>0" shows "Beta (x + 1) y + Beta x (y + 1) = Beta x y" proof - have "Beta (x + 1) y + Beta x (y + 1) = (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)" by (simp add: Beta_altdef add_divide_distrib algebra_simps) also have "\ = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))" by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps) also from assms have "\ = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp finally show ?thesis . qed theorem Beta_plus1_left: assumes "x \ \\<^sub>\\<^sub>0" shows "(x + y) * Beta (x + 1) y = x * Beta x y" proof - have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))" unfolding Beta_altdef by (simp only: ac_simps) also have "\ = x * Beta x y" unfolding Beta_altdef by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps) finally show ?thesis . qed theorem Beta_plus1_right: assumes "y \ \\<^sub>\\<^sub>0" shows "(x + y) * Beta x (y + 1) = y * Beta x y" using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute) lemma Gamma_Gamma_Beta: assumes "x + y \ \\<^sub>\\<^sub>0" shows "Gamma x * Gamma y = Beta x y * Gamma (x + y)" unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"] by (simp add: rGamma_inverse_Gamma) subsection \Legendre duplication theorem\ context begin private lemma Gamma_legendre_duplication_aux: fixes z :: "'a :: Gamma" assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)" proof - let ?powr = "\b a. exp (a * of_real (ln (of_nat b)))" let ?h = "\n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) * exp (1/2 * of_real (ln (real_of_nat n)))" { fix z :: 'a assume z: "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" let ?g = "\n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n / Gamma_series' (2*z) (2*n)" have "eventually (\n. ?g n = ?h n) sequentially" using eventually_gt_at_top proof eventually_elim fix n :: nat assume n: "n > 0" let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a" have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) / (pochhammer z n * pochhammer (z + 1/2) n)" by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac) have B: "Gamma_series' (2*z) (2*n) = ?f' * ?powr 2 (2*z) * ?powr n (2*z) / (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double) from z have "pochhammer z n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) moreover from z have "pochhammer (z + 1/2) n \ 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int) ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) = ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))" using n unfolding A B by (simp add: field_split_simps exp_minus) also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)" by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib) finally show "?g n = ?h n" by (simp only: mult_ac) qed moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \ \\<^sub>\\<^sub>0" by auto hence "?g \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "(*)2" "2*z"] by (intro tendsto_intros Gamma_series'_LIMSEQ) (simp_all add: o_def strict_mono_def Gamma_eq_zero_iff) ultimately have "?h \ ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)" by (blast intro: Lim_transform_eventually) } note lim = this from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \ \\<^sub>\\<^sub>0" by auto from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \ \\<^sub>\\<^sub>0" by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all with lim[of "1/2 :: 'a"] have "?h \ 2 * Gamma (1/2 :: 'a)" by (simp add: exp_of_real) from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis by (simp add: field_split_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps) qed -theorem Gamma_reflection_complex: +text \ + The following lemma is somewhat annoying. With a little bit of complex analysis + (Cauchy's integral theorem, to be exact), this would be completely trivial. However, + we want to avoid depending on the complex analysis session at this point, so we prove it + the hard way. +\ +private lemma Gamma_reflection_aux: + defines "h \ \z::complex. if z \ \ then 0 else + (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" + defines "a \ complex_of_real pi" + obtains h' where "continuous_on UNIV h'" "\z. (h has_field_derivative (h' z)) (at z)" +proof - + define f where "f n = a * of_real (cos_coeff (n+1) - sin_coeff (n+2))" for n + define F where "F z = (if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z)" for z + define g where "g n = complex_of_real (sin_coeff (n+1))" for n + define G where "G z = (if z = 0 then 1 else sin (a*z)/(a*z))" for z + have a_nz: "a \ 0" unfolding a_def by simp + + have "(\n. f n * (a*z)^n) sums (F z) \ (\n. g n * (a*z)^n) sums (G z)" + if "abs (Re z) < 1" for z + proof (cases "z = 0"; rule conjI) + assume "z \ 0" + note z = this that + + from z have sin_nz: "sin (a*z) \ 0" unfolding a_def by (auto simp: sin_eq_0) + have "(\n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"] + by (simp add: scaleR_conv_of_real) + from sums_split_initial_segment[OF this, of 1] + have "(\n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac) + from sums_mult[OF this, of "inverse (a*z)"] z a_nz + have A: "(\n. g n * (a*z)^n) sums (sin (a*z)/(a*z))" + by (simp add: field_simps g_def) + with z show "(\n. g n * (a*z)^n) sums (G z)" by (simp add: G_def) + from A z a_nz sin_nz have g_nz: "(\n. g n * (a*z)^n) \ 0" by (simp add: sums_iff g_def) + + have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def) + from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1] + have "(\n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))" + by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def) + from sums_mult[OF this, of "inverse z"] z assms + show "(\n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def) + next + assume z: "z = 0" + have "(\n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp + with z show "(\n. f n * (a * z) ^ n) sums (F z)" + by (simp add: f_def F_def sin_coeff_def cos_coeff_def) + have "(\n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp + with z show "(\n. g n * (a * z) ^ n) sums (G z)" + by (simp add: g_def G_def sin_coeff_def cos_coeff_def) + qed + note sums = conjunct1[OF this] conjunct2[OF this] + + define h2 where [abs_def]: + "h2 z = (\n. f n * (a*z)^n) / (\n. g n * (a*z)^n) + Digamma (1 + z) - Digamma (1 - z)" for z + define POWSER where [abs_def]: "POWSER f z = (\n. f n * (z^n :: complex))" for f z + define POWSER' where [abs_def]: "POWSER' f z = (\n. diffs f n * (z^n))" for f and z :: complex + define h2' where [abs_def]: + "h2' z = a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / + (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" for z + + have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t + proof - + from that have t: "t \ \ \ t = 0" by (auto elim!: Ints_cases simp: dist_0_norm) + hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)" + unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def) + also have "a*cot (a*t) - 1/t = (F t) / (G t)" + using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def) + also have "\ = (\n. f n * (a*t)^n) / (\n. g n * (a*t)^n)" + using sums[of t] that by (simp add: sums_iff dist_0_norm) + finally show "h t = h2 t" by (simp only: h2_def) + qed + + let ?A = "{z. abs (Re z) < 1}" + have "open ({z. Re z < 1} \ {z. Re z > -1})" + using open_halfspace_Re_gt open_halfspace_Re_lt by auto + also have "({z. Re z < 1} \ {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto + finally have open_A: "open ?A" . + hence [simp]: "interior ?A = ?A" by (simp add: interior_open) + + have summable_f: "summable (\n. f n * z^n)" for z + by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) + (simp_all add: norm_mult a_def del: of_real_add) + have summable_g: "summable (\n. g n * z^n)" for z + by (rule powser_inside, rule sums_summable, rule sums[of "\ * of_real (norm z + 1) / a"]) + (simp_all add: norm_mult a_def del: of_real_add) + have summable_fg': "summable (\n. diffs f n * z^n)" "summable (\n. diffs g n * z^n)" for z + by (intro termdiff_converges_all summable_f summable_g)+ + have "(POWSER f has_field_derivative (POWSER' f z)) (at z)" + "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z + unfolding POWSER_def POWSER'_def + by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+ + note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def] + have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z" + for z unfolding POWSER_def POWSER'_def + by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+ + note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def] + + { + fix z :: complex assume z: "abs (Re z) < 1" + define d where "d = \ * of_real (norm z + 1)" + have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add) + have "eventually (\z. h z = h2 z) (nhds z)" + using eventually_nhds_in_nhd[of z ?A] using h_eq z + by (auto elim!: eventually_mono simp: dist_0_norm) + + moreover from sums(2)[OF z] z have nz: "(\n. g n * (a * z) ^ n) \ 0" + unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def) + have A: "z \ \ \ z = 0" using z by (auto elim!: Ints_cases) + have no_int: "1 + z \ \ \ z = 0" using z Ints_diff[of "1+z" 1] A + by (auto elim!: nonpos_Ints_cases) + have no_int': "1 - z \ \ \ z = 0" using z Ints_diff[of 1 "1-z"] A + by (auto elim!: nonpos_Ints_cases) + from no_int no_int' have no_int: "1 - z \ \\<^sub>\\<^sub>0" "1 + z \ \\<^sub>\\<^sub>0" by auto + have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def + by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+) + (auto simp: h2'_def POWSER_def field_simps power2_eq_square) + ultimately have deriv: "(h has_field_derivative h2' z) (at z)" + by (subst DERIV_cong_ev[OF refl _ refl]) + + from sums(2)[OF z] z have "(\n. g n * (a * z) ^ n) \ 0" + unfolding G_def by (auto simp: sums_iff a_def sin_eq_0) + hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def + by (intro continuous_intros cont + continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto + note deriv and this + } note A = this + + interpret h: periodic_fun_simple' h + proof + fix z :: complex + show "h (z + 1) = h z" + proof (cases "z \ \") + assume z: "z \ \" + hence A: "z + 1 \ \" "z \ 0" using Ints_diff[of "z+1" 1] by auto + hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)" + by (subst (1 2) Digamma_plus1) simp_all + with A z show "h (z + 1) = h z" + by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def) + qed (simp add: h_def) + qed + + have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z + proof - + have "((\z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)" + by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) + (insert z, auto intro!: derivative_eq_intros) + hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1) + moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all + ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique) + qed + + define h2'' where "h2'' z = h2' (z - of_int \Re z\)" for z + have deriv: "(h has_field_derivative h2'' z) (at z)" for z + proof - + fix z :: complex + have B: "\Re z - real_of_int \Re z\\ < 1" by linarith + have "((\t. h (t - of_int \Re z\)) has_field_derivative h2'' z) (at z)" + unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)]) + (insert B, auto intro!: derivative_intros) + thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int) + qed + + have cont: "continuous_on UNIV h2''" + proof (intro continuous_at_imp_continuous_on ballI) + fix z :: complex + define r where "r = \Re z\" + define A where "A = {t. of_int r - 1 < Re t \ Re t < of_int r + 1}" + have "continuous_on A (\t. h2' (t - of_int r))" unfolding A_def + by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros) + (simp_all add: abs_real_def) + moreover have "h2'' t = h2' (t - of_int r)" if t: "t \ A" for t + proof (cases "Re t \ of_int r") + case True + from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) + with True have "\Re t\ = \Re z\" unfolding r_def by linarith + thus ?thesis by (auto simp: r_def h2''_def) + next + case False + from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) + with False have t': "\Re t\ = \Re z\ - 1" unfolding r_def by linarith + moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)" + by (intro h2'_eq) simp_all + ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t') + qed + ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl]) + moreover { + have "open ({t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t})" + by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt) + also have "{t. of_int r - 1 < Re t} \ {t. of_int r + 1 > Re t} = A" + unfolding A_def by blast + finally have "open A" . + } + ultimately have C: "isCont h2'' t" if "t \ A" for t using that + by (subst (asm) continuous_on_eq_continuous_at) auto + have "of_int r - 1 < Re z" "Re z < of_int r + 1" unfolding r_def by linarith+ + thus "isCont h2'' z" by (intro C) (simp_all add: A_def) + qed + + from that[OF cont deriv] show ?thesis . +qed + +lemma Gamma_reflection_complex: fixes z :: complex shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)" proof - let ?g = "\z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)" define g where [abs_def]: "g z = (if z \ \ then of_real pi else ?g z)" for z :: complex let ?h = "\z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))" define h where [abs_def]: "h z = (if z \ \ then 0 else ?h z)" for z :: complex - \ \\<^term>\g\ is periodic with period 1.\ + \ \@{term g} is periodic with period 1.\ interpret g: periodic_fun_simple' g proof fix z :: complex show "g (z + 1) = g z" proof (cases "z \ \") case False hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def) also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)" using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints by (subst Beta_plus1_left [symmetric]) auto also have "\ * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))" using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi) also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)" using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def) finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto qed (simp add: g_def) qed - \ \\<^term>\g\ is entire.\ - have g_g' [derivative_intros]: "(g has_field_derivative (h z * g z)) (at z)" for z :: complex + \ \@{term g} is entire.\ + have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex proof (cases "z \ \") let ?h' = "\z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) + of_real pi * cos (z * of_real pi))" case False from False have "eventually (\t. t \ UNIV - \) (nhds z)" by (intro eventually_nhds_in_open) (auto simp: open_Diff) hence "eventually (\t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def) moreover { from False Ints_diff[of 1 "1-z"] have "1 - z \ \" by auto hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def) also from False have "sin (of_real pi * z) \ 0" by (subst sin_eq_0) auto hence "?h' z = h z * g z" using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def) finally have "(?g has_field_derivative (h z * g z)) (at z)" . } ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl]) next case True then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases) let ?t = "(\z::complex. if z = 0 then 1 else sin z / z) \ (\z. of_real pi * z)" have deriv_0: "(g has_field_derivative 0) (at 0)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "eventually (\z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)" using eventually_nhds_ball[OF zero_less_one, of "0::complex"] proof eventually_elim fix z :: complex assume z: "z \ ball 0 1" show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z" proof (cases "z = 0") assume z': "z \ 0" with z have z'': "z \ \\<^sub>\\<^sub>0" "z \ \" by (auto elim!: Ints_cases simp: dist_0_norm) from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp with z'' z' show ?thesis by (simp add: g_def ac_simps) qed (simp add: g_def) qed have "(?t has_field_derivative (0 * of_real pi)) (at 0)" using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"] by (intro DERIV_chain) simp_all thus "((\z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)" by (auto intro!: derivative_eq_intros simp: o_def) qed have "((g \ (\x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))" using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros) also have "g \ (\x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int) finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def) qed - have g_holo [holomorphic_intros]: "g holomorphic_on A" for A - by (rule holomorphic_on_subset[of _ UNIV]) - (force simp: holomorphic_on_open intro!: derivative_intros)+ - have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z proof (cases "z \ \") case True with that have "z = 0 \ z = 1" by (force elim!: Ints_cases) moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square) moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square Beta_def algebra_simps) ultimately show ?thesis by force next case False hence z: "z/2 \ \" "(z+1)/2 \ \" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases) hence z': "z/2 \ \\<^sub>\\<^sub>0" "(z+1)/2 \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) from z have "1-z/2 \ \" "1-((z+1)/2) \ \" using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto hence z'': "1-z/2 \ \\<^sub>\\<^sub>0" "1-((z+1)/2) \ \\<^sub>\\<^sub>0" by (auto elim!: nonpos_Ints_cases) from z have "g (z/2) * g ((z+1)/2) = (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) * (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))" by (simp add: g_def) also from z' Gamma_legendre_duplication_aux[of "z/2"] have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z" by (simp add: add_divide_distrib) also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"] have "Gamma (1-z/2) * Gamma (1-(z+1)/2) = Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))" by (simp add: add_divide_distrib ac_simps) finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) * (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))" by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real) also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)" using cos_sin_eq[of "- of_real pi * z/2", symmetric] by (simp add: ring_distribs add_divide_distrib ac_simps) also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)" by (subst sin_times_cos) (simp add: field_simps) also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z" using \z \ \\ by (simp add: g_def) finally show ?thesis . qed have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z proof - define r where "r = \Re z / 2\" have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int) also have "of_int (2*r) = 2 * of_int r" by simp also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+ hence "Gamma (1/2)^2 * g (z - 2 * of_int r) = g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)" unfolding r_def by (intro g_eq[symmetric]) simp_all also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp also have "g \ = g (z/2)" by (rule g.minus_of_int) also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp also have "g \ = g ((z+1)/2)" by (rule g.minus_of_int) finally show ?thesis .. qed have g_nz [simp]: "g z \ 0" for z :: complex unfolding g_def using Ints_diff[of 1 "1 - z"] by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int) - have h_altdef: "h z = deriv g z / g z" for z :: complex - using DERIV_imp_deriv[OF g_g', of z] by simp - have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z proof - have "((\t. g (t/2) * g ((t+1)/2)) has_field_derivative (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps) hence "((\t. Gamma (1/2)^2 * g t) has_field_derivative Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)" by (subst (1 2) g_eq[symmetric]) simp from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"] have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)" using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints) moreover have "(g has_field_derivative (g z * h z)) (at z)" using g_g'[of z] by (simp add: ac_simps) ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)" by (intro DERIV_unique) thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp qed - have h_holo [holomorphic_intros]: "h holomorphic_on A" for A - unfolding h_altdef [abs_def] - by (rule holomorphic_on_subset[of _ UNIV]) (auto intro!: holomorphic_intros) - define h' where "h' = deriv h" - have h_h': "(h has_field_derivative h' z) (at z)" for z unfolding h'_def - by (auto intro!: holomorphic_derivI[of _ UNIV] holomorphic_intros) - have h'_holo [holomorphic_intros]: "h' holomorphic_on A" for A unfolding h'_def - by (rule holomorphic_on_subset[of _ UNIV]) (auto intro!: holomorphic_intros) - have h'_cont: "continuous_on UNIV h'" - by (intro holomorphic_on_imp_continuous_on holomorphic_intros) + obtain h' where h'_cont: "continuous_on UNIV h'" and + h_h': "\z. (h has_field_derivative h' z) (at z)" + unfolding h_def by (erule Gamma_reflection_aux) have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z proof - have "((\t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)" by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2]) hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)" by (subst (asm) h_eq[symmetric]) from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique) qed have h'_zero: "h' z = 0" for z proof - define m where "m = max 1 \Re z\" define B where "B = {t. abs (Re t) \ m \ abs (Im t) \ abs (Im z)}" have "closed ({t. Re t \ -m} \ {t. Re t \ m} \ {t. Im t \ -\Im z\} \ {t. Im t \ \Im z\})" (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le closed_halfspace_Im_ge closed_halfspace_Im_le) also have "?B = B" unfolding B_def by fastforce finally have "closed B" . moreover have "bounded B" unfolding bounded_iff proof (intro ballI exI) fix t assume t: "t \ B" have "norm t \ \Re t\ + \Im t\" by (rule cmod_le) also from t have "\Re t\ \ m" unfolding B_def by blast also from t have "\Im t\ \ \Im z\" unfolding B_def by blast finally show "norm t \ m + \Im z\" by - simp qed ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast define M where "M = (SUP z\B. norm (h' z))" have "compact (h' ` B)" by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+ hence bdd: "bdd_above ((\z. norm (h' z)) ` B)" using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded) have "norm (h' z) \ M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def) also have "M \ M/2" proof (subst M_def, subst cSUP_le_iff) have "z \ B" unfolding B_def m_def by simp thus "B \ {}" by auto next show "\z\B. norm (h' z) \ M/2" proof fix t :: complex assume t: "t \ B" from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm) also have "norm \ = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp also have "norm (h' (t/2) + h' ((t+1)/2)) \ norm (h' (t/2)) + norm (h' ((t+1)/2))" by (rule norm_triangle_ineq) also from t have "abs (Re ((t + 1)/2)) \ m" unfolding m_def B_def by auto with t have "t/2 \ B" "(t+1)/2 \ B" unfolding B_def by auto hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \ M + M" unfolding M_def by (intro add_mono cSUP_upper bdd) (auto simp: B_def) also have "(M + M) / 4 = M / 2" by simp finally show "norm (h' t) \ M/2" by - simp_all qed qed (insert bdd, auto simp: cball_eq_empty) hence "M \ 0" by simp finally show "h' z = 0" by simp qed have h_h'_2: "(h has_field_derivative 0) (at z)" for z using h_h'[of z] h'_zero[of z] by simp have g_real: "g z \ \" if "z \ \" for z unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real) have h_real: "h z \ \" if "z \ \" for z unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real) + have g_nz: "g z \ 0" for z unfolding g_def using Ints_diff[of 1 "1-z"] + by (auto simp: Gamma_eq_zero_iff sin_eq_0) from h'_zero h_h'_2 have "\c. \z\UNIV. h z = c" - by (intro has_field_derivative_zero_constant) simp_all + by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm) then obtain c where c: "\z. h z = c" by auto have "\u. u \ closed_segment 0 1 \ Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))" by (intro complex_mvt_line g_g') - then guess u by (elim exE conjE) note u = this + then obtain u where u: "u \ closed_segment 0 1" "Re (g 1) - Re (g 0) = Re (h u * g u)" + by auto from u(1) have u': "u \ \" unfolding closed_segment_def by (auto simp: scaleR_conv_of_real) from u' g_real[of u] g_nz[of u] have "Re (g u) \ 0" by (auto elim!: Reals_cases) with u(2) c[of u] g_real[of u] g_nz[of u] u' have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1) with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases) with c have A: "h z * g z = 0" for z by simp hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp hence "\c'. \z\UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all then obtain c' where c: "\z. g z = c'" by (force simp: dist_0_norm) from this[of 0] have "c' = pi" unfolding g_def by simp with c have "g z = pi" by simp show ?thesis proof (cases "z \ \") case False - with \g z = pi\ show ?thesis by (auto simp: g_def field_split_simps) + with \g z = pi\ show ?thesis by (auto simp: g_def divide_simps) next case True then obtain n where n: "z = of_int n" by (elim Ints_cases) with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force moreover have "of_int (1 - n) \ \\<^sub>\\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp ultimately show ?thesis using n by (cases "n \ 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int) qed qed lemma rGamma_reflection_complex: "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi" using Gamma_reflection_complex[of z] by (simp add: Gamma_def field_split_simps split: if_split_asm) lemma rGamma_reflection_complex': "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi" proof - have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))" using rGamma_plus1[of "-z", symmetric] by simp also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi" by (rule rGamma_reflection_complex) finally show ?thesis by simp qed lemma Gamma_reflection_complex': "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))" using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def field_split_simps mult_ac) lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi" proof - from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1] have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square) hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp also have "\ = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all moreover have "Gamma (1/2 :: real) \ 0" using Gamma_real_pos[of "1/2"] by simp ultimately show ?thesis by (rule real_sqrt_unique [symmetric]) qed lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)" proof - have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp also have "\ = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real) finally show ?thesis . qed theorem Gamma_legendre_duplication: fixes z :: complex assumes "z \ \\<^sub>\\<^sub>0" "z + 1/2 \ \\<^sub>\\<^sub>0" shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)" using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex) end subsection\<^marker>\tag unimportant\ \Limits and residues\ text \ The inverse of the Gamma function has simple zeros: \ lemma rGamma_zeros: "(\z. rGamma z / (z + of_nat n)) \ (- of_nat n) \ ((-1)^n * fact n :: 'a :: Gamma)" proof (subst tendsto_cong) let ?f = "\z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a" from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] show "eventually (\z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))" by (subst pochhammer_rGamma[of _ "Suc n"]) (auto elim!: eventually_mono simp: field_split_simps pochhammer_rec' eq_neg_iff_add_eq_0) have "isCont ?f (- of_nat n)" by (intro continuous_intros) thus "?f \ (- of_nat n) \ (- 1) ^ n * fact n" unfolding isCont_def by (simp add: pochhammer_same) qed text \ The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function, and their residues can easily be computed from the limit we have just proven: \ lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))" proof - from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] have "eventually (\z. rGamma z \ (0 :: 'a)) (at (- of_nat n))" by (auto elim!: eventually_mono nonpos_Ints_cases' simp: rGamma_eq_zero_iff dist_of_nat dist_minus) with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident] have "filterlim (\z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))" unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity]) (simp_all add: filterlim_at) moreover have "(\z. inverse (rGamma z) :: 'a) = Gamma" by (intro ext) (simp add: rGamma_inverse_Gamma) ultimately show ?thesis by (simp only: ) qed lemma Gamma_residues: "(\z. Gamma z * (z + of_nat n)) \ (- of_nat n) \ ((-1)^n / fact n :: 'a :: Gamma)" proof (subst tendsto_cong) let ?c = "(- 1) ^ n / fact n :: 'a" from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV] show "eventually (\z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n))) (at (- of_nat n))" by (auto elim!: eventually_mono simp: field_split_simps rGamma_inverse_Gamma) have "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ inverse ((- 1) ^ n * fact n :: 'a)" by (intro tendsto_intros rGamma_zeros) simp_all also have "inverse ((- 1) ^ n * fact n) = ?c" by (simp_all add: field_simps flip: power_mult_distrib) finally show "(\z. inverse (rGamma z / (z + of_nat n))) \ (- of_nat n) \ ?c" . qed -lemma is_pole_Gamma: "is_pole Gamma (- of_nat n)" - unfolding is_pole_def using Gamma_poles . - -lemma Gamme_residue: - "residue Gamma (- of_nat n) = (-1) ^ n / fact n" -proof (rule residue_simple') - show "open (- (\\<^sub>\\<^sub>0 - {-of_nat n}) :: complex set)" - by (intro open_Compl closed_subset_Ints) auto - show "Gamma holomorphic_on (- (\\<^sub>\\<^sub>0 - {-of_nat n}) - {- of_nat n})" - by (rule holomorphic_Gamma) auto - show "(\w. Gamma w * (w - - of_nat n)) \- of_nat n \ (- 1) ^ n / fact n" - using Gamma_residues[of n] by simp -qed auto - subsection \Alternative definitions\ subsubsection \Variant of the Euler form\ definition Gamma_series_euler' where "Gamma_series_euler' z n = inverse z * (\k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))" context begin private lemma Gamma_euler'_aux1: fixes z :: "'a :: {real_normed_field,banach}" assumes n: "n > 0" shows "exp (z * of_real (ln (of_nat n + 1))) = (\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))" proof - have "(\k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) = exp (z * of_real (\k = 1..n. ln (1 + 1 / real_of_nat k)))" by (subst exp_sum [symmetric]) (simp_all add: sum_distrib_left) also have "(\k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\k=1..n. 1 + 1 / real_of_nat k)" by (subst ln_prod [symmetric]) (auto intro!: add_pos_nonneg) also have "(\k=1..n. 1 + 1 / of_nat k :: real) = (\k=1..n. (of_nat k + 1) / of_nat k)" by (intro prod.cong) (simp_all add: field_split_simps) also have "(\k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1" by (induction n) (simp_all add: prod.nat_ivl_Suc' field_split_simps) finally show ?thesis .. qed theorem Gamma_series_euler': assumes z: "(z :: 'a :: Gamma) \ \\<^sub>\\<^sub>0" shows "(\n. Gamma_series_euler' z n) \ Gamma z" proof (rule Gamma_seriesI, rule Lim_transform_eventually) let ?f = "\n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)" let ?r = "\n. ?f n / Gamma_series z n" let ?r' = "\n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))" from z have z': "z \ 0" by auto have "eventually (\n. ?r' n = ?r n) sequentially" using z by (auto simp: field_split_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac intro: eventually_mono eventually_gt_at_top[of "0::nat"] dest: pochhammer_eq_0_imp_nonpos_Int) moreover have "?r' \ exp (z * of_real (ln 1))" by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all ultimately show "?r \ 1" by (force intro: Lim_transform_eventually) from eventually_gt_at_top[of "0::nat"] show "eventually (\n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially" proof eventually_elim fix n :: nat assume n: "n > 0" from n z' have "Gamma_series_euler' z n = exp (z * of_real (ln (of_nat n + 1))) / (z * (\k=1..n. (1 + z / of_nat k)))" by (subst Gamma_euler'_aux1) (simp_all add: Gamma_series_euler'_def prod.distrib prod_inversef[symmetric] divide_inverse) also have "(\k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n" proof (cases n) case (Suc n') then show ?thesis unfolding pochhammer_prod fact_prod by (simp add: atLeastLessThanSuc_atLeastAtMost field_simps prod_dividef prod.atLeast_Suc_atMost_Suc_shift del: prod.cl_ivl_Suc) qed auto also have "z * \ = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec) finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp qed qed end subsubsection \Weierstrass form\ definition Gamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where "Gamma_series_Weierstrass z n = exp (-euler_mascheroni * z) / z * (\k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))" definition\<^marker>\tag unimportant\ rGamma_series_Weierstrass :: "'a :: {banach,real_normed_field} \ nat \ 'a" where "rGamma_series_Weierstrass z n = exp (euler_mascheroni * z) * z * (\k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))" lemma Gamma_series_Weierstrass_nonpos_Ints: "eventually (\k. Gamma_series_Weierstrass (- of_nat n) k = 0) sequentially" using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_Weierstrass_def) lemma rGamma_series_Weierstrass_nonpos_Ints: "eventually (\k. rGamma_series_Weierstrass (- of_nat n) k = 0) sequentially" using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_Weierstrass_def) theorem Gamma_Weierstrass_complex: "Gamma_series_Weierstrass z \ Gamma (z :: complex)" proof (cases "z \ \\<^sub>\\<^sub>0") case True then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') also from True have "Gamma_series_Weierstrass \ \ Gamma z" by (simp add: tendsto_cong[OF Gamma_series_Weierstrass_nonpos_Ints] Gamma_nonpos_Int) finally show ?thesis . next case False hence z: "z \ 0" by auto let ?f = "(\x. \x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))" have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \ 1" for n :: nat using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp) have "(\n. \k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \ ln_Gamma z + euler_mascheroni * z + ln z" using ln_Gamma_series'_aux[OF False] by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def sum.shift_bounds_Suc_ivl sums_def atLeast0LessThan) from tendsto_exp[OF this] False z have "?f \ z * exp (euler_mascheroni * z) * Gamma z" by (simp add: exp_add exp_sum exp_diff mult_ac Gamma_complex_altdef A) from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z show "Gamma_series_Weierstrass z \ Gamma z" by (simp add: exp_minus field_split_simps Gamma_series_Weierstrass_def [abs_def]) qed lemma tendsto_complex_of_real_iff: "((\x. complex_of_real (f x)) \ of_real c) F = (f \ c) F" by (rule tendsto_of_real_iff) lemma Gamma_Weierstrass_real: "Gamma_series_Weierstrass x \ Gamma (x :: real)" using Gamma_Weierstrass_complex[of "of_real x"] unfolding Gamma_series_Weierstrass_def[abs_def] by (subst tendsto_complex_of_real_iff [symmetric]) (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real) lemma rGamma_Weierstrass_complex: "rGamma_series_Weierstrass z \ rGamma (z :: complex)" proof (cases "z \ \\<^sub>\\<^sub>0") case True then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases') also from True have "rGamma_series_Weierstrass \ \ rGamma z" by (simp add: tendsto_cong[OF rGamma_series_Weierstrass_nonpos_Ints] rGamma_nonpos_Int) finally show ?thesis . next case False have "rGamma_series_Weierstrass z = (\n. inverse (Gamma_series_Weierstrass z n))" by (simp add: rGamma_series_Weierstrass_def[abs_def] Gamma_series_Weierstrass_def exp_minus divide_inverse prod_inversef[symmetric] mult_ac) also from False have "\ \ inverse (Gamma z)" by (intro tendsto_intros Gamma_Weierstrass_complex) (simp add: Gamma_eq_zero_iff) finally show ?thesis by (simp add: Gamma_def) qed subsubsection \Binomial coefficient form\ lemma Gamma_gbinomial: "(\n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \ rGamma (z+1)" proof (cases "z = 0") case False show ?thesis proof (rule Lim_transform_eventually) let ?powr = "\a b. exp (b * of_real (ln (of_nat a)))" show "eventually (\n. rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially" proof (intro always_eventually allI) fix n :: nat from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n" by (simp add: gbinomial_pochhammer' pochhammer_rec) also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z" by (simp add: rGamma_series_def field_split_simps exp_minus) finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" .. qed from False have "(\n. rGamma_series z n / z) \ rGamma z / z" by (intro tendsto_intros) also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z] by (simp add: field_simps) finally show "(\n. rGamma_series z n / z) \ rGamma (z+1)" . qed qed (simp_all add: binomial_gbinomial [symmetric]) lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)" by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric]) lemma gbinomial_asymptotic: fixes z :: "'a :: Gamma" shows "(\n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \ inverse (Gamma (- z))" unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"] by (subst (asm) gbinomial_minus') (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric]) lemma fact_binomial_limit: "(\n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \ 1 / fact k" proof (rule Lim_transform_eventually) have "(\n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n)))) \ 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \ _") using Gamma_gbinomial[of "of_nat k :: 'a"] by (simp add: binomial_gbinomial add_ac Gamma_def field_split_simps exp_of_real [symmetric] exp_minus) also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact) finally show "?f \ 1 / fact k" . show "eventually (\n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially" using eventually_gt_at_top[of "0::nat"] proof eventually_elim fix n :: nat assume n: "n > 0" from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)" by (simp add: exp_of_nat_mult) thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp qed qed lemma binomial_asymptotic': "(\n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \ 1" using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp lemma gbinomial_Beta: assumes "z + 1 \ \\<^sub>\\<^sub>0" shows "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))" using assms proof (induction n arbitrary: z) case 0 hence "z + 2 \ \\<^sub>\\<^sub>0" using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute) with 0 show ?case by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute) next case (Suc n z) show ?case proof (cases "z \ \\<^sub>\\<^sub>0") case True with Suc.prems have "z = 0" by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff) show ?thesis proof (cases "n = 0") case True with \z = 0\ show ?thesis by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric]) next case False with \z = 0\ show ?thesis by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1) qed next case False have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp also have "\ = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)" by (subst gbinomial_factors) (simp add: field_simps) also from False have "\ = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))" (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1) also have "of_nat (Suc n) \ (\\<^sub>\\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)" by (subst Beta_plus1_right [symmetric]) simp_all finally show ?thesis . qed qed theorem gbinomial_Gamma: assumes "z + 1 \ \\<^sub>\\<^sub>0" shows "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))" proof - have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))" by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac) also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)" using Gamma_plus1[of "z+1"] by (auto simp add: field_split_simps mult_ac add_ac) finally show ?thesis . qed subsubsection \Integral form\ lemma integrable_on_powr_from_0': assumes a: "a > (-1::real)" and c: "c \ 0" shows "(\x. x powr a) integrable_on {0<..c}" proof - from c have *: "{0<..c} - {0..c} = {}" "{0..c} - {0<..c} = {0}" by auto show ?thesis by (rule integrable_spike_set [OF integrable_on_powr_from_0[OF a c]]) (simp_all add: *) qed lemma absolutely_integrable_Gamma_integral: assumes "Re z > 0" "a > 0" shows "(\t. complex_of_real t powr (z - 1) / of_real (exp (a * t))) absolutely_integrable_on {0<..}" (is "?f absolutely_integrable_on _") proof - have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" by (intro tendsto_intros ln_x_over_x_tendsto_0) hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp from order_tendstoD(2)[OF this, of "a/2"] and \a > 0\ have "eventually (\x. (Re z - 1) * ln x / x < a/2) at_top" by simp from eventually_conj[OF this eventually_gt_at_top[of 0]] obtain x0 where "\x\x0. (Re z - 1) * ln x / x < a/2 \ x > 0" by (auto simp: eventually_at_top_linorder) hence "x0 > 0" by simp have "x powr (Re z - 1) / exp (a * x) < exp (-(a/2) * x)" if "x \ x0" for x proof - from that and \\x\x0. _\ have x: "(Re z - 1) * ln x / x < a / 2" "x > 0" by auto have "x powr (Re z - 1) = exp ((Re z - 1) * ln x)" using \x > 0\ by (simp add: powr_def) also from x have "(Re z - 1) * ln x < (a * x) / 2" by (simp add: field_simps) finally show ?thesis by (simp add: field_simps exp_add [symmetric]) qed note x0 = \x0 > 0\ this have "?f absolutely_integrable_on ({0<..x0} \ {x0..})" proof (rule set_integrable_Un) show "?f absolutely_integrable_on {0<..x0}" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2]) show "integrable lebesgue (\x. indicat_real {0<..x0} x *\<^sub>R x powr (Re z - 1))" using x0(1) assms by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_powr_from_0') auto show "(\x. indicat_real {0<..x0} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \ borel_measurable lebesgue" by (intro measurable_completion) (auto intro!: borel_measurable_continuous_on_indicator continuous_intros) fix x :: real have "x powr (Re z - 1) / exp (a * x) \ x powr (Re z - 1) / 1" if "x \ 0" using that assms by (intro divide_left_mono) auto thus "norm (indicator {0<..x0} x *\<^sub>R ?f x) \ norm (indicator {0<..x0} x *\<^sub>R x powr (Re z - 1))" by (simp_all add: norm_divide norm_powr_real_powr indicator_def) qed next show "?f absolutely_integrable_on {x0..}" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2]) show "integrable lebesgue (\x. indicat_real {x0..} x *\<^sub>R exp (- (a / 2) * x))" using assms by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_exp_minus_to_infinity) auto show "(\x. indicat_real {x0..} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \ borel_measurable lebesgue" using x0(1) by (intro measurable_completion) (auto intro!: borel_measurable_continuous_on_indicator continuous_intros) fix x :: real show "norm (indicator {x0..} x *\<^sub>R ?f x) \ norm (indicator {x0..} x *\<^sub>R exp (-(a/2) * x))" using x0 by (auto simp: norm_divide norm_powr_real_powr indicator_def less_imp_le) qed qed auto also have "{0<..x0} \ {x0..} = {0<..}" using x0(1) by auto finally show ?thesis . qed lemma integrable_Gamma_integral_bound: fixes a c :: real assumes a: "a > -1" and c: "c \ 0" defines "f \ \x. if x \ {0..c} then x powr a else exp (-x/2)" shows "f integrable_on {0..}" proof - have "f integrable_on {0..c}" by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]]) (insert a c, simp_all add: f_def) moreover have A: "(\x. exp (-x/2)) integrable_on {c..}" using integrable_on_exp_minus_to_infinity[of "1/2"] by simp have "f integrable_on {c..}" by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def) ultimately show "f integrable_on {0..}" by (rule integrable_Un') (insert c, auto simp: max_def) qed theorem Gamma_integral_complex: assumes z: "Re z > 0" shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" proof - have A: "((\t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n)) has_integral (fact n / pochhammer z (n+1))) {0..1}" if "Re z > 0" for n z using that proof (induction n arbitrary: z) case 0 have "((\t. complex_of_real t powr (z - 1)) has_integral (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0 by (intro fundamental_theorem_of_calculus_interior) (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_field) thus ?case by simp next case (Suc n) let ?f = "\t. complex_of_real t powr z / z" let ?f' = "\t. complex_of_real t powr (z - 1)" let ?g = "\t. (1 - complex_of_real t) ^ Suc n" let ?g' = "\t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)" have "((\t. ?f' t * ?g t) has_integral (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}" (is "(_ has_integral ?I) _") proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g]) from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g" by (auto intro!: continuous_intros) next fix t :: real assume t: "t \ {0<..<1}" show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems by (auto intro!: derivative_eq_intros has_vector_derivative_real_field) show "(?g has_vector_derivative ?g' t) (at t)" by (rule has_vector_derivative_real_field derivative_eq_intros refl)+ simp_all next from Suc.prems have [simp]: "z \ 0" by auto from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp have [simp]: "z + of_nat n \ 0" "z + 1 + of_nat n \ 0" for n using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel) have "((\x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}" (is "(?A has_integral ?B) _") using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec) also have "?A = (\t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps) also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))" by (simp add: field_split_simps pochhammer_rec prod.shift_bounds_cl_Suc_ivl del: of_nat_Suc) finally show "((\t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}" by simp qed (simp_all add: bounded_bilinear_mult) thus ?case by simp qed have B: "((\t. if t \ {0..of_nat n} then of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0) has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n proof (cases "n > 0") case [simp]: True hence [simp]: "n \ 0" by auto with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0] have "((\x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n) has_integral fact n * of_nat n / pochhammer z (n+1)) ((\x. real n * x)`{0..1})" (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real) also from True have "((\x. real n*x)`{0..1}) = {0..real n}" by (subst image_mult_atLeastAtMost) simp_all also have "?f = (\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)" using True by (intro ext) (simp add: field_simps) finally have "((\x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" (is ?P) . also have "?P \ ((\x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric]) also have "\ \ ((\x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n) has_integral ?I) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div) finally have \ . note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"] have "((\x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n) has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P) by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric]) (simp add: algebra_simps) also have "?P \ ((\x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n) has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real) also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) = (of_nat n powr z * fact n / pochhammer z (n+1))" by (auto simp add: powr_def algebra_simps exp_diff exp_of_real) finally show ?thesis by (subst has_integral_restrict) simp_all next case False thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl) qed have "eventually (\n. Gamma_series z n = of_nat n powr z * fact n / pochhammer z (n+1)) sequentially" using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def) from this and Gamma_series_LIMSEQ[of z] have C: "(\k. of_nat k powr z * fact k / pochhammer z (k+1)) \ Gamma z" by (blast intro: Lim_transform_eventually) { fix x :: real assume x: "x \ 0" have lim_exp: "(\k. (1 - x / real k) ^ k) \ exp (-x)" using tendsto_exp_limit_sequentially[of "-x"] by simp have "(\k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k)) \ of_real x powr (z - 1) * of_real (exp (-x))" (is ?P) by (intro tendsto_intros lim_exp) also from eventually_gt_at_top[of "nat \x\"] have "eventually (\k. of_nat k > x) sequentially" by eventually_elim linarith hence "?P \ (\k. if x \ of_nat k then of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0) \ of_real x powr (z - 1) * of_real (exp (-x))" by (intro tendsto_cong) (auto elim!: eventually_mono) finally have \ . } hence D: "\x\{0..}. (\k. if x \ {0..real k} then of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0) \ of_real x powr (z - 1) / of_real (exp x)" by (simp add: exp_minus field_simps cong: if_cong) have "((\x. (Re z - 1) * (ln x / x)) \ (Re z - 1) * 0) at_top" by (intro tendsto_intros ln_x_over_x_tendsto_0) hence "((\x. ((Re z - 1) * ln x) / x) \ 0) at_top" by simp from order_tendstoD(2)[OF this, of "1/2"] have "eventually (\x. (Re z - 1) * ln x / x < 1/2) at_top" by simp from eventually_conj[OF this eventually_gt_at_top[of 0]] obtain x0 where "\x\x0. (Re z - 1) * ln x / x < 1/2 \ x > 0" by (auto simp: eventually_at_top_linorder) hence x0: "x0 > 0" "\x. x \ x0 \ (Re z - 1) * ln x < x / 2" by auto define h where "h = (\x. if x \ {0..x0} then x powr (Re z - 1) else exp (-x/2))" have le_h: "x powr (Re z - 1) * exp (-x) \ h x" if x: "x \ 0" for x proof (cases "x > x0") case True from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)" by (simp add: powr_def exp_diff exp_minus field_simps exp_add) also from x0(2)[of x] True have "\ < exp (-x/2)" by (simp add: field_simps) finally show ?thesis using True by (auto simp add: h_def) next case False from x have "x powr (Re z - 1) * exp (- x) \ x powr (Re z - 1) * 1" by (intro mult_left_mono) simp_all with False show ?thesis by (auto simp add: h_def) qed have E: "\x\{0..}. cmod (if x \ {0..real k} then of_real x powr (z - 1) * (1 - complex_of_real x / of_nat k) ^ k else 0) \ h x" (is "\x\_. ?f x \ _") for k proof safe fix x :: real assume x: "x \ 0" { fix x :: real and n :: nat assume x: "x \ of_nat n" have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp also have "norm \ = \(1 - x / real n)\" by (subst norm_of_real) (rule refl) also from x have "\ = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: field_split_simps) finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" . } note D = this from D[of x k] x have "?f x \ (if of_nat k \ x \ k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)" by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg) also have "\ \ x powr (Re z - 1) * exp (-x)" by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n) also from x have "\ \ h x" by (rule le_h) finally show "?f x \ h x" . qed have F: "h integrable_on {0..}" unfolding h_def by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all) show ?thesis by (rule has_integral_dominated_convergence[OF B F E D C]) qed lemma Gamma_integral_real: assumes x: "x > (0 :: real)" shows "((\t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}" proof - have A: "((\t. complex_of_real t powr (complex_of_real x - 1) / complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}" using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real) have "((\t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}" by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric]) from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def) qed lemma absolutely_integrable_Gamma_integral': assumes "Re z > 0" shows "(\t. complex_of_real t powr (z - 1) / of_real (exp t)) absolutely_integrable_on {0<..}" using absolutely_integrable_Gamma_integral [OF assms zero_less_one] by simp lemma Gamma_integral_complex': assumes z: "Re z > 0" shows "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0<..}" proof - have "((\t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}" by (rule Gamma_integral_complex) fact+ hence "((\t. if t \ {0<..} then of_real t powr (z - 1) / of_real (exp t) else 0) has_integral Gamma z) {0..}" by (rule has_integral_spike [of "{0}", rotated 2]) auto also have "?this = ?thesis" by (subst has_integral_restrict) auto finally show ?thesis . qed lemma Gamma_conv_nn_integral_real: assumes "s > (0::real)" shows "Gamma s = nn_integral lborel (\t. ennreal (indicator {0..} t * t powr (s - 1) / exp t))" using nn_integral_has_integral_lebesgue[OF _ Gamma_integral_real[OF assms]] by simp lemma integrable_Beta: assumes "a > 0" "b > (0::real)" shows "set_integrable lborel {0..1} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" proof - define C where "C = max 1 ((1/2) powr (b - 1))" define D where "D = max 1 ((1/2) powr (a - 1))" have C: "(1 - x) powr (b - 1) \ C" if "x \ {0..1/2}" for x proof (cases "b < 1") case False with that have "(1 - x) powr (b - 1) \ (1 powr (b - 1))" by (intro powr_mono2) auto thus ?thesis by (auto simp: C_def) qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 C_def) have D: "x powr (a - 1) \ D" if "x \ {1/2..1}" for x proof (cases "a < 1") case False with that have "x powr (a - 1) \ (1 powr (a - 1))" by (intro powr_mono2) auto thus ?thesis by (auto simp: D_def) next case True qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 D_def) have [simp]: "C \ 0" "D \ 0" by (simp_all add: C_def D_def) have I1: "set_integrable lborel {0..1/2} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "(\t. t powr (a - 1)) integrable_on {0..1/2}" by (rule integrable_on_powr_from_0) (use assms in auto) hence "(\t. t powr (a - 1)) absolutely_integrable_on {0..1/2}" by (subst absolutely_integrable_on_iff_nonneg) auto from integrable_mult_right[OF this [unfolded set_integrable_def], of C] show "integrable lborel (\x. indicat_real {0..1/2} x *\<^sub>R (C * x powr (a - 1)))" by (subst (asm) integrable_completion) (auto simp: mult_ac) next fix x :: real have "x powr (a - 1) * (1 - x) powr (b - 1) \ x powr (a - 1) * C" if "x \ {0..1/2}" using that by (intro mult_left_mono powr_mono2 C) auto thus "norm (indicator {0..1/2} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \ norm (indicator {0..1/2} x *\<^sub>R (C * x powr (a - 1)))" by (auto simp: indicator_def abs_mult mult_ac) qed (auto intro!: AE_I2 simp: indicator_def) have I2: "set_integrable lborel {1/2..1} (\t. t powr (a - 1) * (1 - t) powr (b - 1))" unfolding set_integrable_def proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2]) have "(\t. t powr (b - 1)) integrable_on {0..1/2}" by (rule integrable_on_powr_from_0) (use assms in auto) hence "(\t. t powr (b - 1)) integrable_on (cbox 0 (1/2))" by simp from integrable_affinity[OF this, of "-1" 1] have "(\t. (1 - t) powr (b - 1)) integrable_on {1/2..1}" by simp hence "(\t. (1 - t) powr (b - 1)) absolutely_integrable_on {1/2..1}" by (subst absolutely_integrable_on_iff_nonneg) auto from integrable_mult_right[OF this [unfolded set_integrable_def], of D] show "integrable lborel (\x. indicat_real {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))" by (subst (asm) integrable_completion) (auto simp: mult_ac) next fix x :: real have "x powr (a - 1) * (1 - x) powr (b - 1) \ D * (1 - x) powr (b - 1)" if "x \ {1/2..1}" using that by (intro mult_right_mono powr_mono2 D) auto thus "norm (indicator {1/2..1} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \ norm (indicator {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))" by (auto simp: indicator_def abs_mult mult_ac) qed (auto intro!: AE_I2 simp: indicator_def) have "set_integrable lborel ({0..1/2} \ {1/2..1}) (\t. t powr (a - 1) * (1 - t) powr (b - 1))" by (intro set_integrable_Un I1 I2) auto also have "{0..1/2} \ {1/2..1} = {0..(1::real)}" by auto finally show ?thesis . qed lemma integrable_Beta': assumes "a > 0" "b > (0::real)" shows "(\t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}" using integrable_Beta[OF assms] by (rule set_borel_integral_eq_integral) theorem has_integral_Beta_real: assumes a: "a > 0" and b: "b > (0 :: real)" shows "((\t. t powr (a - 1) * (1 - t) powr (b - 1)) has_integral Beta a b) {0..1}" proof - define B where "B = integral {0..1} (\x. x powr (a - 1) * (1 - x) powr (b - 1))" have [simp]: "B \ 0" unfolding B_def using a b by (intro integral_nonneg integrable_Beta') auto from a b have "ennreal (Gamma a * Gamma b) = (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \lborel) * (\\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \lborel)" by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) * ennreal (indicator {0..} u * u powr (b - 1) / exp u) \lborel \lborel)" by (simp add: nn_integral_cmult nn_integral_multc) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator ({0..}\{0..}) (t,u) * t powr (a - 1) * u powr (b - 1) / exp (t + u)) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add) also have "\ = (\\<^sup>+t. \\<^sup>+u. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel \lborel)" proof (rule nn_integral_cong, goal_cases) case (1 t) have "(\\<^sup>+u. ennreal (indicator ({0..}\{0..}) (t,u) * t powr (a - 1) * u powr (b - 1) / exp (t + u)) \distr lborel borel ((+) (-t))) = (\\<^sup>+u. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel)" by (subst nn_integral_distr) (auto intro!: nn_integral_cong simp: indicator_def) thus ?case by (subst (asm) lborel_distr_plus) qed also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator ({0..}\{t..}) (t,u) * t powr (a - 1) * (u - t) powr (b - 1) / exp u) \lborel \lborel)" by (subst lborel_pair.Fubini') (auto simp: case_prod_unfold indicator_def cong: measurable_cong_sets) also have "\ = (\\<^sup>+u. \\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) * ennreal (indicator {0..} u / exp u) \lborel \lborel)" by (intro nn_integral_cong) (auto simp: indicator_def ennreal_mult' [symmetric]) also have "\ = (\\<^sup>+u. (\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) * ennreal (indicator {0..} u / exp u) \lborel)" by (subst nn_integral_multc [symmetric]) auto also have "\ = (\\<^sup>+u. (\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) * ennreal (indicator {0<..} u / exp u) \lborel)" by (intro nn_integral_cong_AE eventually_mono[OF AE_lborel_singleton[of 0]]) (auto simp: indicator_def) also have "\ = (\\<^sup>+u. ennreal B * ennreal (indicator {0..} u / exp u * u powr (a + b - 1)) \lborel)" proof (intro nn_integral_cong, goal_cases) case (1 u) show ?case proof (cases "u > 0") case True have "(\\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \lborel) = (\\<^sup>+t. ennreal (indicator {0..1} t * (u * t) powr (a - 1) * (u - u * t) powr (b - 1)) \distr lborel borel ((*) (1 / u)))" (is "_ = nn_integral _ ?f") using True by (subst nn_integral_distr) (auto simp: indicator_def field_simps intro!: nn_integral_cong) also have "distr lborel borel ((*) (1 / u)) = density lborel (\_. u)" using \u > 0\ by (subst lborel_distr_mult) auto also have "nn_integral \ ?f = (\\<^sup>+x. ennreal (indicator {0..1} x * (u * (u * x) powr (a - 1) * (u * (1 - x)) powr (b - 1))) \lborel)" using \u > 0\ by (subst nn_integral_density) (auto simp: ennreal_mult' [symmetric] algebra_simps) also have "\ = (\\<^sup>+x. ennreal (u powr (a + b - 1)) * ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel)" using \u > 0\ a b by (intro nn_integral_cong) (auto simp: indicator_def powr_mult powr_add powr_diff mult_ac ennreal_mult' [symmetric]) also have "\ = ennreal (u powr (a + b - 1)) * (\\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel)" by (subst nn_integral_cmult) auto also have "((\x. x powr (a - 1) * (1 - x) powr (b - 1)) has_integral integral {0..1} (\x. x powr (a - 1) * (1 - x) powr (b - 1))) {0..1}" using a b by (intro integrable_integral integrable_Beta') from nn_integral_has_integral_lebesgue[OF _ this] a b have "(\\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) * (1 - x) powr (b - 1)) \lborel) = B" by (simp add: mult_ac B_def) finally show ?thesis using \u > 0\ by (simp add: ennreal_mult' [symmetric] mult_ac) qed auto qed also have "\ = ennreal B * ennreal (Gamma (a + b))" using a b by (subst nn_integral_cmult) (auto simp: Gamma_conv_nn_integral_real) also have "\ = ennreal (B * Gamma (a + b))" by (subst (1 2) mult.commute, intro ennreal_mult' [symmetric]) (use a b in auto) finally have "B = Beta a b" using a b Gamma_real_pos[of "a + b"] by (subst (asm) ennreal_inj) (auto simp: field_simps Beta_def Gamma_eq_zero_iff) moreover have "(\t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}" by (intro integrable_Beta' a b) ultimately show ?thesis by (simp add: has_integral_iff B_def) qed subsection \The Weierstraß product formula for the sine\ theorem sin_product_formula_complex: fixes z :: complex shows "(\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k^2)) \ sin (of_real pi * z)" proof - let ?f = "rGamma_series_Weierstrass" have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n)) \ (- of_real pi * inverse z) * (rGamma z * rGamma (- z))" by (intro tendsto_intros rGamma_Weierstrass_complex) also have "(\n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) = (\n. of_real pi * z * (\k=1..n. 1 - z^2 / of_nat k ^ 2))" proof fix n :: nat have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * (\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)" by (simp add: rGamma_series_Weierstrass_def mult_ac exp_minus divide_simps prod.distrib[symmetric] power2_eq_square) also have "(\k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) = (\k=1..n. 1 - z^2 / of_nat k ^ 2)" by (intro prod.cong) (simp_all add: power2_eq_square field_simps) finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \" by (simp add: field_split_simps) qed also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)" by (subst rGamma_reflection_complex') (simp add: field_split_simps) finally show ?thesis . qed lemma sin_product_formula_real: "(\n. pi * (x::real) * (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x)" proof - from sin_product_formula_complex[of "of_real x"] have "(\n. of_real pi * of_real x * (\k=1..n. 1 - (of_real x)^2 / (of_nat k)^2)) \ sin (of_real pi * of_real x :: complex)" (is "?f \ ?y") . also have "?f = (\n. of_real (pi * x * (\k=1..n. 1 - x^2 / (of_nat k^2))))" by simp also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult) finally show ?thesis by (subst (asm) tendsto_of_real_iff) qed lemma sin_product_formula_real': assumes "x \ (0::real)" shows "(\n. (\k=1..n. 1 - x^2 / of_nat k^2)) \ sin (pi * x) / (pi * x)" using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms by simp theorem wallis: "(\n. \k=1..n. (4*real k^2) / (4*real k^2 - 1)) \ pi / 2" proof - from tendsto_inverse[OF tendsto_mult[OF sin_product_formula_real[of "1/2"] tendsto_const[of "2/pi"]]] have "(\n. (\k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) \ pi/2" by (simp add: prod_inversef [symmetric]) also have "(\n. (\k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) = (\n. (\k=1..n. (4*real k^2)/(4*real k^2 - 1)))" by (intro ext prod.cong refl) (simp add: field_split_simps) finally show ?thesis . qed subsection \The Solution to the Basel problem\ theorem inverse_squares_sums: "(\n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)" proof - define P where "P x n = (\k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n define K where "K = (\n. inverse (real_of_nat (Suc n))^2)" define f where [abs_def]: "f x = (\n. P x n / of_nat (Suc n)^2)" for x define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x have sums: "(\n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x proof (cases "x = 0") assume x: "x = 0" have "summable (\n. inverse ((real_of_nat (Suc n))\<^sup>2))" using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums) next assume x: "x \ 0" have "(\n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))" unfolding P_def using x by (intro telescope_sums' sin_product_formula_real') also have "(\n. P x n - P x (Suc n)) = (\n. (x^2 / of_nat (Suc n)^2) * P x n)" unfolding P_def by (simp add: prod.nat_ivl_Suc' algebra_simps) also have "P x 0 = 1" by (simp add: P_def) finally have "(\n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" . from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp qed have "continuous_on (ball 0 1) f" proof (rule uniform_limit_theorem; (intro always_eventually allI)?) show "uniform_limit (ball 0 1) (\n x. \k ball 0 1" { fix k :: nat assume k: "k \ 1" from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1) also from k have "\ \ of_nat k^2" by simp finally have "(1 - x^2 / of_nat k^2) \ {0..1}" using k by (simp_all add: field_simps del: of_nat_Suc) } hence "(\k=1..n. abs (1 - x^2 / of_nat k^2)) \ (\k=1..n. 1)" by (intro prod_mono) simp thus "norm (P x n / (of_nat (Suc n)^2)) \ 1 / of_nat (Suc n)^2" unfolding P_def by (simp add: field_simps abs_prod del: of_nat_Suc) qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide) qed (auto simp: P_def intro!: continuous_intros) hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all hence "(f \ 0 \ f 0)" by (simp add: isCont_def) also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide) finally have "f \ 0 \ K" . moreover have "f \ 0 \ pi^2 / 6" proof (rule Lim_transform_eventually) define f' where [abs_def]: "f' x = (\n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x have "eventually (\x. x \ (0::real)) (at 0)" by (auto simp add: eventually_at intro!: exI[of _ 1]) thus "eventually (\x. f' x = f x) (at 0)" proof eventually_elim fix x :: real assume x: "x \ 0" have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def) with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"] have "(\n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))" by (simp add: eval_nat_numeral) from sums_divide[OF this, of "x^3 * pi"] x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)" by (simp add: field_split_simps eval_nat_numeral power_mult_distrib mult_ac) with x have "(\n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)" by (simp add: g_def) hence "f' x = g x / x^2" by (simp add: sums_iff f'_def) also have "\ = f x" using sums[of x] x by (simp add: sums_iff g_def f_def) finally show "f' x = f x" . qed have "isCont f' 0" unfolding f'_def proof (intro isCont_powser_converges_everywhere) fix x :: real show "summable (\n. -sin_coeff (n+3) * pi^(n+2) * x^n)" proof (cases "x = 0") assume x: "x \ 0" from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x show ?thesis by (simp add: mult_ac power_mult_distrib field_split_simps eval_nat_numeral) qed (simp only: summable_0_powser) qed hence "f' \ 0 \ f' 0" by (simp add: isCont_def) also have "f' 0 = pi * pi / fact 3" unfolding f'_def by (subst powser_zero) (simp add: sin_coeff_def) finally show "f' \ 0 \ pi^2 / 6" by (simp add: eval_nat_numeral) qed ultimately have "K = pi^2 / 6" by (rule LIM_unique) moreover from inverse_power_summable[of 2] have "summable (\n. (inverse (real_of_nat (Suc n)))\<^sup>2)" by (subst summable_Suc_iff) (simp add: power_inverse) ultimately show ?thesis unfolding K_def by (auto simp add: sums_iff power_divide inverse_eq_divide) qed end diff --git a/src/HOL/Analysis/Great_Picard.thy b/src/HOL/Analysis/Great_Picard.thy deleted file mode 100644 --- a/src/HOL/Analysis/Great_Picard.thy +++ /dev/null @@ -1,1848 +0,0 @@ -section \The Great Picard Theorem and its Applications\ - -text\Ported from HOL Light (cauchy.ml) by L C Paulson, 2017\ - -theory Great_Picard - imports Conformal_Mappings Further_Topology - -begin - -subsection\Schottky's theorem\ - -lemma Schottky_lemma0: - assumes holf: "f holomorphic_on S" and cons: "contractible S" and "a \ S" - and f: "\z. z \ S \ f z \ 1 \ f z \ -1" - obtains g where "g holomorphic_on S" - "norm(g a) \ 1 + norm(f a) / 3" - "\z. z \ S \ f z = cos(of_real pi * g z)" -proof - - obtain g where holg: "g holomorphic_on S" and g: "norm(g a) \ pi + norm(f a)" - and f_eq_cos: "\z. z \ S \ f z = cos(g z)" - using contractible_imp_holomorphic_arccos_bounded [OF assms] - by blast - show ?thesis - proof - show "(\z. g z / pi) holomorphic_on S" - by (auto intro: holomorphic_intros holg) - have "3 \ pi" - using pi_approx by force - have "3 * norm(g a) \ 3 * (pi + norm(f a))" - using g by auto - also have "... \ pi * 3 + pi * cmod (f a)" - using \3 \ pi\ by (simp add: mult_right_mono algebra_simps) - finally show "cmod (g a / complex_of_real pi) \ 1 + cmod (f a) / 3" - by (simp add: field_simps norm_divide) - show "\z. z \ S \ f z = cos (complex_of_real pi * (g z / complex_of_real pi))" - by (simp add: f_eq_cos) - qed -qed - - -lemma Schottky_lemma1: - fixes n::nat - assumes "0 < n" - shows "0 < n + sqrt(real n ^ 2 - 1)" -proof - - have "(n-1)^2 \ n^2 - 1" - using assms by (simp add: algebra_simps power2_eq_square) - then have "real (n - 1) \ sqrt (real (n\<^sup>2 - 1))" - by (metis of_nat_le_iff of_nat_power real_le_rsqrt) - then have "n-1 \ sqrt(real n ^ 2 - 1)" - by (simp add: Suc_leI assms of_nat_diff) - then show ?thesis - using assms by linarith -qed - - -lemma Schottky_lemma2: - fixes x::real - assumes "0 \ x" - obtains n where "0 < n" "\x - ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi\ < 1/2" -proof - - obtain n0::nat where "0 < n0" "ln(n0 + sqrt(real n0 ^ 2 - 1)) / pi \ x" - proof - show "ln(real 1 + sqrt(real 1 ^ 2 - 1))/pi \ x" - by (auto simp: assms) - qed auto - moreover - obtain M::nat where "\n. \0 < n; ln(n + sqrt(real n ^ 2 - 1)) / pi \ x\ \ n \ M" - proof - fix n::nat - assume "0 < n" "ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi \ x" - then have "ln (n + sqrt ((real n)\<^sup>2 - 1)) \ x * pi" - by (simp add: field_split_simps) - then have *: "exp (ln (n + sqrt ((real n)\<^sup>2 - 1))) \ exp (x * pi)" - by blast - have 0: "0 \ sqrt ((real n)\<^sup>2 - 1)" - using \0 < n\ by auto - have "n + sqrt ((real n)\<^sup>2 - 1) = exp (ln (n + sqrt ((real n)\<^sup>2 - 1)))" - by (simp add: Suc_leI \0 < n\ add_pos_nonneg real_of_nat_ge_one_iff) - also have "... \ exp (x * pi)" - using "*" by blast - finally have "real n \ exp (x * pi)" - using 0 by linarith - then show "n \ nat (ceiling (exp(x * pi)))" - by linarith - qed - ultimately obtain n where - "0 < n" and le_x: "ln(n + sqrt(real n ^ 2 - 1)) / pi \ x" - and le_n: "\k. \0 < k; ln(k + sqrt(real k ^ 2 - 1)) / pi \ x\ \ k \ n" - using bounded_Max_nat [of "\n. 0 ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi \ x"] by metis - define a where "a \ ln(n + sqrt(real n ^ 2 - 1)) / pi" - define b where "b \ ln (1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / pi" - have le_xa: "a \ x" - and le_na: "\k. \0 < k; ln(k + sqrt(real k ^ 2 - 1)) / pi \ x\ \ k \ n" - using le_x le_n by (auto simp: a_def) - moreover have "x < b" - using le_n [of "Suc n"] by (force simp: b_def) - moreover have "b - a < 1" - proof - - have "ln (1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) - ln (real n + sqrt ((real n)\<^sup>2 - 1)) = - ln ((1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / (real n + sqrt ((real n)\<^sup>2 - 1)))" - by (simp add: \0 < n\ Schottky_lemma1 add_pos_nonneg ln_div [symmetric]) - also have "... \ 3" - proof (cases "n = 1") - case True - have "sqrt 3 \ 2" - by (simp add: real_le_lsqrt) - then have "(2 + sqrt 3) \ 4" - by simp - also have "... \ exp 3" - using exp_ge_add_one_self [of "3::real"] by simp - finally have "ln (2 + sqrt 3) \ 3" - by (metis add_nonneg_nonneg add_pos_nonneg dbl_def dbl_inc_def dbl_inc_simps(3) - dbl_simps(3) exp_gt_zero ln_exp ln_le_cancel_iff real_sqrt_ge_0_iff zero_le_one zero_less_one) - then show ?thesis - by (simp add: True) - next - case False with \0 < n\ have "1 < n" "2 \ n" - by linarith+ - then have 1: "1 \ real n * real n" - by (metis less_imp_le_nat one_le_power power2_eq_square real_of_nat_ge_one_iff) - have *: "4 + (m+2) * 2 \ (m+2) * ((m+2) * 3)" for m::nat - by simp - have "4 + n * 2 \ n * (n * 3)" - using * [of "n-2"] \2 \ n\ - by (metis le_add_diff_inverse2) - then have **: "4 + real n * 2 \ real n * (real n * 3)" - by (metis (mono_tags, hide_lams) of_nat_le_iff of_nat_add of_nat_mult of_nat_numeral) - have "sqrt ((1 + real n)\<^sup>2 - 1) \ 2 * sqrt ((real n)\<^sup>2 - 1)" - by (auto simp: real_le_lsqrt power2_eq_square algebra_simps 1 **) - then - have "((1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / (real n + sqrt ((real n)\<^sup>2 - 1))) \ 2" - using Schottky_lemma1 \0 < n\ by (simp add: field_split_simps) - then have "ln ((1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / (real n + sqrt ((real n)\<^sup>2 - 1))) \ ln 2" - apply (subst ln_le_cancel_iff) - using Schottky_lemma1 \0 < n\ by auto (force simp: field_split_simps) - also have "... \ 3" - using ln_add_one_self_le_self [of 1] by auto - finally show ?thesis . - qed - also have "... < pi" - using pi_approx by simp - finally show ?thesis - by (simp add: a_def b_def field_split_simps) - qed - ultimately have "\x - a\ < 1/2 \ \x - b\ < 1/2" - by (auto simp: abs_if) - then show thesis - proof - assume "\x - a\ < 1 / 2" - then show ?thesis - by (rule_tac n=n in that) (auto simp: a_def \0 < n\) - next - assume "\x - b\ < 1 / 2" - then show ?thesis - by (rule_tac n="Suc n" in that) (auto simp: b_def \0 < n\) - qed -qed - - -lemma Schottky_lemma3: - fixes z::complex - assumes "z \ (\m \ Ints. \n \ {0<..}. {Complex m (ln(n + sqrt(real n ^ 2 - 1)) / pi)}) - \ (\m \ Ints. \n \ {0<..}. {Complex m (-ln(n + sqrt(real n ^ 2 - 1)) / pi)})" - shows "cos(pi * cos(pi * z)) = 1 \ cos(pi * cos(pi * z)) = -1" -proof - - have sqrt2 [simp]: "complex_of_real (sqrt x) * complex_of_real (sqrt x) = x" if "x \ 0" for x::real - by (metis abs_of_nonneg of_real_mult real_sqrt_mult_self that) - have 1: "\k. exp (\ * (of_int m * complex_of_real pi) - - (ln (real n + sqrt ((real n)\<^sup>2 - 1)))) + - inverse - (exp (\ * (of_int m * complex_of_real pi) - - (ln (real n + sqrt ((real n)\<^sup>2 - 1))))) = of_int k * 2" - if "n > 0" for m n - proof - - have eeq: "e \ 0 \ e + inverse e = n*2 \ inverse e^2 - 2 * n*inverse e + 1 = 0" for n e::complex - by (auto simp: field_simps power2_eq_square) - have [simp]: "1 \ real n * real n" - by (metis One_nat_def add.commute nat_less_real_le of_nat_1 of_nat_Suc one_le_power power2_eq_square that) - show ?thesis - apply (simp add: eeq) - using Schottky_lemma1 [OF that] - apply (auto simp: eeq exp_diff exp_Euler exp_of_real algebra_simps sin_int_times_real cos_int_times_real) - apply (rule_tac x="int n" in exI) - apply (auto simp: power2_eq_square algebra_simps) - apply (rule_tac x="- int n" in exI) - apply (auto simp: power2_eq_square algebra_simps) - done - qed - have 2: "\k. exp (\ * (of_int m * complex_of_real pi) + - (ln (real n + sqrt ((real n)\<^sup>2 - 1)))) + - inverse - (exp (\ * (of_int m * complex_of_real pi) + - (ln (real n + sqrt ((real n)\<^sup>2 - 1))))) = of_int k * 2" - if "n > 0" for m n - proof - - have eeq: "e \ 0 \ e + inverse e = n*2 \ e^2 - 2 * n*e + 1 = 0" for n e::complex - by (auto simp: field_simps power2_eq_square) - have [simp]: "1 \ real n * real n" - by (metis One_nat_def add.commute nat_less_real_le of_nat_1 of_nat_Suc one_le_power power2_eq_square that) - show ?thesis - apply (simp add: eeq) - using Schottky_lemma1 [OF that] - apply (auto simp: exp_add exp_Euler exp_of_real algebra_simps sin_int_times_real cos_int_times_real) - apply (rule_tac x="int n" in exI) - apply (auto simp: power2_eq_square algebra_simps) - apply (rule_tac x="- int n" in exI) - apply (auto simp: power2_eq_square algebra_simps) - done - qed - have "\x. cos (complex_of_real pi * z) = of_int x" - using assms - apply safe - apply (auto simp: Ints_def cos_exp_eq exp_minus Complex_eq) - apply (auto simp: algebra_simps dest: 1 2) - done - then have "sin(pi * cos(pi * z)) ^ 2 = 0" - by (simp add: Complex_Transcendental.sin_eq_0) - then have "1 - cos(pi * cos(pi * z)) ^ 2 = 0" - by (simp add: sin_squared_eq) - then show ?thesis - using power2_eq_1_iff by auto -qed - - -theorem Schottky: - assumes holf: "f holomorphic_on cball 0 1" - and nof0: "norm(f 0) \ r" - and not01: "\z. z \ cball 0 1 \ \(f z = 0 \ f z = 1)" - and "0 < t" "t < 1" "norm z \ t" - shows "norm(f z) \ exp(pi * exp(pi * (2 + 2 * r + 12 * t / (1 - t))))" -proof - - obtain h where holf: "h holomorphic_on cball 0 1" - and nh0: "norm (h 0) \ 1 + norm(2 * f 0 - 1) / 3" - and h: "\z. z \ cball 0 1 \ 2 * f z - 1 = cos(of_real pi * h z)" - proof (rule Schottky_lemma0 [of "\z. 2 * f z - 1" "cball 0 1" 0]) - show "(\z. 2 * f z - 1) holomorphic_on cball 0 1" - by (intro holomorphic_intros holf) - show "contractible (cball (0::complex) 1)" - by (auto simp: convex_imp_contractible) - show "\z. z \ cball 0 1 \ 2 * f z - 1 \ 1 \ 2 * f z - 1 \ - 1" - using not01 by force - qed auto - obtain g where holg: "g holomorphic_on cball 0 1" - and ng0: "norm(g 0) \ 1 + norm(h 0) / 3" - and g: "\z. z \ cball 0 1 \ h z = cos(of_real pi * g z)" - proof (rule Schottky_lemma0 [OF holf convex_imp_contractible, of 0]) - show "\z. z \ cball 0 1 \ h z \ 1 \ h z \ - 1" - using h not01 by fastforce+ - qed auto - have g0_2_f0: "norm(g 0) \ 2 + norm(f 0)" - proof - - have "cmod (2 * f 0 - 1) \ cmod (2 * f 0) + 1" - by (metis norm_one norm_triangle_ineq4) - also have "... \ 2 + cmod (f 0) * 3" - by simp - finally have "1 + norm(2 * f 0 - 1) / 3 \ (2 + norm(f 0) - 1) * 3" - apply (simp add: field_split_simps) - using norm_ge_zero [of "f 0 * 2 - 1"] - by linarith - with nh0 have "norm(h 0) \ (2 + norm(f 0) - 1) * 3" - by linarith - then have "1 + norm(h 0) / 3 \ 2 + norm(f 0)" - by simp - with ng0 show ?thesis - by auto - qed - have "z \ ball 0 1" - using assms by auto - have norm_g_12: "norm(g z - g 0) \ (12 * t) / (1 - t)" - proof - - obtain g' where g': "\x. x \ cball 0 1 \ (g has_field_derivative g' x) (at x within cball 0 1)" - using holg [unfolded holomorphic_on_def field_differentiable_def] by metis - have int_g': "(g' has_contour_integral g z - g 0) (linepath 0 z)" - using contour_integral_primitive [OF g' valid_path_linepath, of 0 z] - using \z \ ball 0 1\ segment_bound1 by fastforce - have "cmod (g' w) \ 12 / (1 - t)" if "w \ closed_segment 0 z" for w - proof - - have w: "w \ ball 0 1" - using segment_bound [OF that] \z \ ball 0 1\ by simp - have ttt: "\z. z \ frontier (cball 0 1) \ 1 - t \ dist w z" - using \norm z \ t\ segment_bound1 [OF \w \ closed_segment 0 z\] - apply (simp add: dist_complex_def) - by (metis diff_left_mono dist_commute dist_complex_def norm_triangle_ineq2 order_trans) - have *: "\\b. (\w \ T \ U. w \ ball b 1); \x. x \ D \ g x \ T \ U\ \ \b. ball b 1 \ g ` D" for T U D - by force - have "\b. ball b 1 \ g ` cball 0 1" - proof (rule *) - show "(\w \ (\m \ Ints. \n \ {0<..}. {Complex m (ln(n + sqrt(real n ^ 2 - 1)) / pi)}) \ - (\m \ Ints. \n \ {0<..}. {Complex m (-ln(n + sqrt(real n ^ 2 - 1)) / pi)}). w \ ball b 1)" for b - proof - - obtain m where m: "m \ \" "\Re b - m\ \ 1/2" - by (metis Ints_of_int abs_minus_commute of_int_round_abs_le) - show ?thesis - proof (cases "0::real" "Im b" rule: le_cases) - case le - then obtain n where "0 < n" and n: "\Im b - ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi\ < 1/2" - using Schottky_lemma2 [of "Im b"] by blast - have "dist b (Complex m (Im b)) \ 1/2" - by (metis cancel_comm_monoid_add_class.diff_cancel cmod_eq_Re complex.sel(1) complex.sel(2) dist_norm m(2) minus_complex.code) - moreover - have "dist (Complex m (Im b)) (Complex m (ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi)) < 1/2" - using n by (simp add: complex_norm cmod_eq_Re complex_diff dist_norm del: Complex_eq) - ultimately have "dist b (Complex m (ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)) < 1" - by (simp add: dist_triangle_lt [of b "Complex m (Im b)"] dist_commute) - with le m \0 < n\ show ?thesis - apply (rule_tac x = "Complex m (ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)" in bexI) - apply (simp_all del: Complex_eq greaterThan_0) - by blast - next - case ge - then obtain n where "0 < n" and n: "\- Im b - ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi\ < 1/2" - using Schottky_lemma2 [of "- Im b"] by auto - have "dist b (Complex m (Im b)) \ 1/2" - by (metis cancel_comm_monoid_add_class.diff_cancel cmod_eq_Re complex.sel(1) complex.sel(2) dist_norm m(2) minus_complex.code) - moreover - have "dist (Complex m (- ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi)) (Complex m (Im b)) < 1/2" - using n - apply (simp add: complex_norm cmod_eq_Re complex_diff dist_norm del: Complex_eq) - by (metis add.commute add_uminus_conv_diff) - ultimately have "dist b (Complex m (- ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)) < 1" - by (simp add: dist_triangle_lt [of b "Complex m (Im b)"] dist_commute) - with ge m \0 < n\ show ?thesis - apply (rule_tac x = "Complex m (- ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)" in bexI) - apply (simp_all del: Complex_eq greaterThan_0) - by blast - qed - qed - show "g v \ (\m \ Ints. \n \ {0<..}. {Complex m (ln(n + sqrt(real n ^ 2 - 1)) / pi)}) \ - (\m \ Ints. \n \ {0<..}. {Complex m (-ln(n + sqrt(real n ^ 2 - 1)) / pi)})" - if "v \ cball 0 1" for v - using not01 [OF that] - by (force simp: g [OF that, symmetric] h [OF that, symmetric] dest!: Schottky_lemma3 [of "g v"]) - qed - then have 12: "(1 - t) * cmod (deriv g w) / 12 < 1" - using Bloch_general [OF holg _ ttt, of 1] w by force - have "g field_differentiable at w within cball 0 1" - using holg w by (simp add: holomorphic_on_def) - then have "g field_differentiable at w within ball 0 1" - using ball_subset_cball field_differentiable_within_subset by blast - with w have der_gw: "(g has_field_derivative deriv g w) (at w)" - by (simp add: field_differentiable_within_open [of _ "ball 0 1"] field_differentiable_derivI) - with DERIV_unique [OF der_gw] g' w have "deriv g w = g' w" - by (metis open_ball at_within_open ball_subset_cball has_field_derivative_subset subsetCE) - then show "cmod (g' w) \ 12 / (1 - t)" - using g' 12 \t < 1\ by (simp add: field_simps) - qed - then have "cmod (g z - g 0) \ 12 / (1 - t) * cmod z" - using has_contour_integral_bound_linepath [OF int_g', of "12/(1 - t)"] assms - by simp - with \cmod z \ t\ \t < 1\ show ?thesis - by (simp add: field_split_simps) - qed - have fz: "f z = (1 + cos(of_real pi * h z)) / 2" - using h \z \ ball 0 1\ by (auto simp: field_simps) - have "cmod (f z) \ exp (cmod (complex_of_real pi * h z))" - by (simp add: fz mult.commute norm_cos_plus1_le) - also have "... \ exp (pi * exp (pi * (2 + 2 * r + 12 * t / (1 - t))))" - proof (simp add: norm_mult) - have "cmod (g z - g 0) \ 12 * t / (1 - t)" - using norm_g_12 \t < 1\ by (simp add: norm_mult) - then have "cmod (g z) - cmod (g 0) \ 12 * t / (1 - t)" - using norm_triangle_ineq2 order_trans by blast - then have *: "cmod (g z) \ 2 + 2 * r + 12 * t / (1 - t)" - using g0_2_f0 norm_ge_zero [of "f 0"] nof0 - by linarith - have "cmod (h z) \ exp (cmod (complex_of_real pi * g z))" - using \z \ ball 0 1\ by (simp add: g norm_cos_le) - also have "... \ exp (pi * (2 + 2 * r + 12 * t / (1 - t)))" - using \t < 1\ nof0 * by (simp add: norm_mult) - finally show "cmod (h z) \ exp (pi * (2 + 2 * r + 12 * t / (1 - t)))" . - qed - finally show ?thesis . -qed - - -subsection\The Little Picard Theorem\ - -theorem Landau_Picard: - obtains R - where "\z. 0 < R z" - "\f. \f holomorphic_on cball 0 (R(f 0)); - \z. norm z \ R(f 0) \ f z \ 0 \ f z \ 1\ \ norm(deriv f 0) < 1" -proof - - define R where "R \ \z. 3 * exp(pi * exp(pi*(2 + 2 * cmod z + 12)))" - show ?thesis - proof - show Rpos: "\z. 0 < R z" - by (auto simp: R_def) - show "norm(deriv f 0) < 1" - if holf: "f holomorphic_on cball 0 (R(f 0))" - and Rf: "\z. norm z \ R(f 0) \ f z \ 0 \ f z \ 1" for f - proof - - let ?r = "R(f 0)" - define g where "g \ f \ (\z. of_real ?r * z)" - have "0 < ?r" - using Rpos by blast - have holg: "g holomorphic_on cball 0 1" - unfolding g_def - apply (intro holomorphic_intros holomorphic_on_compose holomorphic_on_subset [OF holf]) - using Rpos by (auto simp: less_imp_le norm_mult) - have *: "norm(g z) \ exp(pi * exp(pi * (2 + 2 * norm (f 0) + 12 * t / (1 - t))))" - if "0 < t" "t < 1" "norm z \ t" for t z - proof (rule Schottky [OF holg]) - show "cmod (g 0) \ cmod (f 0)" - by (simp add: g_def) - show "\z. z \ cball 0 1 \ \ (g z = 0 \ g z = 1)" - using Rpos by (simp add: g_def less_imp_le norm_mult Rf) - qed (auto simp: that) - have C1: "g holomorphic_on ball 0 (1 / 2)" - by (rule holomorphic_on_subset [OF holg]) auto - have C2: "continuous_on (cball 0 (1 / 2)) g" - by (meson cball_divide_subset_numeral holg holomorphic_on_imp_continuous_on holomorphic_on_subset) - have C3: "cmod (g z) \ R (f 0) / 3" if "cmod (0 - z) = 1/2" for z - proof - - have "norm(g z) \ exp(pi * exp(pi * (2 + 2 * norm (f 0) + 12)))" - using * [of "1/2"] that by simp - also have "... = ?r / 3" - by (simp add: R_def) - finally show ?thesis . - qed - then have cmod_g'_le: "cmod (deriv g 0) * 3 \ R (f 0) * 2" - using Cauchy_inequality [OF C1 C2 _ C3, of 1] by simp - have holf': "f holomorphic_on ball 0 (R(f 0))" - by (rule holomorphic_on_subset [OF holf]) auto - then have fd0: "f field_differentiable at 0" - by (rule holomorphic_on_imp_differentiable_at [OF _ open_ball]) - (auto simp: Rpos [of "f 0"]) - have g_eq: "deriv g 0 = of_real ?r * deriv f 0" - apply (rule DERIV_imp_deriv) - apply (simp add: g_def) - by (metis DERIV_chain DERIV_cmult_Id fd0 field_differentiable_derivI mult.commute mult_zero_right) - show ?thesis - using cmod_g'_le Rpos [of "f 0"] by (simp add: g_eq norm_mult) - qed - qed -qed - -lemma little_Picard_01: - assumes holf: "f holomorphic_on UNIV" and f01: "\z. f z \ 0 \ f z \ 1" - obtains c where "f = (\x. c)" -proof - - obtain R - where Rpos: "\z. 0 < R z" - and R: "\h. \h holomorphic_on cball 0 (R(h 0)); - \z. norm z \ R(h 0) \ h z \ 0 \ h z \ 1\ \ norm(deriv h 0) < 1" - using Landau_Picard by metis - have contf: "continuous_on UNIV f" - by (simp add: holf holomorphic_on_imp_continuous_on) - show ?thesis - proof (cases "\x. deriv f x = 0") - case True - obtain c where "\x. f(x) = c" - apply (rule DERIV_zero_connected_constant [OF connected_UNIV open_UNIV finite.emptyI contf]) - apply (metis True DiffE holf holomorphic_derivI open_UNIV, auto) - done - then show ?thesis - using that by auto - next - case False - then obtain w where w: "deriv f w \ 0" by auto - define fw where "fw \ (f \ (\z. w + z / deriv f w))" - have norm_let1: "norm(deriv fw 0) < 1" - proof (rule R) - show "fw holomorphic_on cball 0 (R (fw 0))" - unfolding fw_def - by (intro holomorphic_intros holomorphic_on_compose w holomorphic_on_subset [OF holf] subset_UNIV) - show "fw z \ 0 \ fw z \ 1" if "cmod z \ R (fw 0)" for z - using f01 by (simp add: fw_def) - qed - have "(fw has_field_derivative deriv f w * inverse (deriv f w)) (at 0)" - apply (simp add: fw_def) - apply (rule DERIV_chain) - using holf holomorphic_derivI apply force - apply (intro derivative_eq_intros w) - apply (auto simp: field_simps) - done - then show ?thesis - using norm_let1 w by (simp add: DERIV_imp_deriv) - qed -qed - - -theorem little_Picard: - assumes holf: "f holomorphic_on UNIV" - and "a \ b" "range f \ {a,b} = {}" - obtains c where "f = (\x. c)" -proof - - let ?g = "\x. 1/(b - a)*(f x - b) + 1" - obtain c where "?g = (\x. c)" - proof (rule little_Picard_01) - show "?g holomorphic_on UNIV" - by (intro holomorphic_intros holf) - show "\z. ?g z \ 0 \ ?g z \ 1" - using assms by (auto simp: field_simps) - qed auto - then have "?g x = c" for x - by meson - then have "f x = c * (b-a) + a" for x - using assms by (auto simp: field_simps) - then show ?thesis - using that by blast -qed - - -text\A couple of little applications of Little Picard\ - -lemma holomorphic_periodic_fixpoint: - assumes holf: "f holomorphic_on UNIV" - and "p \ 0" and per: "\z. f(z + p) = f z" - obtains x where "f x = x" -proof - - have False if non: "\x. f x \ x" - proof - - obtain c where "(\z. f z - z) = (\z. c)" - proof (rule little_Picard) - show "(\z. f z - z) holomorphic_on UNIV" - by (simp add: holf holomorphic_on_diff) - show "range (\z. f z - z) \ {p,0} = {}" - using assms non by auto (metis add.commute diff_eq_eq) - qed (auto simp: assms) - with per show False - by (metis add.commute add_cancel_left_left \p \ 0\ diff_add_cancel) - qed - then show ?thesis - using that by blast -qed - - -lemma holomorphic_involution_point: - assumes holfU: "f holomorphic_on UNIV" and non: "\a. f \ (\x. a + x)" - obtains x where "f(f x) = x" -proof - - { assume non_ff [simp]: "\x. f(f x) \ x" - then have non_fp [simp]: "f z \ z" for z - by metis - have holf: "f holomorphic_on X" for X - using assms holomorphic_on_subset by blast - obtain c where c: "(\x. (f(f x) - x)/(f x - x)) = (\x. c)" - proof (rule little_Picard_01) - show "(\x. (f(f x) - x)/(f x - x)) holomorphic_on UNIV" - apply (intro holomorphic_intros holf holomorphic_on_compose [unfolded o_def, OF holf]) - using non_fp by auto - qed auto - then obtain "c \ 0" "c \ 1" - by (metis (no_types, hide_lams) non_ff diff_zero divide_eq_0_iff right_inverse_eq) - have eq: "f(f x) - c * f x = x*(1 - c)" for x - using fun_cong [OF c, of x] by (simp add: field_simps) - have df_times_dff: "deriv f z * (deriv f (f z) - c) = 1 * (1 - c)" for z - proof (rule DERIV_unique) - show "((\x. f (f x) - c * f x) has_field_derivative - deriv f z * (deriv f (f z) - c)) (at z)" - apply (intro derivative_eq_intros) - apply (rule DERIV_chain [unfolded o_def, of f]) - apply (auto simp: algebra_simps intro!: holomorphic_derivI [OF holfU]) - done - show "((\x. f (f x) - c * f x) has_field_derivative 1 * (1 - c)) (at z)" - by (simp add: eq mult_commute_abs) - qed - { fix z::complex - obtain k where k: "deriv f \ f = (\x. k)" - proof (rule little_Picard) - show "(deriv f \ f) holomorphic_on UNIV" - by (meson holfU holomorphic_deriv holomorphic_on_compose holomorphic_on_subset open_UNIV subset_UNIV) - obtain "deriv f (f x) \ 0" "deriv f (f x) \ c" for x - using df_times_dff \c \ 1\ eq_iff_diff_eq_0 - by (metis lambda_one mult_zero_left mult_zero_right) - then show "range (deriv f \ f) \ {0,c} = {}" - by force - qed (use \c \ 0\ in auto) - have "\ f constant_on UNIV" - by (meson UNIV_I non_ff constant_on_def) - with holf open_mapping_thm have "open(range f)" - by blast - obtain l where l: "\x. f x - k * x = l" - proof (rule DERIV_zero_connected_constant [of UNIV "{}" "\x. f x - k * x"], simp_all) - have "deriv f w - k = 0" for w - proof (rule analytic_continuation [OF _ open_UNIV connected_UNIV subset_UNIV, of "\z. deriv f z - k" "f z" "range f" w]) - show "(\z. deriv f z - k) holomorphic_on UNIV" - by (intro holomorphic_intros holf open_UNIV) - show "f z islimpt range f" - by (metis (no_types, lifting) IntI UNIV_I \open (range f)\ image_eqI inf.absorb_iff2 inf_aci(1) islimpt_UNIV islimpt_eq_acc_point open_Int top_greatest) - show "\z. z \ range f \ deriv f z - k = 0" - by (metis comp_def diff_self image_iff k) - qed auto - moreover - have "((\x. f x - k * x) has_field_derivative deriv f x - k) (at x)" for x - by (metis DERIV_cmult_Id Deriv.field_differentiable_diff UNIV_I field_differentiable_derivI holf holomorphic_on_def) - ultimately - show "\x. ((\x. f x - k * x) has_field_derivative 0) (at x)" - by auto - show "continuous_on UNIV (\x. f x - k * x)" - by (simp add: continuous_on_diff holf holomorphic_on_imp_continuous_on) - qed (auto simp: connected_UNIV) - have False - proof (cases "k=1") - case True - then have "\x. k * x + l \ a + x" for a - using l non [of a] ext [of f "(+) a"] - by (metis add.commute diff_eq_eq) - with True show ?thesis by auto - next - case False - have "\x. (1 - k) * x \ f 0" - using l [of 0] apply (simp add: algebra_simps) - by (metis diff_add_cancel l mult.commute non_fp) - then show False - by (metis False eq_iff_diff_eq_0 mult.commute nonzero_mult_div_cancel_right times_divide_eq_right) - qed - } - } - then show thesis - using that by blast -qed - - -subsection\The Arzelà--Ascoli theorem\ - -lemma subsequence_diagonalization_lemma: - fixes P :: "nat \ (nat \ 'a) \ bool" - assumes sub: "\i r. \k. strict_mono (k :: nat \ nat) \ P i (r \ k)" - and P_P: "\i r::nat \ 'a. \k1 k2 N. - \P i (r \ k1); \j. N \ j \ \j'. j \ j' \ k2 j = k1 j'\ \ P i (r \ k2)" - obtains k where "strict_mono (k :: nat \ nat)" "\i. P i (r \ k)" -proof - - obtain kk where "\i r. strict_mono (kk i r :: nat \ nat) \ P i (r \ (kk i r))" - using sub by metis - then have sub_kk: "\i r. strict_mono (kk i r)" and P_kk: "\i r. P i (r \ (kk i r))" - by auto - define rr where "rr \ rec_nat (kk 0 r) (\n x. x \ kk (Suc n) (r \ x))" - then have [simp]: "rr 0 = kk 0 r" "\n. rr(Suc n) = rr n \ kk (Suc n) (r \ rr n)" - by auto - show thesis - proof - have sub_rr: "strict_mono (rr i)" for i - using sub_kk by (induction i) (auto simp: strict_mono_def o_def) - have P_rr: "P i (r \ rr i)" for i - using P_kk by (induction i) (auto simp: o_def) - have "i \ i+d \ rr i n \ rr (i+d) n" for d i n - proof (induction d) - case 0 then show ?case - by simp - next - case (Suc d) then show ?case - apply simp - using seq_suble [OF sub_kk] order_trans strict_mono_less_eq [OF sub_rr] by blast - qed - then have "\i j n. i \ j \ rr i n \ rr j n" - by (metis le_iff_add) - show "strict_mono (\n. rr n n)" - apply (simp add: strict_mono_Suc_iff) - by (meson lessI less_le_trans seq_suble strict_monoD sub_kk sub_rr) - have "\j. i \ j \ rr (n+d) i = rr n j" for d n i - apply (induction d arbitrary: i, auto) - by (meson order_trans seq_suble sub_kk) - then have "\m n i. n \ m \ \j. i \ j \ rr m i = rr n j" - by (metis le_iff_add) - then show "P i (r \ (\n. rr n n))" for i - by (meson P_rr P_P) - qed -qed - -lemma function_convergent_subsequence: - fixes f :: "[nat,'a] \ 'b::{real_normed_vector,heine_borel}" - assumes "countable S" and M: "\n::nat. \x. x \ S \ norm(f n x) \ M" - obtains k where "strict_mono (k::nat\nat)" "\x. x \ S \ \l. (\n. f (k n) x) \ l" -proof (cases "S = {}") - case True - then show ?thesis - using strict_mono_id that by fastforce -next - case False - with \countable S\ obtain \ :: "nat \ 'a" where \: "S = range \" - using uncountable_def by blast - obtain k where "strict_mono k" and k: "\i. \l. (\n. (f \ k) n (\ i)) \ l" - proof (rule subsequence_diagonalization_lemma - [of "\i r. \l. ((\n. (f \ r) n (\ i)) \ l) sequentially" id]) - show "\k::nat\nat. strict_mono k \ (\l. (\n. (f \ (r \ k)) n (\ i)) \ l)" for i r - proof - - have "f (r n) (\ i) \ cball 0 M" for n - by (simp add: \ M) - then show ?thesis - using compact_def [of "cball (0::'b) M"] apply simp - apply (drule_tac x="(\n. f (r n) (\ i))" in spec) - apply (force simp: o_def) - done - qed - show "\i r k1 k2 N. - \\l. (\n. (f \ (r \ k1)) n (\ i)) \ l; \j. N \ j \ \j'\j. k2 j = k1 j'\ - \ \l. (\n. (f \ (r \ k2)) n (\ i)) \ l" - apply (simp add: lim_sequentially) - apply (erule ex_forward all_forward imp_forward)+ - apply auto - by (metis (no_types, hide_lams) le_cases order_trans) - qed auto - with \ that show ?thesis - by force -qed - - -theorem Arzela_Ascoli: - fixes \ :: "[nat,'a::euclidean_space] \ 'b::{real_normed_vector,heine_borel}" - assumes "compact S" - and M: "\n x. x \ S \ norm(\ n x) \ M" - and equicont: - "\x e. \x \ S; 0 < e\ - \ \d. 0 < d \ (\n y. y \ S \ norm(x - y) < d \ norm(\ n x - \ n y) < e)" - obtains g k where "continuous_on S g" "strict_mono (k :: nat \ nat)" - "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ norm(\(k n) x - g x) < e" -proof - - have UEQ: "\e. 0 < e \ \d. 0 < d \ (\n. \x \ S. \x' \ S. dist x' x < d \ dist (\ n x') (\ n x) < e)" - apply (rule compact_uniformly_equicontinuous [OF \compact S\, of "range \"]) - using equicont by (force simp: dist_commute dist_norm)+ - have "continuous_on S g" - if "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ norm(\(r n) x - g x) < e" - for g:: "'a \ 'b" and r :: "nat \ nat" - proof (rule uniform_limit_theorem [of _ "\ \ r"]) - show "\\<^sub>F n in sequentially. continuous_on S ((\ \ r) n)" - apply (simp add: eventually_sequentially) - apply (rule_tac x=0 in exI) - using UEQ apply (force simp: continuous_on_iff) - done - show "uniform_limit S (\ \ r) g sequentially" - apply (simp add: uniform_limit_iff eventually_sequentially) - by (metis dist_norm that) - qed auto - moreover - obtain R where "countable R" "R \ S" and SR: "S \ closure R" - by (metis separable that) - obtain k where "strict_mono k" and k: "\x. x \ R \ \l. (\n. \ (k n) x) \ l" - apply (rule function_convergent_subsequence [OF \countable R\ M]) - using \R \ S\ apply force+ - done - then have Cauchy: "Cauchy ((\n. \ (k n) x))" if "x \ R" for x - using convergent_eq_Cauchy that by blast - have "\N. \m n x. N \ m \ N \ n \ x \ S \ dist ((\ \ k) m x) ((\ \ k) n x) < e" - if "0 < e" for e - proof - - obtain d where "0 < d" - and d: "\n. \x \ S. \x' \ S. dist x' x < d \ dist (\ n x') (\ n x) < e/3" - by (metis UEQ \0 < e\ divide_pos_pos zero_less_numeral) - obtain T where "T \ R" and "finite T" and T: "S \ (\c\T. ball c d)" - proof (rule compactE_image [OF \compact S\, of R "(\x. ball x d)"]) - have "closure R \ (\c\R. ball c d)" - apply clarsimp - using \0 < d\ closure_approachable by blast - with SR show "S \ (\c\R. ball c d)" - by auto - qed auto - have "\M. \m\M. \n\M. dist (\ (k m) x) (\ (k n) x) < e/3" if "x \ R" for x - using Cauchy \0 < e\ that unfolding Cauchy_def - by (metis less_divide_eq_numeral1(1) mult_zero_left) - then obtain MF where MF: "\x m n. \x \ R; m \ MF x; n \ MF x\ \ norm (\ (k m) x - \ (k n) x) < e/3" - using dist_norm by metis - have "dist ((\ \ k) m x) ((\ \ k) n x) < e" - if m: "Max (MF ` T) \ m" and n: "Max (MF ` T) \ n" "x \ S" for m n x - proof - - obtain t where "t \ T" and t: "x \ ball t d" - using \x \ S\ T by auto - have "norm(\ (k m) t - \ (k m) x) < e / 3" - by (metis \R \ S\ \T \ R\ \t \ T\ d dist_norm mem_ball subset_iff t \x \ S\) - moreover - have "norm(\ (k n) t - \ (k n) x) < e / 3" - by (metis \R \ S\ \T \ R\ \t \ T\ subsetD d dist_norm mem_ball t \x \ S\) - moreover - have "norm(\ (k m) t - \ (k n) t) < e / 3" - proof (rule MF) - show "t \ R" - using \T \ R\ \t \ T\ by blast - show "MF t \ m" "MF t \ n" - by (meson Max_ge \finite T\ \t \ T\ finite_imageI imageI le_trans m n)+ - qed - ultimately - show ?thesis - unfolding dist_norm [symmetric] o_def - by (metis dist_triangle_third dist_commute) - qed - then show ?thesis - by force - qed - then have "\g. \e>0. \N. \n\N. \x \ S. norm(\(k n) x - g x) < e" - using uniformly_convergent_eq_cauchy [of "\x. x \ S" "\ \ k"] - apply (simp add: o_def dist_norm) - by meson - ultimately show thesis - by (metis that \strict_mono k\) -qed - - - -subsubsection\<^marker>\tag important\\Montel's theorem\ - -text\a sequence of holomorphic functions uniformly bounded -on compact subsets of an open set S has a subsequence that converges to a -holomorphic function, and converges \emph{uniformly} on compact subsets of S.\ - - -theorem Montel: - fixes \ :: "[nat,complex] \ complex" - assumes "open S" - and \: "\h. h \ \ \ h holomorphic_on S" - and bounded: "\K. \compact K; K \ S\ \ \B. \h \ \. \ z \ K. norm(h z) \ B" - and rng_f: "range \ \ \" - obtains g r - where "g holomorphic_on S" "strict_mono (r :: nat \ nat)" - "\x. x \ S \ ((\n. \ (r n) x) \ g x) sequentially" - "\K. \compact K; K \ S\ \ uniform_limit K (\ \ r) g sequentially" -proof - - obtain K where comK: "\n. compact(K n)" and KS: "\n::nat. K n \ S" - and subK: "\X. \compact X; X \ S\ \ \N. \n\N. X \ K n" - using open_Union_compact_subsets [OF \open S\] by metis - then have "\i. \B. \h \ \. \ z \ K i. norm(h z) \ B" - by (simp add: bounded) - then obtain B where B: "\i h z. \h \ \; z \ K i\ \ norm(h z) \ B i" - by metis - have *: "\r g. strict_mono (r::nat\nat) \ (\e > 0. \N. \n\N. \x \ K i. norm((\ \ r) n x - g x) < e)" - if "\n. \ n \ \" for \ i - proof - - obtain g k where "continuous_on (K i) g" "strict_mono (k::nat\nat)" - "\e. 0 < e \ \N. \n\N. \x \ K i. norm(\(k n) x - g x) < e" - proof (rule Arzela_Ascoli [of "K i" "\" "B i"]) - show "\d>0. \n y. y \ K i \ cmod (z - y) < d \ cmod (\ n z - \ n y) < e" - if z: "z \ K i" and "0 < e" for z e - proof - - obtain r where "0 < r" and r: "cball z r \ S" - using z KS [of i] \open S\ by (force simp: open_contains_cball) - have "cball z (2 / 3 * r) \ cball z r" - using \0 < r\ by (simp add: cball_subset_cball_iff) - then have z23S: "cball z (2 / 3 * r) \ S" - using r by blast - obtain M where "0 < M" and M: "\n w. dist z w \ 2/3 * r \ norm(\ n w) \ M" - proof - - obtain N where N: "\n\N. cball z (2/3 * r) \ K n" - using subK compact_cball [of z "(2 / 3 * r)"] z23S by force - have "cmod (\ n w) \ \B N\ + 1" if "dist z w \ 2 / 3 * r" for n w - proof - - have "w \ K N" - using N mem_cball that by blast - then have "cmod (\ n w) \ B N" - using B \\n. \ n \ \\ by blast - also have "... \ \B N\ + 1" - by simp - finally show ?thesis . - qed - then show ?thesis - by (rule_tac M="\B N\ + 1" in that) auto - qed - have "cmod (\ n z - \ n y) < e" - if "y \ K i" and y_near_z: "cmod (z - y) < r/3" "cmod (z - y) < e * r / (6 * M)" - for n y - proof - - have "((\w. \ n w / (w - \)) has_contour_integral - (2 * pi) * \ * winding_number (circlepath z (2 / 3 * r)) \ * \ n \) - (circlepath z (2 / 3 * r))" - if "dist \ z < (2 / 3 * r)" for \ - proof (rule Cauchy_integral_formula_convex_simple) - have "\ n holomorphic_on S" - by (simp add: \ \\n. \ n \ \\) - with z23S show "\ n holomorphic_on cball z (2 / 3 * r)" - using holomorphic_on_subset by blast - qed (use that \0 < r\ in \auto simp: dist_commute\) - then have *: "((\w. \ n w / (w - \)) has_contour_integral (2 * pi) * \ * \ n \) - (circlepath z (2 / 3 * r))" - if "dist \ z < (2 / 3 * r)" for \ - using that by (simp add: winding_number_circlepath dist_norm) - have y: "((\w. \ n w / (w - y)) has_contour_integral (2 * pi) * \ * \ n y) - (circlepath z (2 / 3 * r))" - apply (rule *) - using that \0 < r\ by (simp only: dist_norm norm_minus_commute) - have z: "((\w. \ n w / (w - z)) has_contour_integral (2 * pi) * \ * \ n z) - (circlepath z (2 / 3 * r))" - apply (rule *) - using \0 < r\ by simp - have le_er: "cmod (\ n x / (x - y) - \ n x / (x - z)) \ e / r" - if "cmod (x - z) = r/3 + r/3" for x - proof - - have "\ (cmod (x - y) < r/3)" - using y_near_z(1) that \M > 0\ \r > 0\ - by (metis (full_types) norm_diff_triangle_less norm_minus_commute order_less_irrefl) - then have r4_le_xy: "r/4 \ cmod (x - y)" - using \r > 0\ by simp - then have neq: "x \ y" "x \ z" - using that \r > 0\ by (auto simp: field_split_simps norm_minus_commute) - have leM: "cmod (\ n x) \ M" - by (simp add: M dist_commute dist_norm that) - have "cmod (\ n x / (x - y) - \ n x / (x - z)) = cmod (\ n x) * cmod (1 / (x - y) - 1 / (x - z))" - by (metis (no_types, lifting) divide_inverse mult.left_neutral norm_mult right_diff_distrib') - also have "... = cmod (\ n x) * cmod ((y - z) / ((x - y) * (x - z)))" - using neq by (simp add: field_split_simps) - also have "... = cmod (\ n x) * (cmod (y - z) / (cmod(x - y) * (2/3 * r)))" - by (simp add: norm_mult norm_divide that) - also have "... \ M * (cmod (y - z) / (cmod(x - y) * (2/3 * r)))" - apply (rule mult_mono) - apply (rule leM) - using \r > 0\ \M > 0\ neq by auto - also have "... < M * ((e * r / (6 * M)) / (cmod(x - y) * (2/3 * r)))" - unfolding mult_less_cancel_left - using y_near_z(2) \M > 0\ \r > 0\ neq - apply (simp add: field_simps mult_less_0_iff norm_minus_commute) - done - also have "... \ e/r" - using \e > 0\ \r > 0\ r4_le_xy by (simp add: field_split_simps) - finally show ?thesis by simp - qed - have "(2 * pi) * cmod (\ n y - \ n z) = cmod ((2 * pi) * \ * \ n y - (2 * pi) * \ * \ n z)" - by (simp add: right_diff_distrib [symmetric] norm_mult) - also have "cmod ((2 * pi) * \ * \ n y - (2 * pi) * \ * \ n z) \ e / r * (2 * pi * (2 / 3 * r))" - apply (rule has_contour_integral_bound_circlepath [OF has_contour_integral_diff [OF y z], of "e/r"]) - using \e > 0\ \r > 0\ le_er by auto - also have "... = (2 * pi) * e * ((2 / 3))" - using \r > 0\ by (simp add: field_split_simps) - finally have "cmod (\ n y - \ n z) \ e * (2 / 3)" - by simp - also have "... < e" - using \e > 0\ by simp - finally show ?thesis by (simp add: norm_minus_commute) - qed - then show ?thesis - apply (rule_tac x="min (r/3) ((e * r)/(6 * M))" in exI) - using \0 < e\ \0 < r\ \0 < M\ by simp - qed - show "\n x. x \ K i \ cmod (\ n x) \ B i" - using B \\n. \ n \ \\ by blast - qed (use comK in \fastforce+\) - then show ?thesis - by fastforce - qed - have "\k g. strict_mono (k::nat\nat) \ (\e > 0. \N. \n\N. \x \ K i. norm((\ \ r \ k) n x - g x) < e)" - for i r - apply (rule *) - using rng_f by auto - then have **: "\i r. \k. strict_mono (k::nat\nat) \ (\g. \e>0. \N. \n\N. \x \ K i. norm((\ \ (r \ k)) n x - g x) < e)" - by (force simp: o_assoc) - obtain k :: "nat \ nat" where "strict_mono k" - and "\i. \g. \e>0. \N. \n\N. \x\K i. cmod ((\ \ (id \ k)) n x - g x) < e" - apply (rule subsequence_diagonalization_lemma [OF **, of id]) - apply (erule ex_forward all_forward imp_forward)+ - apply auto - apply (rule_tac x="max N Na" in exI, fastforce+) - done - then have lt_e: "\i. \g. \e>0. \N. \n\N. \x\K i. cmod ((\ \ k) n x - g x) < e" - by simp - have "\l. \e>0. \N. \n\N. norm(\ (k n) z - l) < e" if "z \ S" for z - proof - - obtain G where G: "\i e. e > 0 \ \M. \n\M. \x\K i. cmod ((\ \ k) n x - G i x) < e" - using lt_e by metis - obtain N where "\n. n \ N \ z \ K n" - using subK [of "{z}"] that \z \ S\ by auto - moreover have "\e. e > 0 \ \M. \n\M. \x\K N. cmod ((\ \ k) n x - G N x) < e" - using G by auto - ultimately show ?thesis - by (metis comp_apply order_refl) - qed - then obtain g where g: "\z e. \z \ S; e > 0\ \ \N. \n\N. norm(\ (k n) z - g z) < e" - by metis - show ?thesis - proof - show g_lim: "\x. x \ S \ (\n. \ (k n) x) \ g x" - by (simp add: lim_sequentially g dist_norm) - have dg_le_e: "\N. \n\N. \x\T. cmod (\ (k n) x - g x) < e" - if T: "compact T" "T \ S" and "0 < e" for T e - proof - - obtain N where N: "\n. n \ N \ T \ K n" - using subK [OF T] by blast - obtain h where h: "\e. e>0 \ \M. \n\M. \x\K N. cmod ((\ \ k) n x - h x) < e" - using lt_e by blast - have geq: "g w = h w" if "w \ T" for w - apply (rule LIMSEQ_unique [of "\n. \(k n) w"]) - using \T \ S\ g_lim that apply blast - using h N that by (force simp: lim_sequentially dist_norm) - show ?thesis - using T h N \0 < e\ by (fastforce simp add: geq) - qed - then show "\K. \compact K; K \ S\ - \ uniform_limit K (\ \ k) g sequentially" - by (simp add: uniform_limit_iff dist_norm eventually_sequentially) - show "g holomorphic_on S" - proof (rule holomorphic_uniform_sequence [OF \open S\ \]) - show "\n. (\ \ k) n \ \" - by (simp add: range_subsetD rng_f) - show "\d>0. cball z d \ S \ uniform_limit (cball z d) (\n. (\ \ k) n) g sequentially" - if "z \ S" for z - proof - - obtain d where d: "d>0" "cball z d \ S" - using \open S\ \z \ S\ open_contains_cball by blast - then have "uniform_limit (cball z d) (\ \ k) g sequentially" - using dg_le_e compact_cball by (auto simp: uniform_limit_iff eventually_sequentially dist_norm) - with d show ?thesis by blast - qed - qed - qed (auto simp: \strict_mono k\) -qed - - - -subsection\Some simple but useful cases of Hurwitz's theorem\ - -proposition Hurwitz_no_zeros: - assumes S: "open S" "connected S" - and holf: "\n::nat. \ n holomorphic_on S" - and holg: "g holomorphic_on S" - and ul_g: "\K. \compact K; K \ S\ \ uniform_limit K \ g sequentially" - and nonconst: "\ g constant_on S" - and nz: "\n z. z \ S \ \ n z \ 0" - and "z0 \ S" - shows "g z0 \ 0" -proof - assume g0: "g z0 = 0" - obtain h r m - where "0 < m" "0 < r" and subS: "ball z0 r \ S" - and holh: "h holomorphic_on ball z0 r" - and geq: "\w. w \ ball z0 r \ g w = (w - z0)^m * h w" - and hnz: "\w. w \ ball z0 r \ h w \ 0" - by (blast intro: holomorphic_factor_zero_nonconstant [OF holg S \z0 \ S\ g0 nonconst]) - then have holf0: "\ n holomorphic_on ball z0 r" for n - by (meson holf holomorphic_on_subset) - have *: "((\z. deriv (\ n) z / \ n z) has_contour_integral 0) (circlepath z0 (r/2))" for n - proof (rule Cauchy_theorem_disc_simple [of _ z0 r]) - show "(\z. deriv (\ n) z / \ n z) holomorphic_on ball z0 r" - apply (intro holomorphic_intros holomorphic_deriv holf holf0 open_ball nz) - using \ball z0 r \ S\ by blast - qed (use \0 < r\ in auto) - have hol_dg: "deriv g holomorphic_on S" - by (simp add: \open S\ holg holomorphic_deriv) - have "continuous_on (sphere z0 (r/2)) (deriv g)" - apply (intro holomorphic_on_imp_continuous_on holomorphic_on_subset [OF hol_dg]) - using \0 < r\ subS by auto - then have "compact (deriv g ` (sphere z0 (r/2)))" - by (rule compact_continuous_image [OF _ compact_sphere]) - then have bo_dg: "bounded (deriv g ` (sphere z0 (r/2)))" - using compact_imp_bounded by blast - have "continuous_on (sphere z0 (r/2)) (cmod \ g)" - apply (intro continuous_intros holomorphic_on_imp_continuous_on holomorphic_on_subset [OF holg]) - using \0 < r\ subS by auto - then have "compact ((cmod \ g) ` sphere z0 (r/2))" - by (rule compact_continuous_image [OF _ compact_sphere]) - moreover have "(cmod \ g) ` sphere z0 (r/2) \ {}" - using \0 < r\ by auto - ultimately obtain b where b: "b \ (cmod \ g) ` sphere z0 (r/2)" - "\t. t \ (cmod \ g) ` sphere z0 (r/2) \ b \ t" - using compact_attains_inf [of "(norm \ g) ` (sphere z0 (r/2))"] by blast - have "(\n. contour_integral (circlepath z0 (r/2)) (\z. deriv (\ n) z / \ n z)) \ - contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z)" - proof (rule contour_integral_uniform_limit_circlepath) - show "\\<^sub>F n in sequentially. (\z. deriv (\ n) z / \ n z) contour_integrable_on circlepath z0 (r/2)" - using * contour_integrable_on_def eventually_sequentiallyI by meson - show "uniform_limit (sphere z0 (r/2)) (\n z. deriv (\ n) z / \ n z) (\z. deriv g z / g z) sequentially" - proof (rule uniform_lim_divide [OF _ _ bo_dg]) - show "uniform_limit (sphere z0 (r/2)) (\a. deriv (\ a)) (deriv g) sequentially" - proof (rule uniform_limitI) - fix e::real - assume "0 < e" - have *: "dist (deriv (\ n) w) (deriv g w) < e" - if e8: "\x. dist z0 x \ 3 * r / 4 \ dist (\ n x) (g x) * 8 < r * e" - and w: "dist w z0 = r/2" for n w - proof - - have "ball w (r/4) \ ball z0 r" "cball w (r/4) \ ball z0 r" - using \0 < r\ by (simp_all add: ball_subset_ball_iff cball_subset_ball_iff w) - with subS have wr4_sub: "ball w (r/4) \ S" "cball w (r/4) \ S" by force+ - moreover - have "(\z. \ n z - g z) holomorphic_on S" - by (intro holomorphic_intros holf holg) - ultimately have hol: "(\z. \ n z - g z) holomorphic_on ball w (r/4)" - and cont: "continuous_on (cball w (r / 4)) (\z. \ n z - g z)" - using holomorphic_on_subset by (blast intro: holomorphic_on_imp_continuous_on)+ - have "w \ S" - using \0 < r\ wr4_sub by auto - have "\y. dist w y < r / 4 \ dist z0 y \ 3 * r / 4" - apply (rule dist_triangle_le [where z=w]) - using w by (simp add: dist_commute) - with e8 have in_ball: "\y. y \ ball w (r/4) \ \ n y - g y \ ball 0 (r/4 * e/2)" - by (simp add: dist_norm [symmetric]) - have "\ n field_differentiable at w" - by (metis holomorphic_on_imp_differentiable_at \w \ S\ holf \open S\) - moreover - have "g field_differentiable at w" - using \w \ S\ \open S\ holg holomorphic_on_imp_differentiable_at by auto - moreover - have "cmod (deriv (\w. \ n w - g w) w) * 2 \ e" - apply (rule Cauchy_higher_deriv_bound [OF hol cont in_ball, of 1, simplified]) - using \r > 0\ by auto - ultimately have "dist (deriv (\ n) w) (deriv g w) \ e/2" - by (simp add: dist_norm) - then show ?thesis - using \e > 0\ by auto - qed - have "cball z0 (3 * r / 4) \ ball z0 r" - by (simp add: cball_subset_ball_iff \0 < r\) - with subS have "uniform_limit (cball z0 (3 * r/4)) \ g sequentially" - by (force intro: ul_g) - then have "\\<^sub>F n in sequentially. \x\cball z0 (3 * r / 4). dist (\ n x) (g x) < r / 4 * e / 2" - using \0 < e\ \0 < r\ by (force simp: intro!: uniform_limitD) - then show "\\<^sub>F n in sequentially. \x \ sphere z0 (r/2). dist (deriv (\ n) x) (deriv g x) < e" - apply (simp add: eventually_sequentially) - apply (elim ex_forward all_forward imp_forward asm_rl) - using * apply (force simp: dist_commute) - done - qed - show "uniform_limit (sphere z0 (r/2)) \ g sequentially" - proof (rule uniform_limitI) - fix e::real - assume "0 < e" - have "sphere z0 (r/2) \ ball z0 r" - using \0 < r\ by auto - with subS have "uniform_limit (sphere z0 (r/2)) \ g sequentially" - by (force intro: ul_g) - then show "\\<^sub>F n in sequentially. \x \ sphere z0 (r/2). dist (\ n x) (g x) < e" - apply (rule uniform_limitD) - using \0 < e\ by force - qed - show "b > 0" "\x. x \ sphere z0 (r/2) \ b \ cmod (g x)" - using b \0 < r\ by (fastforce simp: geq hnz)+ - qed - qed (use \0 < r\ in auto) - then have "(\n. 0) \ contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z)" - by (simp add: contour_integral_unique [OF *]) - then have "contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z) = 0" - by (simp add: LIMSEQ_const_iff) - moreover - have "contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z) = - contour_integral (circlepath z0 (r/2)) (\z. m / (z - z0) + deriv h z / h z)" - proof (rule contour_integral_eq, use \0 < r\ in simp) - fix w - assume w: "dist z0 w * 2 = r" - then have w_inb: "w \ ball z0 r" - using \0 < r\ by auto - have h_der: "(h has_field_derivative deriv h w) (at w)" - using holh holomorphic_derivI w_inb by blast - have "deriv g w = ((of_nat m * h w + deriv h w * (w - z0)) * (w - z0) ^ m) / (w - z0)" - if "r = dist z0 w * 2" "w \ z0" - proof - - have "((\w. (w - z0) ^ m * h w) has_field_derivative - (m * h w + deriv h w * (w - z0)) * (w - z0) ^ m / (w - z0)) (at w)" - apply (rule derivative_eq_intros h_der refl)+ - using that \m > 0\ \0 < r\ apply (simp add: divide_simps distrib_right) - apply (metis Suc_pred mult.commute power_Suc) - done - then show ?thesis - apply (rule DERIV_imp_deriv [OF has_field_derivative_transform_within_open [where S = "ball z0 r"]]) - using that \m > 0\ \0 < r\ - apply (simp_all add: hnz geq) - done - qed - with \0 < r\ \0 < m\ w w_inb show "deriv g w / g w = of_nat m / (w - z0) + deriv h w / h w" - by (auto simp: geq field_split_simps hnz) - qed - moreover - have "contour_integral (circlepath z0 (r/2)) (\z. m / (z - z0) + deriv h z / h z) = - 2 * of_real pi * \ * m + 0" - proof (rule contour_integral_unique [OF has_contour_integral_add]) - show "((\x. m / (x - z0)) has_contour_integral 2 * of_real pi * \ * m) (circlepath z0 (r/2))" - by (force simp: \0 < r\ intro: Cauchy_integral_circlepath_simple) - show "((\x. deriv h x / h x) has_contour_integral 0) (circlepath z0 (r/2))" - apply (rule Cauchy_theorem_disc_simple [of _ z0 r]) - using hnz holh holomorphic_deriv holomorphic_on_divide \0 < r\ - apply force+ - done - qed - ultimately show False using \0 < m\ by auto -qed - -corollary Hurwitz_injective: - assumes S: "open S" "connected S" - and holf: "\n::nat. \ n holomorphic_on S" - and holg: "g holomorphic_on S" - and ul_g: "\K. \compact K; K \ S\ \ uniform_limit K \ g sequentially" - and nonconst: "\ g constant_on S" - and inj: "\n. inj_on (\ n) S" - shows "inj_on g S" -proof - - have False if z12: "z1 \ S" "z2 \ S" "z1 \ z2" "g z2 = g z1" for z1 z2 - proof - - obtain z0 where "z0 \ S" and z0: "g z0 \ g z2" - using constant_on_def nonconst by blast - have "(\z. g z - g z1) holomorphic_on S" - by (intro holomorphic_intros holg) - then obtain r where "0 < r" "ball z2 r \ S" "\z. dist z2 z < r \ z \ z2 \ g z \ g z1" - apply (rule isolated_zeros [of "\z. g z - g z1" S z2 z0]) - using S \z0 \ S\ z0 z12 by auto - have "g z2 - g z1 \ 0" - proof (rule Hurwitz_no_zeros [of "S - {z1}" "\n z. \ n z - \ n z1" "\z. g z - g z1"]) - show "open (S - {z1})" - by (simp add: S open_delete) - show "connected (S - {z1})" - by (simp add: connected_open_delete [OF S]) - show "\n. (\z. \ n z - \ n z1) holomorphic_on S - {z1}" - by (intro holomorphic_intros holomorphic_on_subset [OF holf]) blast - show "(\z. g z - g z1) holomorphic_on S - {z1}" - by (intro holomorphic_intros holomorphic_on_subset [OF holg]) blast - show "uniform_limit K (\n z. \ n z - \ n z1) (\z. g z - g z1) sequentially" - if "compact K" "K \ S - {z1}" for K - proof (rule uniform_limitI) - fix e::real - assume "e > 0" - have "uniform_limit K \ g sequentially" - using that ul_g by fastforce - then have K: "\\<^sub>F n in sequentially. \x \ K. dist (\ n x) (g x) < e/2" - using \0 < e\ by (force simp: intro!: uniform_limitD) - have "uniform_limit {z1} \ g sequentially" - by (simp add: ul_g z12) - then have "\\<^sub>F n in sequentially. \x \ {z1}. dist (\ n x) (g x) < e/2" - using \0 < e\ by (force simp: intro!: uniform_limitD) - then have z1: "\\<^sub>F n in sequentially. dist (\ n z1) (g z1) < e/2" - by simp - have "\\<^sub>F n in sequentially. \x\K. dist (\ n x - \ n z1) (g x - g z1) < e/2 + e/2" - apply (rule eventually_mono [OF eventually_conj [OF K z1]]) - apply (simp add: dist_norm algebra_simps del: divide_const_simps) - by (metis add.commute dist_commute dist_norm dist_triangle_add_half) - have "\\<^sub>F n in sequentially. \x\K. dist (\ n x - \ n z1) (g x - g z1) < e/2 + e/2" - using eventually_conj [OF K z1] - apply (rule eventually_mono) - by (metis (no_types, hide_lams) diff_add_eq diff_diff_eq2 dist_commute dist_norm dist_triangle_add_half field_sum_of_halves) - then - show "\\<^sub>F n in sequentially. \x\K. dist (\ n x - \ n z1) (g x - g z1) < e" - by simp - qed - show "\ (\z. g z - g z1) constant_on S - {z1}" - unfolding constant_on_def - by (metis Diff_iff \z0 \ S\ empty_iff insert_iff right_minus_eq z0 z12) - show "\n z. z \ S - {z1} \ \ n z - \ n z1 \ 0" - by (metis DiffD1 DiffD2 eq_iff_diff_eq_0 inj inj_onD insertI1 \z1 \ S\) - show "z2 \ S - {z1}" - using \z2 \ S\ \z1 \ z2\ by auto - qed - with z12 show False by auto - qed - then show ?thesis by (auto simp: inj_on_def) -qed - - - -subsection\The Great Picard theorem\ - -lemma GPicard1: - assumes S: "open S" "connected S" and "w \ S" "0 < r" "Y \ X" - and holX: "\h. h \ X \ h holomorphic_on S" - and X01: "\h z. \h \ X; z \ S\ \ h z \ 0 \ h z \ 1" - and r: "\h. h \ Y \ norm(h w) \ r" - obtains B Z where "0 < B" "open Z" "w \ Z" "Z \ S" "\h z. \h \ Y; z \ Z\ \ norm(h z) \ B" -proof - - obtain e where "e > 0" and e: "cball w e \ S" - using assms open_contains_cball_eq by blast - show ?thesis - proof - show "0 < exp(pi * exp(pi * (2 + 2 * r + 12)))" - by simp - show "ball w (e / 2) \ S" - using e ball_divide_subset_numeral ball_subset_cball by blast - show "cmod (h z) \ exp (pi * exp (pi * (2 + 2 * r + 12)))" - if "h \ Y" "z \ ball w (e / 2)" for h z - proof - - have "h \ X" - using \Y \ X\ \h \ Y\ by blast - with holX have "h holomorphic_on S" - by auto - then have "h holomorphic_on cball w e" - by (metis e holomorphic_on_subset) - then have hol_h_o: "(h \ (\z. (w + of_real e * z))) holomorphic_on cball 0 1" - apply (intro holomorphic_intros holomorphic_on_compose) - apply (erule holomorphic_on_subset) - using that \e > 0\ by (auto simp: dist_norm norm_mult) - have norm_le_r: "cmod ((h \ (\z. w + complex_of_real e * z)) 0) \ r" - by (auto simp: r \h \ Y\) - have le12: "norm (of_real(inverse e) * (z - w)) \ 1/2" - using that \e > 0\ by (simp add: inverse_eq_divide dist_norm norm_minus_commute norm_divide) - have non01: "\z::complex. cmod z \ 1 \ h (w + e * z) \ 0 \ h (w + e * z) \ 1" - apply (rule X01 [OF \h \ X\]) - apply (rule subsetD [OF e]) - using \0 < e\ by (auto simp: dist_norm norm_mult) - have "cmod (h z) \ cmod (h (w + of_real e * (inverse e * (z - w))))" - using \0 < e\ by (simp add: field_split_simps) - also have "... \ exp (pi * exp (pi * (14 + 2 * r)))" - using r [OF \h \ Y\] Schottky [OF hol_h_o norm_le_r _ _ _ le12] non01 by auto - finally - show ?thesis by simp - qed - qed (use \e > 0\ in auto) -qed - -lemma GPicard2: - assumes "S \ T" "connected T" "S \ {}" "open S" "\x. \x islimpt S; x \ T\ \ x \ S" - shows "S = T" - by (metis assms open_subset connected_clopen closedin_limpt) - - -lemma GPicard3: - assumes S: "open S" "connected S" "w \ S" and "Y \ X" - and holX: "\h. h \ X \ h holomorphic_on S" - and X01: "\h z. \h \ X; z \ S\ \ h z \ 0 \ h z \ 1" - and no_hw_le1: "\h. h \ Y \ norm(h w) \ 1" - and "compact K" "K \ S" - obtains B where "\h z. \h \ Y; z \ K\ \ norm(h z) \ B" -proof - - define U where "U \ {z \ S. \B Z. 0 < B \ open Z \ z \ Z \ Z \ S \ - (\h z'. h \ Y \ z' \ Z \ norm(h z') \ B)}" - then have "U \ S" by blast - have "U = S" - proof (rule GPicard2 [OF \U \ S\ \connected S\]) - show "U \ {}" - proof - - obtain B Z where "0 < B" "open Z" "w \ Z" "Z \ S" - and "\h z. \h \ Y; z \ Z\ \ norm(h z) \ B" - apply (rule GPicard1 [OF S zero_less_one \Y \ X\ holX]) - using no_hw_le1 X01 by force+ - then show ?thesis - unfolding U_def using \w \ S\ by blast - qed - show "open U" - unfolding open_subopen [of U] by (auto simp: U_def) - fix v - assume v: "v islimpt U" "v \ S" - have "\ (\r>0. \h\Y. r < cmod (h v))" - proof - assume "\r>0. \h\Y. r < cmod (h v)" - then have "\n. \h\Y. Suc n < cmod (h v)" - by simp - then obtain \ where FY: "\n. \ n \ Y" and ltF: "\n. Suc n < cmod (\ n v)" - by metis - define \ where "\ \ \n z. inverse(\ n z)" - have hol\: "\ n holomorphic_on S" for n - apply (simp add: \_def) - using FY X01 \Y \ X\ holX apply (blast intro: holomorphic_on_inverse) - done - have \not0: "\ n z \ 0" and \not1: "\ n z \ 1" if "z \ S" for n z - using FY X01 \Y \ X\ that by (force simp: \_def)+ - have \_le1: "cmod (\ n v) \ 1" for n - using less_le_trans linear ltF - by (fastforce simp add: \_def norm_inverse inverse_le_1_iff) - define W where "W \ {h. h holomorphic_on S \ (\z \ S. h z \ 0 \ h z \ 1)}" - obtain B Z where "0 < B" "open Z" "v \ Z" "Z \ S" - and B: "\h z. \h \ range \; z \ Z\ \ norm(h z) \ B" - apply (rule GPicard1 [OF \open S\ \connected S\ \v \ S\ zero_less_one, of "range \" W]) - using hol\ \not0 \not1 \_le1 by (force simp: W_def)+ - then obtain e where "e > 0" and e: "ball v e \ Z" - by (meson open_contains_ball) - obtain h j where holh: "h holomorphic_on ball v e" and "strict_mono j" - and lim: "\x. x \ ball v e \ (\n. \ (j n) x) \ h x" - and ulim: "\K. \compact K; K \ ball v e\ - \ uniform_limit K (\ \ j) h sequentially" - proof (rule Montel) - show "\h. h \ range \ \ h holomorphic_on ball v e" - by (metis \Z \ S\ e hol\ holomorphic_on_subset imageE) - show "\K. \compact K; K \ ball v e\ \ \B. \h\range \. \z\K. cmod (h z) \ B" - using B e by blast - qed auto - have "h v = 0" - proof (rule LIMSEQ_unique) - show "(\n. \ (j n) v) \ h v" - using \e > 0\ lim by simp - have lt_Fj: "real x \ cmod (\ (j x) v)" for x - by (metis of_nat_Suc ltF \strict_mono j\ add.commute less_eq_real_def less_le_trans nat_le_real_less seq_suble) - show "(\n. \ (j n) v) \ 0" - proof (rule Lim_null_comparison [OF eventually_sequentiallyI lim_inverse_n]) - show "cmod (\ (j x) v) \ inverse (real x)" if "1 \ x" for x - using that by (simp add: \_def norm_inverse_le_norm [OF lt_Fj]) - qed - qed - have "h v \ 0" - proof (rule Hurwitz_no_zeros [of "ball v e" "\ \ j" h]) - show "\n. (\ \ j) n holomorphic_on ball v e" - using \Z \ S\ e hol\ by force - show "\n z. z \ ball v e \ (\ \ j) n z \ 0" - using \not0 \Z \ S\ e by fastforce - show "\ h constant_on ball v e" - proof (clarsimp simp: constant_on_def) - fix c - have False if "\z. dist v z < e \ h z = c" - proof - - have "h v = c" - by (simp add: \0 < e\ that) - obtain y where "y \ U" "y \ v" and y: "dist y v < e" - using v \e > 0\ by (auto simp: islimpt_approachable) - then obtain C T where "y \ S" "C > 0" "open T" "y \ T" "T \ S" - and "\h z'. \h \ Y; z' \ T\ \ cmod (h z') \ C" - using \y \ U\ by (auto simp: U_def) - then have le_C: "\n. cmod (\ n y) \ C" - using FY by blast - have "\\<^sub>F n in sequentially. dist (\ (j n) y) (h y) < inverse C" - using uniform_limitD [OF ulim [of "{y}"], of "inverse C"] \C > 0\ y - by (simp add: dist_commute) - then obtain n where "dist (\ (j n) y) (h y) < inverse C" - by (meson eventually_at_top_linorder order_refl) - moreover - have "h y = h v" - by (metis \h v = c\ dist_commute that y) - ultimately have "norm (\ (j n) y) < inverse C" - by (simp add: \h v = 0\) - then have "C < norm (\ (j n) y)" - apply (simp add: \_def) - by (metis FY X01 \0 < C\ \y \ S\ \Y \ X\ inverse_less_iff_less norm_inverse subsetD zero_less_norm_iff) - show False - using \C < cmod (\ (j n) y)\ le_C not_less by blast - qed - then show "\x\ball v e. h x \ c" by force - qed - show "h holomorphic_on ball v e" - by (simp add: holh) - show "\K. \compact K; K \ ball v e\ \ uniform_limit K (\ \ j) h sequentially" - by (simp add: ulim) - qed (use \e > 0\ in auto) - with \h v = 0\ show False by blast - qed - then show "v \ U" - apply (clarsimp simp add: U_def v) - apply (rule GPicard1[OF \open S\ \connected S\ \v \ S\ _ \Y \ X\ holX]) - using X01 no_hw_le1 apply (meson | force simp: not_less)+ - done - qed - have "\x. x \ K \ x \ U" - using \U = S\ \K \ S\ by blast - then have "\x. x \ K \ (\B Z. 0 < B \ open Z \ x \ Z \ - (\h z'. h \ Y \ z' \ Z \ norm(h z') \ B))" - unfolding U_def by blast - then obtain F Z where F: "\x. x \ K \ open (Z x) \ x \ Z x \ - (\h z'. h \ Y \ z' \ Z x \ norm(h z') \ F x)" - by metis - then obtain L where "L \ K" "finite L" and L: "K \ (\c \ L. Z c)" - by (auto intro: compactE_image [OF \compact K\, of K Z]) - then have *: "\x h z'. \x \ L; h \ Y \ z' \ Z x\ \ cmod (h z') \ F x" - using F by blast - have "\B. \h z. h \ Y \ z \ K \ norm(h z) \ B" - proof (cases "L = {}") - case True with L show ?thesis by simp - next - case False - with \finite L\ show ?thesis - apply (rule_tac x = "Max (F ` L)" in exI) - apply (simp add: linorder_class.Max_ge_iff) - using * F by (metis L UN_E subsetD) - qed - with that show ?thesis by metis -qed - - -lemma GPicard4: - assumes "0 < k" and holf: "f holomorphic_on (ball 0 k - {0})" - and AE: "\e. \0 < e; e < k\ \ \d. 0 < d \ d < e \ (\z \ sphere 0 d. norm(f z) \ B)" - obtains \ where "0 < \" "\ < k" "\z. z \ ball 0 \ - {0} \ norm(f z) \ B" -proof - - obtain \ where "0 < \" "\ < k/2" and \: "\z. norm z = \ \ norm(f z) \ B" - using AE [of "k/2"] \0 < k\ by auto - show ?thesis - proof - show "\ < k" - using \0 < k\ \\ < k/2\ by auto - show "cmod (f \) \ B" if \: "\ \ ball 0 \ - {0}" for \ - proof - - obtain d where "0 < d" "d < norm \" and d: "\z. norm z = d \ norm(f z) \ B" - using AE [of "norm \"] \\ < k\ \ by auto - have [simp]: "closure (cball 0 \ - ball 0 d) = cball 0 \ - ball 0 d" - by (blast intro!: closure_closed) - have [simp]: "interior (cball 0 \ - ball 0 d) = ball 0 \ - cball (0::complex) d" - using \0 < \\ \0 < d\ by (simp add: interior_diff) - have *: "norm(f w) \ B" if "w \ cball 0 \ - ball 0 d" for w - proof (rule maximum_modulus_frontier [of f "cball 0 \ - ball 0 d"]) - show "f holomorphic_on interior (cball 0 \ - ball 0 d)" - apply (rule holomorphic_on_subset [OF holf]) - using \\ < k\ \0 < d\ that by auto - show "continuous_on (closure (cball 0 \ - ball 0 d)) f" - apply (rule holomorphic_on_imp_continuous_on) - apply (rule holomorphic_on_subset [OF holf]) - using \0 < d\ \\ < k\ by auto - show "\z. z \ frontier (cball 0 \ - ball 0 d) \ cmod (f z) \ B" - apply (simp add: frontier_def) - using \ d less_eq_real_def by blast - qed (use that in auto) - show ?thesis - using * \d < cmod \\ that by auto - qed - qed (use \0 < \\ in auto) -qed - - -lemma GPicard5: - assumes holf: "f holomorphic_on (ball 0 1 - {0})" - and f01: "\z. z \ ball 0 1 - {0} \ f z \ 0 \ f z \ 1" - obtains e B where "0 < e" "e < 1" "0 < B" - "(\z \ ball 0 e - {0}. norm(f z) \ B) \ - (\z \ ball 0 e - {0}. norm(f z) \ B)" -proof - - have [simp]: "1 + of_nat n \ (0::complex)" for n - using of_nat_eq_0_iff by fastforce - have [simp]: "cmod (1 + of_nat n) = 1 + of_nat n" for n - by (metis norm_of_nat of_nat_Suc) - have *: "(\x::complex. x / of_nat (Suc n)) ` (ball 0 1 - {0}) \ ball 0 1 - {0}" for n - by (auto simp: norm_divide field_split_simps split: if_split_asm) - define h where "h \ \n z::complex. f (z / (Suc n))" - have holh: "(h n) holomorphic_on ball 0 1 - {0}" for n - unfolding h_def - proof (rule holomorphic_on_compose_gen [unfolded o_def, OF _ holf *]) - show "(\x. x / of_nat (Suc n)) holomorphic_on ball 0 1 - {0}" - by (intro holomorphic_intros) auto - qed - have h01: "\n z. z \ ball 0 1 - {0} \ h n z \ 0 \ h n z \ 1" - unfolding h_def - apply (rule f01) - using * by force - obtain w where w: "w \ ball 0 1 - {0::complex}" - by (rule_tac w = "1/2" in that) auto - consider "infinite {n. norm(h n w) \ 1}" | "infinite {n. 1 \ norm(h n w)}" - by (metis (mono_tags, lifting) infinite_nat_iff_unbounded_le le_cases mem_Collect_eq) - then show ?thesis - proof cases - case 1 - with infinite_enumerate obtain r :: "nat \ nat" - where "strict_mono r" and r: "\n. r n \ {n. norm(h n w) \ 1}" - by blast - obtain B where B: "\j z. \norm z = 1/2; j \ range (h \ r)\ \ norm(j z) \ B" - proof (rule GPicard3 [OF _ _ w, where K = "sphere 0 (1/2)"]) - show "range (h \ r) \ - {g. g holomorphic_on ball 0 1 - {0} \ (\z\ball 0 1 - {0}. g z \ 0 \ g z \ 1)}" - apply clarsimp - apply (intro conjI holomorphic_intros holomorphic_on_compose holh) - using h01 apply auto - done - show "connected (ball 0 1 - {0::complex})" - by (simp add: connected_open_delete) - qed (use r in auto) - have normf_le_B: "cmod(f z) \ B" if "norm z = 1 / (2 * (1 + of_nat (r n)))" for z n - proof - - have *: "\w. norm w = 1/2 \ cmod((f (w / (1 + of_nat (r n))))) \ B" - using B by (auto simp: h_def o_def) - have half: "norm (z * (1 + of_nat (r n))) = 1/2" - by (simp add: norm_mult divide_simps that) - show ?thesis - using * [OF half] by simp - qed - obtain \ where "0 < \" "\ < 1" "\z. z \ ball 0 \ - {0} \ cmod(f z) \ B" - proof (rule GPicard4 [OF zero_less_one holf, of B]) - fix e::real - assume "0 < e" "e < 1" - obtain n where "(1/e - 2) / 2 < real n" - using reals_Archimedean2 by blast - also have "... \ r n" - using \strict_mono r\ by (simp add: seq_suble) - finally have "(1/e - 2) / 2 < real (r n)" . - with \0 < e\ have e: "e > 1 / (2 + 2 * real (r n))" - by (simp add: field_simps) - show "\d>0. d < e \ (\z\sphere 0 d. cmod (f z) \ B)" - apply (rule_tac x="1 / (2 * (1 + of_nat (r n)))" in exI) - using normf_le_B by (simp add: e) - qed blast - then have \: "cmod (f z) \ \B\ + 1" if "cmod z < \" "z \ 0" for z - using that by fastforce - have "0 < \B\ + 1" - by simp - then show ?thesis - apply (rule that [OF \0 < \\ \\ < 1\]) - using \ by auto - next - case 2 - with infinite_enumerate obtain r :: "nat \ nat" - where "strict_mono r" and r: "\n. r n \ {n. norm(h n w) \ 1}" - by blast - obtain B where B: "\j z. \norm z = 1/2; j \ range (\n. inverse \ h (r n))\ \ norm(j z) \ B" - proof (rule GPicard3 [OF _ _ w, where K = "sphere 0 (1/2)"]) - show "range (\n. inverse \ h (r n)) \ - {g. g holomorphic_on ball 0 1 - {0} \ (\z\ball 0 1 - {0}. g z \ 0 \ g z \ 1)}" - apply clarsimp - apply (intro conjI holomorphic_intros holomorphic_on_compose_gen [unfolded o_def, OF _ holh] holomorphic_on_compose) - using h01 apply auto - done - show "connected (ball 0 1 - {0::complex})" - by (simp add: connected_open_delete) - show "\j. j \ range (\n. inverse \ h (r n)) \ cmod (j w) \ 1" - using r norm_inverse_le_norm by fastforce - qed (use r in auto) - have norm_if_le_B: "cmod(inverse (f z)) \ B" if "norm z = 1 / (2 * (1 + of_nat (r n)))" for z n - proof - - have *: "inverse (cmod((f (z / (1 + of_nat (r n)))))) \ B" if "norm z = 1/2" for z - using B [OF that] by (force simp: norm_inverse h_def) - have half: "norm (z * (1 + of_nat (r n))) = 1/2" - by (simp add: norm_mult divide_simps that) - show ?thesis - using * [OF half] by (simp add: norm_inverse) - qed - have hol_if: "(inverse \ f) holomorphic_on (ball 0 1 - {0})" - by (metis (no_types, lifting) holf comp_apply f01 holomorphic_on_inverse holomorphic_transform) - obtain \ where "0 < \" "\ < 1" and leB: "\z. z \ ball 0 \ - {0} \ cmod((inverse \ f) z) \ B" - proof (rule GPicard4 [OF zero_less_one hol_if, of B]) - fix e::real - assume "0 < e" "e < 1" - obtain n where "(1/e - 2) / 2 < real n" - using reals_Archimedean2 by blast - also have "... \ r n" - using \strict_mono r\ by (simp add: seq_suble) - finally have "(1/e - 2) / 2 < real (r n)" . - with \0 < e\ have e: "e > 1 / (2 + 2 * real (r n))" - by (simp add: field_simps) - show "\d>0. d < e \ (\z\sphere 0 d. cmod ((inverse \ f) z) \ B)" - apply (rule_tac x="1 / (2 * (1 + of_nat (r n)))" in exI) - using norm_if_le_B by (simp add: e) - qed blast - have \: "cmod (f z) \ inverse B" and "B > 0" if "cmod z < \" "z \ 0" for z - proof - - have "inverse (cmod (f z)) \ B" - using leB that by (simp add: norm_inverse) - moreover - have "f z \ 0" - using \\ < 1\ f01 that by auto - ultimately show "cmod (f z) \ inverse B" - by (simp add: norm_inverse inverse_le_imp_le) - show "B > 0" - using \f z \ 0\ \inverse (cmod (f z)) \ B\ not_le order.trans by fastforce - qed - then have "B > 0" - by (metis \0 < \\ dense leI order.asym vector_choose_size) - then have "inverse B > 0" - by (simp add: field_split_simps) - then show ?thesis - apply (rule that [OF \0 < \\ \\ < 1\]) - using \ by auto - qed -qed - - -lemma GPicard6: - assumes "open M" "z \ M" "a \ 0" and holf: "f holomorphic_on (M - {z})" - and f0a: "\w. w \ M - {z} \ f w \ 0 \ f w \ a" - obtains r where "0 < r" "ball z r \ M" - "bounded(f ` (ball z r - {z})) \ - bounded((inverse \ f) ` (ball z r - {z}))" -proof - - obtain r where "0 < r" and r: "ball z r \ M" - using assms openE by blast - let ?g = "\w. f (z + of_real r * w) / a" - obtain e B where "0 < e" "e < 1" "0 < B" - and B: "(\z \ ball 0 e - {0}. norm(?g z) \ B) \ (\z \ ball 0 e - {0}. norm(?g z) \ B)" - proof (rule GPicard5) - show "?g holomorphic_on ball 0 1 - {0}" - apply (intro holomorphic_intros holomorphic_on_compose_gen [unfolded o_def, OF _ holf]) - using \0 < r\ \a \ 0\ r - by (auto simp: dist_norm norm_mult subset_eq) - show "\w. w \ ball 0 1 - {0} \ f (z + of_real r * w) / a \ 0 \ f (z + of_real r * w) / a \ 1" - apply (simp add: field_split_simps \a \ 0\) - apply (rule f0a) - using \0 < r\ r by (auto simp: dist_norm norm_mult subset_eq) - qed - show ?thesis - proof - show "0 < e*r" - by (simp add: \0 < e\ \0 < r\) - have "ball z (e * r) \ ball z r" - by (simp add: \0 < r\ \e < 1\ order.strict_implies_order subset_ball) - then show "ball z (e * r) \ M" - using r by blast - consider "\z. z \ ball 0 e - {0} \ norm(?g z) \ B" | "\z. z \ ball 0 e - {0} \ norm(?g z) \ B" - using B by blast - then show "bounded (f ` (ball z (e * r) - {z})) \ - bounded ((inverse \ f) ` (ball z (e * r) - {z}))" - proof cases - case 1 - have "\dist z w < e * r; w \ z\ \ cmod (f w) \ B * norm a" for w - using \a \ 0\ \0 < r\ 1 [of "(w - z) / r"] - by (simp add: norm_divide dist_norm field_split_simps) - then show ?thesis - by (force simp: intro!: boundedI) - next - case 2 - have "\dist z w < e * r; w \ z\ \ cmod (f w) \ B * norm a" for w - using \a \ 0\ \0 < r\ 2 [of "(w - z) / r"] - by (simp add: norm_divide dist_norm field_split_simps) - then have "\dist z w < e * r; w \ z\ \ inverse (cmod (f w)) \ inverse (B * norm a)" for w - by (metis \0 < B\ \a \ 0\ mult_pos_pos norm_inverse norm_inverse_le_norm zero_less_norm_iff) - then show ?thesis - by (force simp: norm_inverse intro!: boundedI) - qed - qed -qed - - -theorem great_Picard: - assumes "open M" "z \ M" "a \ b" and holf: "f holomorphic_on (M - {z})" - and fab: "\w. w \ M - {z} \ f w \ a \ f w \ b" - obtains l where "(f \ l) (at z) \ ((inverse \ f) \ l) (at z)" -proof - - obtain r where "0 < r" and zrM: "ball z r \ M" - and r: "bounded((\z. f z - a) ` (ball z r - {z})) \ - bounded((inverse \ (\z. f z - a)) ` (ball z r - {z}))" - proof (rule GPicard6 [OF \open M\ \z \ M\]) - show "b - a \ 0" - using assms by auto - show "(\z. f z - a) holomorphic_on M - {z}" - by (intro holomorphic_intros holf) - qed (use fab in auto) - have holfb: "f holomorphic_on ball z r - {z}" - apply (rule holomorphic_on_subset [OF holf]) - using zrM by auto - have holfb_i: "(\z. inverse(f z - a)) holomorphic_on ball z r - {z}" - apply (intro holomorphic_intros holfb) - using fab zrM by fastforce - show ?thesis - using r - proof - assume "bounded ((\z. f z - a) ` (ball z r - {z}))" - then obtain B where B: "\w. w \ (\z. f z - a) ` (ball z r - {z}) \ norm w \ B" - by (force simp: bounded_iff) - have "\\<^sub>F w in at z. cmod (f w - a) \ B" - apply (simp add: eventually_at) - apply (rule_tac x=r in exI) - using \0 < r\ by (auto simp: dist_commute intro!: B) - then have "\B. \\<^sub>F w in at z. cmod (f w) \ B" - apply (rule_tac x="B + norm a" in exI) - apply (erule eventually_mono) - by (metis add.commute add_le_cancel_right norm_triangle_sub order.trans) - then obtain g where holg: "g holomorphic_on ball z r" and gf: "\w. w \ ball z r - {z} \ g w = f w" - using \0 < r\ holomorphic_on_extend_bounded [OF holfb] by auto - then have "g \z\ g z" - apply (simp add: continuous_at [symmetric]) - using \0 < r\ centre_in_ball field_differentiable_imp_continuous_at holomorphic_on_imp_differentiable_at by blast - then have "(f \ g z) (at z)" - apply (rule Lim_transform_within_open [of g "g z" z UNIV "ball z r"]) - using \0 < r\ by (auto simp: gf) - then show ?thesis - using that by blast - next - assume "bounded((inverse \ (\z. f z - a)) ` (ball z r - {z}))" - then obtain B where B: "\w. w \ (inverse \ (\z. f z - a)) ` (ball z r - {z}) \ norm w \ B" - by (force simp: bounded_iff) - have "\\<^sub>F w in at z. cmod (inverse (f w - a)) \ B" - apply (simp add: eventually_at) - apply (rule_tac x=r in exI) - using \0 < r\ by (auto simp: dist_commute intro!: B) - then have "\B. \\<^sub>F z in at z. cmod (inverse (f z - a)) \ B" - by blast - then obtain g where holg: "g holomorphic_on ball z r" and gf: "\w. w \ ball z r - {z} \ g w = inverse (f w - a)" - using \0 < r\ holomorphic_on_extend_bounded [OF holfb_i] by auto - then have gz: "g \z\ g z" - apply (simp add: continuous_at [symmetric]) - using \0 < r\ centre_in_ball field_differentiable_imp_continuous_at holomorphic_on_imp_differentiable_at by blast - have gnz: "\w. w \ ball z r - {z} \ g w \ 0" - using gf fab zrM by fastforce - show ?thesis - proof (cases "g z = 0") - case True - have *: "\g \ 0; inverse g = f - a\ \ g / (1 + a * g) = inverse f" for f g::complex - by (auto simp: field_simps) - have "(inverse \ f) \z\ 0" - proof (rule Lim_transform_within_open [of "\w. g w / (1 + a * g w)" _ _ UNIV "ball z r"]) - show "(\w. g w / (1 + a * g w)) \z\ 0" - using True by (auto simp: intro!: tendsto_eq_intros gz) - show "\x. \x \ ball z r; x \ z\ \ g x / (1 + a * g x) = (inverse \ f) x" - using * gf gnz by simp - qed (use \0 < r\ in auto) - with that show ?thesis by blast - next - case False - show ?thesis - proof (cases "1 + a * g z = 0") - case True - have "(f \ 0) (at z)" - proof (rule Lim_transform_within_open [of "\w. (1 + a * g w) / g w" _ _ _ "ball z r"]) - show "(\w. (1 + a * g w) / g w) \z\ 0" - apply (rule tendsto_eq_intros refl gz \g z \ 0\)+ - by (simp add: True) - show "\x. \x \ ball z r; x \ z\ \ (1 + a * g x) / g x = f x" - using fab fab zrM by (fastforce simp add: gf field_split_simps) - qed (use \0 < r\ in auto) - then show ?thesis - using that by blast - next - case False - have *: "\g \ 0; inverse g = f - a\ \ g / (1 + a * g) = inverse f" for f g::complex - by (auto simp: field_simps) - have "(inverse \ f) \z\ g z / (1 + a * g z)" - proof (rule Lim_transform_within_open [of "\w. g w / (1 + a * g w)" _ _ UNIV "ball z r"]) - show "(\w. g w / (1 + a * g w)) \z\ g z / (1 + a * g z)" - using False by (auto simp: False intro!: tendsto_eq_intros gz) - show "\x. \x \ ball z r; x \ z\ \ g x / (1 + a * g x) = (inverse \ f) x" - using * gf gnz by simp - qed (use \0 < r\ in auto) - with that show ?thesis by blast - qed - qed - qed -qed - - -corollary great_Picard_alt: - assumes M: "open M" "z \ M" and holf: "f holomorphic_on (M - {z})" - and non: "\l. \ (f \ l) (at z)" "\l. \ ((inverse \ f) \ l) (at z)" - obtains a where "- {a} \ f ` (M - {z})" - apply (simp add: subset_iff image_iff) - by (metis great_Picard [OF M _ holf] non) - - -corollary great_Picard_infinite: - assumes M: "open M" "z \ M" and holf: "f holomorphic_on (M - {z})" - and non: "\l. \ (f \ l) (at z)" "\l. \ ((inverse \ f) \ l) (at z)" - obtains a where "\w. w \ a \ infinite {x. x \ M - {z} \ f x = w}" -proof - - have False if "a \ b" and ab: "finite {x. x \ M - {z} \ f x = a}" "finite {x. x \ M - {z} \ f x = b}" for a b - proof - - have finab: "finite {x. x \ M - {z} \ f x \ {a,b}}" - using finite_UnI [OF ab] unfolding mem_Collect_eq insert_iff empty_iff - by (simp add: conj_disj_distribL) - obtain r where "0 < r" and zrM: "ball z r \ M" and r: "\x. \x \ M - {z}; f x \ {a,b}\ \ x \ ball z r" - proof - - obtain e where "e > 0" and e: "ball z e \ M" - using assms openE by blast - show ?thesis - proof (cases "{x \ M - {z}. f x \ {a, b}} = {}") - case True - then show ?thesis - apply (rule_tac r=e in that) - using e \e > 0\ by auto - next - case False - let ?r = "min e (Min (dist z ` {x \ M - {z}. f x \ {a,b}}))" - show ?thesis - proof - show "0 < ?r" - using min_less_iff_conj Min_gr_iff finab False \0 < e\ by auto - have "ball z ?r \ ball z e" - by (simp add: subset_ball) - with e show "ball z ?r \ M" by blast - show "\x. \x \ M - {z}; f x \ {a, b}\ \ x \ ball z ?r" - using min_less_iff_conj Min_gr_iff finab False \0 < e\ by auto - qed - qed - qed - have holfb: "f holomorphic_on (ball z r - {z})" - apply (rule holomorphic_on_subset [OF holf]) - using zrM by auto - show ?thesis - apply (rule great_Picard [OF open_ball _ \a \ b\ holfb]) - using non \0 < r\ r zrM by auto - qed - with that show thesis - by meson -qed - -theorem Casorati_Weierstrass: - assumes "open M" "z \ M" "f holomorphic_on (M - {z})" - and "\l. \ (f \ l) (at z)" "\l. \ ((inverse \ f) \ l) (at z)" - shows "closure(f ` (M - {z})) = UNIV" -proof - - obtain a where a: "- {a} \ f ` (M - {z})" - using great_Picard_alt [OF assms] . - have "UNIV = closure(- {a})" - by (simp add: closure_interior) - also have "... \ closure(f ` (M - {z}))" - by (simp add: a closure_mono) - finally show ?thesis - by blast -qed - -end diff --git a/src/HOL/Analysis/Line_Segment.thy b/src/HOL/Analysis/Line_Segment.thy --- a/src/HOL/Analysis/Line_Segment.thy +++ b/src/HOL/Analysis/Line_Segment.thy @@ -1,1012 +1,1048 @@ (* Title: HOL/Analysis/Line_Segment.thy Author: L C Paulson, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Bogdan Grechuk, University of Edinburgh Author: Armin Heller, TU Muenchen Author: Johannes Hoelzl, TU Muenchen *) section \Line Segment\ theory Line_Segment imports Convex Topology_Euclidean_Space begin subsection\<^marker>\tag unimportant\ \Topological Properties of Convex Sets, Metric Spaces and Functions\ lemma convex_supp_sum: assumes "convex S" and 1: "supp_sum u I = 1" and "\i. i \ I \ 0 \ u i \ (u i = 0 \ f i \ S)" shows "supp_sum (\i. u i *\<^sub>R f i) I \ S" proof - have fin: "finite {i \ I. u i \ 0}" using 1 sum.infinite by (force simp: supp_sum_def support_on_def) then have "supp_sum (\i. u i *\<^sub>R f i) I = sum (\i. u i *\<^sub>R f i) {i \ I. u i \ 0}" by (force intro: sum.mono_neutral_left simp: supp_sum_def support_on_def) also have "... \ S" using 1 assms by (force simp: supp_sum_def support_on_def intro: convex_sum [OF fin \convex S\]) finally show ?thesis . qed lemma closure_bounded_linear_image_subset: assumes f: "bounded_linear f" shows "f ` closure S \ closure (f ` S)" using linear_continuous_on [OF f] closed_closure closure_subset by (rule image_closure_subset) lemma closure_linear_image_subset: fixes f :: "'m::euclidean_space \ 'n::real_normed_vector" assumes "linear f" shows "f ` (closure S) \ closure (f ` S)" using assms unfolding linear_conv_bounded_linear by (rule closure_bounded_linear_image_subset) lemma closed_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes S: "closed S" and f: "linear f" "inj f" shows "closed (f ` S)" proof - obtain g where g: "linear g" "g \ f = id" using linear_injective_left_inverse [OF f] by blast then have confg: "continuous_on (range f) g" using linear_continuous_on linear_conv_bounded_linear by blast have [simp]: "g ` f ` S = S" using g by (simp add: image_comp) have cgf: "closed (g ` f ` S)" by (simp add: \g \ f = id\ S image_comp) have [simp]: "(range f \ g -` S) = f ` S" using g unfolding o_def id_def image_def by auto metis+ show ?thesis proof (rule closedin_closed_trans [of "range f"]) show "closedin (top_of_set (range f)) (f ` S)" using continuous_closedin_preimage [OF confg cgf] by simp show "closed (range f)" apply (rule closed_injective_image_subspace) using f apply (auto simp: linear_linear linear_injective_0) done qed qed lemma closed_injective_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes f: "linear f" "inj f" shows "(closed(image f s) \ closed s)" by (metis closed_injective_linear_image closure_eq closure_linear_image_subset closure_subset_eq f(1) f(2) inj_image_subset_iff) lemma closure_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\linear f; inj f\ \ f ` (closure S) = closure (f ` S)" apply (rule subset_antisym) apply (simp add: closure_linear_image_subset) by (simp add: closure_minimal closed_injective_linear_image closure_subset image_mono) lemma closure_bounded_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" shows "\linear f; bounded S\ \ f ` (closure S) = closure (f ` S)" apply (rule subset_antisym, simp add: closure_linear_image_subset) apply (rule closure_minimal, simp add: closure_subset image_mono) by (meson bounded_closure closed_closure compact_continuous_image compact_eq_bounded_closed linear_continuous_on linear_conv_bounded_linear) lemma closure_scaleR: fixes S :: "'a::real_normed_vector set" shows "((*\<^sub>R) c) ` (closure S) = closure (((*\<^sub>R) c) ` S)" proof show "((*\<^sub>R) c) ` (closure S) \ closure (((*\<^sub>R) c) ` S)" using bounded_linear_scaleR_right by (rule closure_bounded_linear_image_subset) show "closure (((*\<^sub>R) c) ` S) \ ((*\<^sub>R) c) ` (closure S)" by (intro closure_minimal image_mono closure_subset closed_scaling closed_closure) qed lemma sphere_eq_empty [simp]: fixes a :: "'a::{real_normed_vector, perfect_space}" shows "sphere a r = {} \ r < 0" by (auto simp: sphere_def dist_norm) (metis dist_norm le_less_linear vector_choose_dist) lemma cone_closure: fixes S :: "'a::real_normed_vector set" assumes "cone S" shows "cone (closure S)" proof (cases "S = {}") case True then show ?thesis by auto next case False then have "0 \ S \ (\c. c > 0 \ (*\<^sub>R) c ` S = S)" using cone_iff[of S] assms by auto then have "0 \ closure S \ (\c. c > 0 \ (*\<^sub>R) c ` closure S = closure S)" using closure_subset by (auto simp: closure_scaleR) then show ?thesis using False cone_iff[of "closure S"] by auto qed corollary component_complement_connected: fixes S :: "'a::real_normed_vector set" assumes "connected S" "C \ components (-S)" shows "connected(-C)" using component_diff_connected [of S UNIV] assms by (auto simp: Compl_eq_Diff_UNIV) proposition clopen: fixes S :: "'a :: real_normed_vector set" shows "closed S \ open S \ S = {} \ S = UNIV" by (force intro!: connected_UNIV [unfolded connected_clopen, rule_format]) corollary compact_open: fixes S :: "'a :: euclidean_space set" shows "compact S \ open S \ S = {}" by (auto simp: compact_eq_bounded_closed clopen) corollary finite_imp_not_open: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "\finite S; open S\ \ S={}" using clopen [of S] finite_imp_closed not_bounded_UNIV by blast corollary empty_interior_finite: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "finite S \ interior S = {}" by (metis interior_subset finite_subset open_interior [of S] finite_imp_not_open) text \Balls, being convex, are connected.\ lemma convex_local_global_minimum: fixes s :: "'a::real_normed_vector set" assumes "e > 0" and "convex_on s f" and "ball x e \ s" and "\y\ball x e. f x \ f y" shows "\y\s. f x \ f y" proof (rule ccontr) have "x \ s" using assms(1,3) by auto assume "\ ?thesis" then obtain y where "y\s" and y: "f x > f y" by auto then have xy: "0 < dist x y" by auto then obtain u where "0 < u" "u \ 1" and u: "u < e / dist x y" using field_lbound_gt_zero[of 1 "e / dist x y"] xy \e>0\ by auto then have "f ((1-u) *\<^sub>R x + u *\<^sub>R y) \ (1-u) * f x + u * f y" using \x\s\ \y\s\ using assms(2)[unfolded convex_on_def, THEN bspec[where x=x], THEN bspec[where x=y], THEN spec[where x="1-u"]] by auto moreover have *: "x - ((1 - u) *\<^sub>R x + u *\<^sub>R y) = u *\<^sub>R (x - y)" by (simp add: algebra_simps) have "(1 - u) *\<^sub>R x + u *\<^sub>R y \ ball x e" unfolding mem_ball dist_norm unfolding * and norm_scaleR and abs_of_pos[OF \0] unfolding dist_norm[symmetric] using u unfolding pos_less_divide_eq[OF xy] by auto then have "f x \ f ((1 - u) *\<^sub>R x + u *\<^sub>R y)" using assms(4) by auto ultimately show False using mult_strict_left_mono[OF y \u>0\] unfolding left_diff_distrib by auto qed lemma convex_ball [iff]: fixes x :: "'a::real_normed_vector" shows "convex (ball x e)" proof (auto simp: convex_def) fix y z assume yz: "dist x y < e" "dist x z < e" fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ u * dist x y + v * dist x z" using uv yz using convex_on_dist [of "ball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto then show "dist x (u *\<^sub>R y + v *\<^sub>R z) < e" using convex_bound_lt[OF yz uv] by auto qed lemma convex_cball [iff]: fixes x :: "'a::real_normed_vector" shows "convex (cball x e)" proof - { fix y z assume yz: "dist x y \ e" "dist x z \ e" fix u v :: real assume uv: "0 \ u" "0 \ v" "u + v = 1" have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ u * dist x y + v * dist x z" using uv yz using convex_on_dist [of "cball x e" x, unfolded convex_on_def, THEN bspec[where x=y], THEN bspec[where x=z]] by auto then have "dist x (u *\<^sub>R y + v *\<^sub>R z) \ e" using convex_bound_le[OF yz uv] by auto } then show ?thesis by (auto simp: convex_def Ball_def) qed lemma connected_ball [iff]: fixes x :: "'a::real_normed_vector" shows "connected (ball x e)" using convex_connected convex_ball by auto lemma connected_cball [iff]: fixes x :: "'a::real_normed_vector" shows "connected (cball x e)" using convex_connected convex_cball by auto lemma bounded_convex_hull: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "bounded (convex hull s)" proof - from assms obtain B where B: "\x\s. norm x \ B" unfolding bounded_iff by auto show ?thesis by (simp add: bounded_subset[OF bounded_cball, of _ 0 B] B subsetI subset_hull) qed lemma finite_imp_bounded_convex_hull: fixes s :: "'a::real_normed_vector set" shows "finite s \ bounded (convex hull s)" using bounded_convex_hull finite_imp_bounded by auto subsection \Midpoint\ definition\<^marker>\tag important\ midpoint :: "'a::real_vector \ 'a \ 'a" where "midpoint a b = (inverse (2::real)) *\<^sub>R (a + b)" lemma midpoint_idem [simp]: "midpoint x x = x" unfolding midpoint_def by simp lemma midpoint_sym: "midpoint a b = midpoint b a" unfolding midpoint_def by (auto simp add: scaleR_right_distrib) lemma midpoint_eq_iff: "midpoint a b = c \ a + b = c + c" proof - have "midpoint a b = c \ scaleR 2 (midpoint a b) = scaleR 2 c" by simp then show ?thesis unfolding midpoint_def scaleR_2 [symmetric] by simp qed lemma fixes a::real assumes "a \ b" shows ge_midpoint_1: "a \ midpoint a b" and le_midpoint_1: "midpoint a b \ b" by (simp_all add: midpoint_def assms) lemma dist_midpoint: fixes a b :: "'a::real_normed_vector" shows "dist a (midpoint a b) = (dist a b) / 2" (is ?t1) "dist b (midpoint a b) = (dist a b) / 2" (is ?t2) "dist (midpoint a b) a = (dist a b) / 2" (is ?t3) "dist (midpoint a b) b = (dist a b) / 2" (is ?t4) proof - have *: "\x y::'a. 2 *\<^sub>R x = - y \ norm x = (norm y) / 2" unfolding equation_minus_iff by auto have **: "\x y::'a. 2 *\<^sub>R x = y \ norm x = (norm y) / 2" by auto note scaleR_right_distrib [simp] show ?t1 unfolding midpoint_def dist_norm apply (rule **) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done show ?t2 unfolding midpoint_def dist_norm apply (rule *) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done show ?t3 unfolding midpoint_def dist_norm apply (rule *) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done show ?t4 unfolding midpoint_def dist_norm apply (rule **) apply (simp add: scaleR_right_diff_distrib) apply (simp add: scaleR_2) done qed lemma midpoint_eq_endpoint [simp]: "midpoint a b = a \ a = b" "midpoint a b = b \ a = b" unfolding midpoint_eq_iff by auto lemma midpoint_plus_self [simp]: "midpoint a b + midpoint a b = a + b" using midpoint_eq_iff by metis lemma midpoint_linear_image: "linear f \ midpoint(f a)(f b) = f(midpoint a b)" by (simp add: linear_iff midpoint_def) subsection \Open and closed segments\ definition\<^marker>\tag important\ closed_segment :: "'a::real_vector \ 'a \ 'a set" where "closed_segment a b = {(1 - u) *\<^sub>R a + u *\<^sub>R b | u::real. 0 \ u \ u \ 1}" definition\<^marker>\tag important\ open_segment :: "'a::real_vector \ 'a \ 'a set" where "open_segment a b \ closed_segment a b - {a,b}" lemmas segment = open_segment_def closed_segment_def lemma in_segment: "x \ closed_segment a b \ (\u. 0 \ u \ u \ 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" "x \ open_segment a b \ a \ b \ (\u. 0 < u \ u < 1 \ x = (1 - u) *\<^sub>R a + u *\<^sub>R b)" using less_eq_real_def by (auto simp: segment algebra_simps) lemma closed_segment_linear_image: "closed_segment (f a) (f b) = f ` (closed_segment a b)" if "linear f" proof - interpret linear f by fact show ?thesis by (force simp add: in_segment add scale) qed lemma open_segment_linear_image: "\linear f; inj f\ \ open_segment (f a) (f b) = f ` (open_segment a b)" by (force simp: open_segment_def closed_segment_linear_image inj_on_def) lemma closed_segment_translation: "closed_segment (c + a) (c + b) = image (\x. c + x) (closed_segment a b)" apply safe apply (rule_tac x="x-c" in image_eqI) apply (auto simp: in_segment algebra_simps) done lemma open_segment_translation: "open_segment (c + a) (c + b) = image (\x. c + x) (open_segment a b)" by (simp add: open_segment_def closed_segment_translation translation_diff) lemma closed_segment_of_real: "closed_segment (of_real x) (of_real y) = of_real ` closed_segment x y" apply (auto simp: image_iff in_segment scaleR_conv_of_real) apply (rule_tac x="(1-u)*x + u*y" in bexI) apply (auto simp: in_segment) done lemma open_segment_of_real: "open_segment (of_real x) (of_real y) = of_real ` open_segment x y" apply (auto simp: image_iff in_segment scaleR_conv_of_real) apply (rule_tac x="(1-u)*x + u*y" in bexI) apply (auto simp: in_segment) done lemma closed_segment_Reals: "\x \ Reals; y \ Reals\ \ closed_segment x y = of_real ` closed_segment (Re x) (Re y)" by (metis closed_segment_of_real of_real_Re) lemma open_segment_Reals: "\x \ Reals; y \ Reals\ \ open_segment x y = of_real ` open_segment (Re x) (Re y)" by (metis open_segment_of_real of_real_Re) lemma open_segment_PairD: "(x, x') \ open_segment (a, a') (b, b') \ (x \ open_segment a b \ a = b) \ (x' \ open_segment a' b' \ a' = b')" by (auto simp: in_segment) lemma closed_segment_PairD: "(x, x') \ closed_segment (a, a') (b, b') \ x \ closed_segment a b \ x' \ closed_segment a' b'" by (auto simp: closed_segment_def) lemma closed_segment_translation_eq [simp]: "d + x \ closed_segment (d + a) (d + b) \ x \ closed_segment a b" proof - have *: "\d x a b. x \ closed_segment a b \ d + x \ closed_segment (d + a) (d + b)" apply (simp add: closed_segment_def) apply (erule ex_forward) apply (simp add: algebra_simps) done show ?thesis using * [where d = "-d"] * by (fastforce simp add:) qed lemma open_segment_translation_eq [simp]: "d + x \ open_segment (d + a) (d + b) \ x \ open_segment a b" by (simp add: open_segment_def) lemma of_real_closed_segment [simp]: "of_real x \ closed_segment (of_real a) (of_real b) \ x \ closed_segment a b" apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward) using of_real_eq_iff by fastforce lemma of_real_open_segment [simp]: "of_real x \ open_segment (of_real a) (of_real b) \ x \ open_segment a b" apply (auto simp: in_segment scaleR_conv_of_real elim!: ex_forward del: exE) using of_real_eq_iff by fastforce lemma convex_contains_segment: "convex S \ (\a\S. \b\S. closed_segment a b \ S)" unfolding convex_alt closed_segment_def by auto lemma closed_segment_in_Reals: "\x \ closed_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" by (meson subsetD convex_Reals convex_contains_segment) lemma open_segment_in_Reals: "\x \ open_segment a b; a \ Reals; b \ Reals\ \ x \ Reals" by (metis Diff_iff closed_segment_in_Reals open_segment_def) lemma closed_segment_subset: "\x \ S; y \ S; convex S\ \ closed_segment x y \ S" by (simp add: convex_contains_segment) lemma closed_segment_subset_convex_hull: "\x \ convex hull S; y \ convex hull S\ \ closed_segment x y \ convex hull S" using convex_contains_segment by blast lemma segment_convex_hull: "closed_segment a b = convex hull {a,b}" proof - have *: "\x. {x} \ {}" by auto show ?thesis unfolding segment convex_hull_insert[OF *] convex_hull_singleton by (safe; rule_tac x="1 - u" in exI; force) qed lemma open_closed_segment: "u \ open_segment w z \ u \ closed_segment w z" by (auto simp add: closed_segment_def open_segment_def) lemma segment_open_subset_closed: "open_segment a b \ closed_segment a b" by (auto simp: closed_segment_def open_segment_def) lemma bounded_closed_segment: fixes a :: "'a::real_normed_vector" shows "bounded (closed_segment a b)" by (rule boundedI[where B="max (norm a) (norm b)"]) (auto simp: closed_segment_def max_def convex_bound_le intro!: norm_triangle_le) lemma bounded_open_segment: fixes a :: "'a::real_normed_vector" shows "bounded (open_segment a b)" by (rule bounded_subset [OF bounded_closed_segment segment_open_subset_closed]) lemmas bounded_segment = bounded_closed_segment open_closed_segment lemma ends_in_segment [iff]: "a \ closed_segment a b" "b \ closed_segment a b" unfolding segment_convex_hull by (auto intro!: hull_subset[unfolded subset_eq, rule_format]) lemma eventually_closed_segment: fixes x0::"'a::real_normed_vector" assumes "open X0" "x0 \ X0" shows "\\<^sub>F x in at x0 within U. closed_segment x0 x \ X0" proof - from openE[OF assms] obtain e where e: "0 < e" "ball x0 e \ X0" . then have "\\<^sub>F x in at x0 within U. x \ ball x0 e" by (auto simp: dist_commute eventually_at) then show ?thesis proof eventually_elim case (elim x) have "x0 \ ball x0 e" using \e > 0\ by simp from convex_ball[unfolded convex_contains_segment, rule_format, OF this elim] have "closed_segment x0 x \ ball x0 e" . also note \\ \ X0\ finally show ?case . qed qed lemma closed_segment_commute: "closed_segment a b = closed_segment b a" proof - have "{a, b} = {b, a}" by auto thus ?thesis by (simp add: segment_convex_hull) qed lemma segment_bound1: assumes "x \ closed_segment a b" shows "norm (x - a) \ norm (b - a)" proof - obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" using assms by (auto simp add: closed_segment_def) then show "norm (x - a) \ norm (b - a)" apply clarify apply (auto simp: algebra_simps) apply (simp add: scaleR_diff_right [symmetric] mult_left_le_one_le) done qed lemma segment_bound: assumes "x \ closed_segment a b" shows "norm (x - a) \ norm (b - a)" "norm (x - b) \ norm (b - a)" by (metis assms closed_segment_commute dist_commute dist_norm segment_bound1)+ lemma open_segment_commute: "open_segment a b = open_segment b a" proof - have "{a, b} = {b, a}" by auto thus ?thesis by (simp add: closed_segment_commute open_segment_def) qed lemma closed_segment_idem [simp]: "closed_segment a a = {a}" unfolding segment by (auto simp add: algebra_simps) lemma open_segment_idem [simp]: "open_segment a a = {}" by (simp add: open_segment_def) lemma closed_segment_eq_open: "closed_segment a b = open_segment a b \ {a,b}" using open_segment_def by auto lemma convex_contains_open_segment: "convex s \ (\a\s. \b\s. open_segment a b \ s)" by (simp add: convex_contains_segment closed_segment_eq_open) lemma closed_segment_eq_real_ivl1: fixes a b::real assumes "a \ b" shows "closed_segment a b = {a .. b}" proof safe fix x assume "x \ closed_segment a b" then obtain u where u: "0 \ u" "u \ 1" and x_def: "x = (1 - u) * a + u * b" by (auto simp: closed_segment_def) have "u * a \ u * b" "(1 - u) * a \ (1 - u) * b" by (auto intro!: mult_left_mono u assms) then show "x \ {a .. b}" unfolding x_def by (auto simp: algebra_simps) next show "\x. x \ {a..b} \ x \ closed_segment a b" by (force simp: closed_segment_def divide_simps algebra_simps intro: exI[where x="(x - a) / (b - a)" for x]) qed lemma closed_segment_eq_real_ivl: fixes a b::real shows "closed_segment a b = (if a \ b then {a .. b} else {b .. a})" using closed_segment_eq_real_ivl1[of a b] closed_segment_eq_real_ivl1[of b a] by (auto simp: closed_segment_commute) lemma open_segment_eq_real_ivl: fixes a b::real shows "open_segment a b = (if a \ b then {a<..x. (v - u) * x + u) ` {0..1}" by (simp add: add.commute [of u] image_affinity_atLeastAtMost [where c=u] closed_segment_eq_real_ivl) +lemma closed_segment_same_Re: + assumes "Re a = Re b" + shows "closed_segment a b = {z. Re z = Re a \ Im z \ closed_segment (Im a) (Im b)}" +proof safe + fix z assume "z \ closed_segment a b" + then obtain u where u: "u \ {0..1}" "z = a + of_real u * (b - a)" + by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) + from assms show "Re z = Re a" by (auto simp: u) + from u(1) show "Im z \ closed_segment (Im a) (Im b)" + by (force simp: u closed_segment_def algebra_simps) +next + fix z assume [simp]: "Re z = Re a" and "Im z \ closed_segment (Im a) (Im b)" + then obtain u where u: "u \ {0..1}" "Im z = Im a + of_real u * (Im b - Im a)" + by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) + from u(1) show "z \ closed_segment a b" using assms + by (force simp: u closed_segment_def algebra_simps scaleR_conv_of_real complex_eq_iff) +qed + +lemma closed_segment_same_Im: + assumes "Im a = Im b" + shows "closed_segment a b = {z. Im z = Im a \ Re z \ closed_segment (Re a) (Re b)}" +proof safe + fix z assume "z \ closed_segment a b" + then obtain u where u: "u \ {0..1}" "z = a + of_real u * (b - a)" + by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) + from assms show "Im z = Im a" by (auto simp: u) + from u(1) show "Re z \ closed_segment (Re a) (Re b)" + by (force simp: u closed_segment_def algebra_simps) +next + fix z assume [simp]: "Im z = Im a" and "Re z \ closed_segment (Re a) (Re b)" + then obtain u where u: "u \ {0..1}" "Re z = Re a + of_real u * (Re b - Re a)" + by (auto simp: closed_segment_def scaleR_conv_of_real algebra_simps) + from u(1) show "z \ closed_segment a b" using assms + by (force simp: u closed_segment_def algebra_simps scaleR_conv_of_real complex_eq_iff) +qed + lemma dist_in_closed_segment: fixes a :: "'a :: euclidean_space" assumes "x \ closed_segment a b" shows "dist x a \ dist a b \ dist x b \ dist a b" proof (intro conjI) obtain u where u: "0 \ u" "u \ 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" using assms by (force simp: in_segment algebra_simps) have "dist x a = u * dist a b" apply (simp add: dist_norm algebra_simps x) by (metis \0 \ u\ abs_of_nonneg norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib) also have "... \ dist a b" by (simp add: mult_left_le_one_le u) finally show "dist x a \ dist a b" . have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" by (simp add: dist_norm algebra_simps x) also have "... = (1-u) * dist a b" proof - have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" using \u \ 1\ by force then show ?thesis by (simp add: dist_norm real_vector.scale_right_diff_distrib) qed also have "... \ dist a b" by (simp add: mult_left_le_one_le u) finally show "dist x b \ dist a b" . qed lemma dist_in_open_segment: fixes a :: "'a :: euclidean_space" assumes "x \ open_segment a b" shows "dist x a < dist a b \ dist x b < dist a b" proof (intro conjI) obtain u where u: "0 < u" "u < 1" and x: "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" using assms by (force simp: in_segment algebra_simps) have "dist x a = u * dist a b" apply (simp add: dist_norm algebra_simps x) by (metis abs_of_nonneg less_eq_real_def norm_minus_commute norm_scaleR real_vector.scale_right_diff_distrib \0 < u\) also have *: "... < dist a b" by (metis (no_types) assms dist_eq_0_iff dist_not_less_zero in_segment(2) linorder_neqE_linordered_idom mult.left_neutral real_mult_less_iff1 \u < 1\) finally show "dist x a < dist a b" . have ab_ne0: "dist a b \ 0" using * by fastforce have "dist x b = norm ((1-u) *\<^sub>R a - (1-u) *\<^sub>R b)" by (simp add: dist_norm algebra_simps x) also have "... = (1-u) * dist a b" proof - have "norm ((1 - 1 * u) *\<^sub>R (a - b)) = (1 - 1 * u) * norm (a - b)" using \u < 1\ by force then show ?thesis by (simp add: dist_norm real_vector.scale_right_diff_distrib) qed also have "... < dist a b" using ab_ne0 \0 < u\ by simp finally show "dist x b < dist a b" . qed lemma dist_decreases_open_segment_0: fixes x :: "'a :: euclidean_space" assumes "x \ open_segment 0 b" shows "dist c x < dist c 0 \ dist c x < dist c b" proof (rule ccontr, clarsimp simp: not_less) obtain u where u: "0 \ b" "0 < u" "u < 1" and x: "x = u *\<^sub>R b" using assms by (auto simp: in_segment) have xb: "x \ b < b \ b" using u x by auto assume "norm c \ dist c x" then have "c \ c \ (c - x) \ (c - x)" by (simp add: dist_norm norm_le) moreover have "0 < x \ b" using u x by auto ultimately have less: "c \ b < x \ b" by (simp add: x algebra_simps inner_commute u) assume "dist c b \ dist c x" then have "(c - b) \ (c - b) \ (c - x) \ (c - x)" by (simp add: dist_norm norm_le) then have "(b \ b) * (1 - u*u) \ 2 * (b \ c) * (1-u)" by (simp add: x algebra_simps inner_commute) then have "(1+u) * (b \ b) * (1-u) \ 2 * (b \ c) * (1-u)" by (simp add: algebra_simps) then have "(1+u) * (b \ b) \ 2 * (b \ c)" using \u < 1\ by auto with xb have "c \ b \ x \ b" by (auto simp: x algebra_simps inner_commute) with less show False by auto qed proposition dist_decreases_open_segment: fixes a :: "'a :: euclidean_space" assumes "x \ open_segment a b" shows "dist c x < dist c a \ dist c x < dist c b" proof - have *: "x - a \ open_segment 0 (b - a)" using assms by (metis diff_self open_segment_translation_eq uminus_add_conv_diff) show ?thesis using dist_decreases_open_segment_0 [OF *, of "c-a"] assms by (simp add: dist_norm) qed corollary open_segment_furthest_le: fixes a b x y :: "'a::euclidean_space" assumes "x \ open_segment a b" shows "norm (y - x) < norm (y - a) \ norm (y - x) < norm (y - b)" by (metis assms dist_decreases_open_segment dist_norm) corollary dist_decreases_closed_segment: fixes a :: "'a :: euclidean_space" assumes "x \ closed_segment a b" shows "dist c x \ dist c a \ dist c x \ dist c b" apply (cases "x \ open_segment a b") using dist_decreases_open_segment less_eq_real_def apply blast by (metis DiffI assms empty_iff insertE open_segment_def order_refl) corollary segment_furthest_le: fixes a b x y :: "'a::euclidean_space" assumes "x \ closed_segment a b" shows "norm (y - x) \ norm (y - a) \ norm (y - x) \ norm (y - b)" by (metis assms dist_decreases_closed_segment dist_norm) lemma convex_intermediate_ball: fixes a :: "'a :: euclidean_space" shows "\ball a r \ T; T \ cball a r\ \ convex T" apply (simp add: convex_contains_open_segment, clarify) by (metis (no_types, hide_lams) less_le_trans mem_ball mem_cball subsetCE dist_decreases_open_segment) lemma csegment_midpoint_subset: "closed_segment (midpoint a b) b \ closed_segment a b" apply (clarsimp simp: midpoint_def in_segment) apply (rule_tac x="(1 + u) / 2" in exI) apply (auto simp: algebra_simps add_divide_distrib diff_divide_distrib) by (metis field_sum_of_halves scaleR_left.add) lemma notin_segment_midpoint: fixes a :: "'a :: euclidean_space" shows "a \ b \ a \ closed_segment (midpoint a b) b" by (auto simp: dist_midpoint dest!: dist_in_closed_segment) subsubsection\More lemmas, especially for working with the underlying formula\ lemma segment_eq_compose: fixes a :: "'a :: real_vector" shows "(\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) = (\x. a + x) o (\u. u *\<^sub>R (b - a))" by (simp add: o_def algebra_simps) lemma segment_degen_1: fixes a :: "'a :: real_vector" shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = b \ a=b \ u=1" proof - { assume "(1 - u) *\<^sub>R a + u *\<^sub>R b = b" then have "(1 - u) *\<^sub>R a = (1 - u) *\<^sub>R b" by (simp add: algebra_simps) then have "a=b \ u=1" by simp } then show ?thesis by (auto simp: algebra_simps) qed lemma segment_degen_0: fixes a :: "'a :: real_vector" shows "(1 - u) *\<^sub>R a + u *\<^sub>R b = a \ a=b \ u=0" using segment_degen_1 [of "1-u" b a] by (auto simp: algebra_simps) lemma add_scaleR_degen: fixes a b ::"'a::real_vector" assumes "(u *\<^sub>R b + v *\<^sub>R a) = (u *\<^sub>R a + v *\<^sub>R b)" "u \ v" shows "a=b" by (metis (no_types, hide_lams) add.commute add_diff_eq diff_add_cancel real_vector.scale_cancel_left real_vector.scale_left_diff_distrib assms) lemma closed_segment_image_interval: "closed_segment a b = (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0..1}" by (auto simp: set_eq_iff image_iff closed_segment_def) lemma open_segment_image_interval: "open_segment a b = (if a=b then {} else (\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1})" by (auto simp: open_segment_def closed_segment_def segment_degen_0 segment_degen_1) lemmas segment_image_interval = closed_segment_image_interval open_segment_image_interval lemma open_segment_bound1: assumes "x \ open_segment a b" shows "norm (x - a) < norm (b - a)" proof - obtain u where "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 < u" "u < 1" "a \ b" using assms by (auto simp add: open_segment_image_interval split: if_split_asm) then show "norm (x - a) < norm (b - a)" apply clarify apply (auto simp: algebra_simps) apply (simp add: scaleR_diff_right [symmetric]) done qed lemma compact_segment [simp]: fixes a :: "'a::real_normed_vector" shows "compact (closed_segment a b)" by (auto simp: segment_image_interval intro!: compact_continuous_image continuous_intros) lemma closed_segment [simp]: fixes a :: "'a::real_normed_vector" shows "closed (closed_segment a b)" by (simp add: compact_imp_closed) lemma closure_closed_segment [simp]: fixes a :: "'a::real_normed_vector" shows "closure(closed_segment a b) = closed_segment a b" by simp lemma open_segment_bound: assumes "x \ open_segment a b" shows "norm (x - a) < norm (b - a)" "norm (x - b) < norm (b - a)" apply (simp add: assms open_segment_bound1) by (metis assms norm_minus_commute open_segment_bound1 open_segment_commute) lemma closure_open_segment [simp]: "closure (open_segment a b) = (if a = b then {} else closed_segment a b)" for a :: "'a::euclidean_space" proof (cases "a = b") case True then show ?thesis by simp next case False have "closure ((\u. u *\<^sub>R (b - a)) ` {0<..<1}) = (\u. u *\<^sub>R (b - a)) ` closure {0<..<1}" apply (rule closure_injective_linear_image [symmetric]) apply (use False in \auto intro!: injI\) done then have "closure ((\u. (1 - u) *\<^sub>R a + u *\<^sub>R b) ` {0<..<1}) = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b) ` closure {0<..<1}" using closure_translation [of a "((\x. x *\<^sub>R b - x *\<^sub>R a) ` {0<..<1})"] by (simp add: segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right image_image) then show ?thesis by (simp add: segment_image_interval closure_greaterThanLessThan [symmetric] del: closure_greaterThanLessThan) qed lemma closed_open_segment_iff [simp]: fixes a :: "'a::euclidean_space" shows "closed(open_segment a b) \ a = b" by (metis open_segment_def DiffE closure_eq closure_open_segment ends_in_segment(1) insert_iff segment_image_interval(2)) lemma compact_open_segment_iff [simp]: fixes a :: "'a::euclidean_space" shows "compact(open_segment a b) \ a = b" by (simp add: bounded_open_segment compact_eq_bounded_closed) lemma convex_closed_segment [iff]: "convex (closed_segment a b)" unfolding segment_convex_hull by(rule convex_convex_hull) lemma convex_open_segment [iff]: "convex (open_segment a b)" proof - have "convex ((\u. u *\<^sub>R (b - a)) ` {0<..<1})" by (rule convex_linear_image) auto then have "convex ((+) a ` (\u. u *\<^sub>R (b - a)) ` {0<..<1})" by (rule convex_translation) then show ?thesis by (simp add: image_image open_segment_image_interval segment_eq_compose field_simps scaleR_diff_left scaleR_diff_right) qed lemmas convex_segment = convex_closed_segment convex_open_segment lemma connected_segment [iff]: fixes x :: "'a :: real_normed_vector" shows "connected (closed_segment x y)" by (simp add: convex_connected) lemma is_interval_closed_segment_1[intro, simp]: "is_interval (closed_segment a b)" for a b::real unfolding closed_segment_eq_real_ivl by (auto simp: is_interval_def) lemma IVT'_closed_segment_real: fixes f :: "real \ real" assumes "y \ closed_segment (f a) (f b)" assumes "continuous_on (closed_segment a b) f" shows "\x \ closed_segment a b. f x = y" using IVT'[of f a y b] IVT'[of "-f" a "-y" b] IVT'[of f b y a] IVT'[of "-f" b "-y" a] assms by (cases "a \ b"; cases "f b \ f a") (auto simp: closed_segment_eq_real_ivl continuous_on_minus) subsection \Betweenness\ definition\<^marker>\tag important\ "between = (\(a,b) x. x \ closed_segment a b)" lemma betweenI: assumes "0 \ u" "u \ 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" shows "between (a, b) x" using assms unfolding between_def closed_segment_def by auto lemma betweenE: assumes "between (a, b) x" obtains u where "0 \ u" "u \ 1" "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" using assms unfolding between_def closed_segment_def by auto lemma between_implies_scaled_diff: assumes "between (S, T) X" "between (S, T) Y" "S \ Y" obtains c where "(X - Y) = c *\<^sub>R (S - Y)" proof - from \between (S, T) X\ obtain u\<^sub>X where X: "X = u\<^sub>X *\<^sub>R S + (1 - u\<^sub>X) *\<^sub>R T" by (metis add.commute betweenE eq_diff_eq) from \between (S, T) Y\ obtain u\<^sub>Y where Y: "Y = u\<^sub>Y *\<^sub>R S + (1 - u\<^sub>Y) *\<^sub>R T" by (metis add.commute betweenE eq_diff_eq) have "X - Y = (u\<^sub>X - u\<^sub>Y) *\<^sub>R (S - T)" proof - from X Y have "X - Y = u\<^sub>X *\<^sub>R S - u\<^sub>Y *\<^sub>R S + ((1 - u\<^sub>X) *\<^sub>R T - (1 - u\<^sub>Y) *\<^sub>R T)" by simp also have "\ = (u\<^sub>X - u\<^sub>Y) *\<^sub>R S - (u\<^sub>X - u\<^sub>Y) *\<^sub>R T" by (simp add: scaleR_left.diff) finally show ?thesis by (simp add: real_vector.scale_right_diff_distrib) qed moreover from Y have "S - Y = (1 - u\<^sub>Y) *\<^sub>R (S - T)" by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) moreover note \S \ Y\ ultimately have "(X - Y) = ((u\<^sub>X - u\<^sub>Y) / (1 - u\<^sub>Y)) *\<^sub>R (S - Y)" by auto from this that show thesis by blast qed lemma between_mem_segment: "between (a,b) x \ x \ closed_segment a b" unfolding between_def by auto lemma between: "between (a, b) (x::'a::euclidean_space) \ dist a b = (dist a x) + (dist x b)" proof (cases "a = b") case True then show ?thesis by (auto simp add: between_def dist_commute) next case False then have Fal: "norm (a - b) \ 0" and Fal2: "norm (a - b) > 0" by auto have *: "\u. a - ((1 - u) *\<^sub>R a + u *\<^sub>R b) = u *\<^sub>R (a - b)" by (auto simp add: algebra_simps) have "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" if "x = (1 - u) *\<^sub>R a + u *\<^sub>R b" "0 \ u" "u \ 1" for u proof - have *: "a - x = u *\<^sub>R (a - b)" "x - b = (1 - u) *\<^sub>R (a - b)" unfolding that(1) by (auto simp add:algebra_simps) show "norm (a - x) *\<^sub>R (x - b) = norm (x - b) *\<^sub>R (a - x)" unfolding norm_minus_commute[of x a] * using \0 \ u\ \u \ 1\ by simp qed moreover have "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" if "dist a b = dist a x + dist x b" proof - let ?\ = "norm (a - x) / norm (a - b)" show "\u. x = (1 - u) *\<^sub>R a + u *\<^sub>R b \ 0 \ u \ u \ 1" proof (intro exI conjI) show "?\ \ 1" using Fal2 unfolding that[unfolded dist_norm] norm_ge_zero by auto show "x = (1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b" proof (subst euclidean_eq_iff; intro ballI) fix i :: 'a assume i: "i \ Basis" have "((1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b) \ i = ((norm (a - b) - norm (a - x)) * (a \ i) + norm (a - x) * (b \ i)) / norm (a - b)" using Fal by (auto simp add: field_simps inner_simps) also have "\ = x\i" apply (rule divide_eq_imp[OF Fal]) unfolding that[unfolded dist_norm] using that[unfolded dist_triangle_eq] i apply (subst (asm) euclidean_eq_iff) apply (auto simp add: field_simps inner_simps) done finally show "x \ i = ((1 - ?\) *\<^sub>R a + (?\) *\<^sub>R b) \ i" by auto qed qed (use Fal2 in auto) qed ultimately show ?thesis by (force simp add: between_def closed_segment_def dist_triangle_eq) qed lemma between_midpoint: fixes a :: "'a::euclidean_space" shows "between (a,b) (midpoint a b)" (is ?t1) and "between (b,a) (midpoint a b)" (is ?t2) proof - have *: "\x y z. x = (1/2::real) *\<^sub>R z \ y = (1/2) *\<^sub>R z \ norm z = norm x + norm y" by auto show ?t1 ?t2 unfolding between midpoint_def dist_norm by (auto simp add: field_simps inner_simps euclidean_eq_iff[where 'a='a] intro!: *) qed lemma between_mem_convex_hull: "between (a,b) x \ x \ convex hull {a,b}" unfolding between_mem_segment segment_convex_hull .. lemma between_triv_iff [simp]: "between (a,a) b \ a=b" by (auto simp: between_def) lemma between_triv1 [simp]: "between (a,b) a" by (auto simp: between_def) lemma between_triv2 [simp]: "between (a,b) b" by (auto simp: between_def) lemma between_commute: "between (a,b) = between (b,a)" by (auto simp: between_def closed_segment_commute) lemma between_antisym: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,c) b\ \ a = b" by (auto simp: between dist_commute) lemma between_trans: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,c) d\ \ between (b,c) d" using dist_triangle2 [of b c d] dist_triangle3 [of b d a] by (auto simp: between dist_commute) lemma between_norm: fixes a :: "'a :: euclidean_space" shows "between (a,b) x \ norm(x - a) *\<^sub>R (b - x) = norm(b - x) *\<^sub>R (x - a)" by (auto simp: between dist_triangle_eq norm_minus_commute algebra_simps) lemma between_swap: fixes A B X Y :: "'a::euclidean_space" assumes "between (A, B) X" assumes "between (A, B) Y" shows "between (X, B) Y \ between (A, Y) X" using assms by (auto simp add: between) lemma between_translation [simp]: "between (a + y,a + z) (a + x) \ between (y,z) x" by (auto simp: between_def) lemma between_trans_2: fixes a :: "'a :: euclidean_space" shows "\between (b,c) a; between (a,b) d\ \ between (c,d) a" by (metis between_commute between_swap between_trans) lemma between_scaleR_lift [simp]: fixes v :: "'a::euclidean_space" shows "between (a *\<^sub>R v, b *\<^sub>R v) (c *\<^sub>R v) \ v = 0 \ between (a, b) c" by (simp add: between dist_norm scaleR_left_diff_distrib [symmetric] distrib_right [symmetric]) lemma between_1: fixes x::real shows "between (a,b) x \ (a \ x \ x \ b) \ (b \ x \ x \ a)" by (auto simp: between_mem_segment closed_segment_eq_real_ivl) end \ No newline at end of file diff --git a/src/HOL/Analysis/Path_Connected.thy b/src/HOL/Analysis/Path_Connected.thy --- a/src/HOL/Analysis/Path_Connected.thy +++ b/src/HOL/Analysis/Path_Connected.thy @@ -1,4006 +1,4114 @@ (* Title: HOL/Analysis/Path_Connected.thy Authors: LC Paulson and Robert Himmelmann (TU Muenchen), based on material from HOL Light *) section \Path-Connectedness\ theory Path_Connected imports Starlike T1_Spaces begin subsection \Paths and Arcs\ definition\<^marker>\tag important\ path :: "(real \ 'a::topological_space) \ bool" where "path g \ continuous_on {0..1} g" definition\<^marker>\tag important\ pathstart :: "(real \ 'a::topological_space) \ 'a" where "pathstart g = g 0" definition\<^marker>\tag important\ pathfinish :: "(real \ 'a::topological_space) \ 'a" where "pathfinish g = g 1" definition\<^marker>\tag important\ path_image :: "(real \ 'a::topological_space) \ 'a set" where "path_image g = g ` {0 .. 1}" definition\<^marker>\tag important\ reversepath :: "(real \ 'a::topological_space) \ real \ 'a" where "reversepath g = (\x. g(1 - x))" definition\<^marker>\tag important\ joinpaths :: "(real \ 'a::topological_space) \ (real \ 'a) \ real \ 'a" (infixr "+++" 75) where "g1 +++ g2 = (\x. if x \ 1/2 then g1 (2 * x) else g2 (2 * x - 1))" definition\<^marker>\tag important\ simple_path :: "(real \ 'a::topological_space) \ bool" where "simple_path g \ path g \ (\x\{0..1}. \y\{0..1}. g x = g y \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" definition\<^marker>\tag important\ arc :: "(real \ 'a :: topological_space) \ bool" where "arc g \ path g \ inj_on g {0..1}" subsection\<^marker>\tag unimportant\\Invariance theorems\ lemma path_eq: "path p \ (\t. t \ {0..1} \ p t = q t) \ path q" using continuous_on_eq path_def by blast lemma path_continuous_image: "path g \ continuous_on (path_image g) f \ path(f \ g)" unfolding path_def path_image_def using continuous_on_compose by blast lemma path_translation_eq: fixes g :: "real \ 'a :: real_normed_vector" shows "path((\x. a + x) \ g) = path g" proof - have g: "g = (\x. -a + x) \ ((\x. a + x) \ g)" by (rule ext) simp show ?thesis unfolding path_def apply safe apply (subst g) apply (rule continuous_on_compose) apply (auto intro: continuous_intros) done qed lemma path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "path(f \ g) = path g" proof - from linear_injective_left_inverse [OF assms] obtain h where h: "linear h" "h \ f = id" by blast then have g: "g = h \ (f \ g)" by (metis comp_assoc id_comp) show ?thesis unfolding path_def using h assms by (metis g continuous_on_compose linear_continuous_on linear_conv_bounded_linear) qed lemma pathstart_translation: "pathstart((\x. a + x) \ g) = a + pathstart g" by (simp add: pathstart_def) lemma pathstart_linear_image_eq: "linear f \ pathstart(f \ g) = f(pathstart g)" by (simp add: pathstart_def) lemma pathfinish_translation: "pathfinish((\x. a + x) \ g) = a + pathfinish g" by (simp add: pathfinish_def) lemma pathfinish_linear_image: "linear f \ pathfinish(f \ g) = f(pathfinish g)" by (simp add: pathfinish_def) lemma path_image_translation: "path_image((\x. a + x) \ g) = (\x. a + x) ` (path_image g)" by (simp add: image_comp path_image_def) lemma path_image_linear_image: "linear f \ path_image(f \ g) = f ` (path_image g)" by (simp add: image_comp path_image_def) lemma reversepath_translation: "reversepath((\x. a + x) \ g) = (\x. a + x) \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma reversepath_linear_image: "linear f \ reversepath(f \ g) = f \ reversepath g" by (rule ext) (simp add: reversepath_def) lemma joinpaths_translation: "((\x. a + x) \ g1) +++ ((\x. a + x) \ g2) = (\x. a + x) \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma joinpaths_linear_image: "linear f \ (f \ g1) +++ (f \ g2) = f \ (g1 +++ g2)" by (rule ext) (simp add: joinpaths_def) lemma simple_path_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "simple_path((\x. a + x) \ g) = simple_path g" by (simp add: simple_path_def path_translation_eq) lemma simple_path_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "simple_path(f \ g) = simple_path g" using assms inj_on_eq_iff [of f] by (auto simp: path_linear_image_eq simple_path_def path_translation_eq) lemma arc_translation_eq: fixes g :: "real \ 'a::euclidean_space" shows "arc((\x. a + x) \ g) = arc g" by (auto simp: arc_def inj_on_def path_translation_eq) lemma arc_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "arc(f \ g) = arc g" using assms inj_on_eq_iff [of f] by (auto simp: arc_def inj_on_def path_linear_image_eq) subsection\<^marker>\tag unimportant\\Basic lemmas about paths\ lemma pathin_iff_path_real [simp]: "pathin euclideanreal g \ path g" by (simp add: pathin_def path_def) lemma continuous_on_path: "path f \ t \ {0..1} \ continuous_on t f" using continuous_on_subset path_def by blast lemma arc_imp_simple_path: "arc g \ simple_path g" by (simp add: arc_def inj_on_def simple_path_def) lemma arc_imp_path: "arc g \ path g" using arc_def by blast lemma arc_imp_inj_on: "arc g \ inj_on g {0..1}" by (auto simp: arc_def) lemma simple_path_imp_path: "simple_path g \ path g" using simple_path_def by blast lemma simple_path_cases: "simple_path g \ arc g \ pathfinish g = pathstart g" unfolding simple_path_def arc_def inj_on_def pathfinish_def pathstart_def by force lemma simple_path_imp_arc: "simple_path g \ pathfinish g \ pathstart g \ arc g" using simple_path_cases by auto lemma arc_distinct_ends: "arc g \ pathfinish g \ pathstart g" unfolding arc_def inj_on_def pathfinish_def pathstart_def by fastforce lemma arc_simple_path: "arc g \ simple_path g \ pathfinish g \ pathstart g" using arc_distinct_ends arc_imp_simple_path simple_path_cases by blast lemma simple_path_eq_arc: "pathfinish g \ pathstart g \ (simple_path g = arc g)" by (simp add: arc_simple_path) lemma path_image_const [simp]: "path_image (\t. a) = {a}" by (force simp: path_image_def) lemma path_image_nonempty [simp]: "path_image g \ {}" unfolding path_image_def image_is_empty box_eq_empty by auto lemma pathstart_in_path_image[intro]: "pathstart g \ path_image g" unfolding pathstart_def path_image_def by auto lemma pathfinish_in_path_image[intro]: "pathfinish g \ path_image g" unfolding pathfinish_def path_image_def by auto lemma connected_path_image[intro]: "path g \ connected (path_image g)" unfolding path_def path_image_def using connected_continuous_image connected_Icc by blast lemma compact_path_image[intro]: "path g \ compact (path_image g)" unfolding path_def path_image_def using compact_continuous_image connected_Icc by blast lemma reversepath_reversepath[simp]: "reversepath (reversepath g) = g" unfolding reversepath_def by auto lemma pathstart_reversepath[simp]: "pathstart (reversepath g) = pathfinish g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathfinish_reversepath[simp]: "pathfinish (reversepath g) = pathstart g" unfolding pathstart_def reversepath_def pathfinish_def by auto lemma pathstart_join[simp]: "pathstart (g1 +++ g2) = pathstart g1" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma pathfinish_join[simp]: "pathfinish (g1 +++ g2) = pathfinish g2" unfolding pathstart_def joinpaths_def pathfinish_def by auto lemma path_image_reversepath[simp]: "path_image (reversepath g) = path_image g" proof - have *: "\g. path_image (reversepath g) \ path_image g" unfolding path_image_def subset_eq reversepath_def Ball_def image_iff by force show ?thesis using *[of g] *[of "reversepath g"] unfolding reversepath_reversepath by auto qed lemma path_reversepath [simp]: "path (reversepath g) \ path g" proof - have *: "\g. path g \ path (reversepath g)" unfolding path_def reversepath_def apply (rule continuous_on_compose[unfolded o_def, of _ "\x. 1 - x"]) apply (auto intro: continuous_intros continuous_on_subset[of "{0..1}"]) done show ?thesis using *[of "reversepath g"] *[of g] unfolding reversepath_reversepath by (rule iffI) qed lemma arc_reversepath: assumes "arc g" shows "arc(reversepath g)" proof - have injg: "inj_on g {0..1}" using assms by (simp add: arc_def) have **: "\x y::real. 1-x = 1-y \ x = y" by simp show ?thesis using assms by (clarsimp simp: arc_def intro!: inj_onI) (simp add: inj_onD reversepath_def **) qed lemma simple_path_reversepath: "simple_path g \ simple_path (reversepath g)" apply (simp add: simple_path_def) apply (force simp: reversepath_def) done lemmas reversepath_simps = path_reversepath path_image_reversepath pathstart_reversepath pathfinish_reversepath lemma path_join[simp]: assumes "pathfinish g1 = pathstart g2" shows "path (g1 +++ g2) \ path g1 \ path g2" unfolding path_def pathfinish_def pathstart_def proof safe assume cont: "continuous_on {0..1} (g1 +++ g2)" have g1: "continuous_on {0..1} g1 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2))" by (intro continuous_on_cong refl) (auto simp: joinpaths_def) have g2: "continuous_on {0..1} g2 \ continuous_on {0..1} ((g1 +++ g2) \ (\x. x / 2 + 1/2))" using assms by (intro continuous_on_cong refl) (auto simp: joinpaths_def pathfinish_def pathstart_def) show "continuous_on {0..1} g1" and "continuous_on {0..1} g2" unfolding g1 g2 by (auto intro!: continuous_intros continuous_on_subset[OF cont] simp del: o_apply) next assume g1g2: "continuous_on {0..1} g1" "continuous_on {0..1} g2" have 01: "{0 .. 1} = {0..1/2} \ {1/2 .. 1::real}" by auto { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2) ` {0..1 / 2}" by (intro image_eqI[where x="x/2"]) auto } note 1 = this { fix x :: real assume "0 \ x" and "x \ 1" then have "x \ (\x. x * 2 - 1) ` {1 / 2..1}" by (intro image_eqI[where x="x/2 + 1/2"]) auto } note 2 = this show "continuous_on {0..1} (g1 +++ g2)" using assms unfolding joinpaths_def 01 apply (intro continuous_on_cases closed_atLeastAtMost g1g2[THEN continuous_on_compose2] continuous_intros) apply (auto simp: field_simps pathfinish_def pathstart_def intro!: 1 2) done qed subsection\<^marker>\tag unimportant\ \Path Images\ lemma bounded_path_image: "path g \ bounded(path_image g)" by (simp add: compact_imp_bounded compact_path_image) lemma closed_path_image: fixes g :: "real \ 'a::t2_space" shows "path g \ closed(path_image g)" by (metis compact_path_image compact_imp_closed) lemma connected_simple_path_image: "simple_path g \ connected(path_image g)" by (metis connected_path_image simple_path_imp_path) lemma compact_simple_path_image: "simple_path g \ compact(path_image g)" by (metis compact_path_image simple_path_imp_path) lemma bounded_simple_path_image: "simple_path g \ bounded(path_image g)" by (metis bounded_path_image simple_path_imp_path) lemma closed_simple_path_image: fixes g :: "real \ 'a::t2_space" shows "simple_path g \ closed(path_image g)" by (metis closed_path_image simple_path_imp_path) lemma connected_arc_image: "arc g \ connected(path_image g)" by (metis connected_path_image arc_imp_path) lemma compact_arc_image: "arc g \ compact(path_image g)" by (metis compact_path_image arc_imp_path) lemma bounded_arc_image: "arc g \ bounded(path_image g)" by (metis bounded_path_image arc_imp_path) lemma closed_arc_image: fixes g :: "real \ 'a::t2_space" shows "arc g \ closed(path_image g)" by (metis closed_path_image arc_imp_path) lemma path_image_join_subset: "path_image (g1 +++ g2) \ path_image g1 \ path_image g2" unfolding path_image_def joinpaths_def by auto lemma subset_path_image_join: assumes "path_image g1 \ s" and "path_image g2 \ s" shows "path_image (g1 +++ g2) \ s" using path_image_join_subset[of g1 g2] and assms by auto lemma path_image_join: "pathfinish g1 = pathstart g2 \ path_image(g1 +++ g2) = path_image g1 \ path_image g2" apply (rule subset_antisym [OF path_image_join_subset]) apply (auto simp: pathfinish_def pathstart_def path_image_def joinpaths_def image_def) apply (drule sym) apply (rule_tac x="xa/2" in bexI, auto) apply (rule ccontr) apply (drule_tac x="(xa+1)/2" in bspec) apply (auto simp: field_simps) apply (drule_tac x="1/2" in bspec, auto) done lemma not_in_path_image_join: assumes "x \ path_image g1" and "x \ path_image g2" shows "x \ path_image (g1 +++ g2)" using assms and path_image_join_subset[of g1 g2] by auto lemma pathstart_compose: "pathstart(f \ p) = f(pathstart p)" by (simp add: pathstart_def) lemma pathfinish_compose: "pathfinish(f \ p) = f(pathfinish p)" by (simp add: pathfinish_def) lemma path_image_compose: "path_image (f \ p) = f ` (path_image p)" by (simp add: image_comp path_image_def) lemma path_compose_join: "f \ (p +++ q) = (f \ p) +++ (f \ q)" by (rule ext) (simp add: joinpaths_def) lemma path_compose_reversepath: "f \ reversepath p = reversepath(f \ p)" by (rule ext) (simp add: reversepath_def) lemma joinpaths_eq: "(\t. t \ {0..1} \ p t = p' t) \ (\t. t \ {0..1} \ q t = q' t) \ t \ {0..1} \ (p +++ q) t = (p' +++ q') t" by (auto simp: joinpaths_def) lemma simple_path_inj_on: "simple_path g \ inj_on g {0<..<1}" by (auto simp: simple_path_def path_image_def inj_on_def less_eq_real_def Ball_def) subsection\<^marker>\tag unimportant\\Simple paths with the endpoints removed\ lemma simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} = c ` {0<..<1}" apply (auto simp: simple_path_def path_image_def pathstart_def pathfinish_def Ball_def Bex_def image_def) apply (metis eq_iff le_less_linear) apply (metis leD linear) using less_eq_real_def zero_le_one apply blast using less_eq_real_def zero_le_one apply blast done lemma connected_simple_path_endless: "simple_path c \ connected(path_image c - {pathstart c,pathfinish c})" apply (simp add: simple_path_endless) apply (rule connected_continuous_image) apply (meson continuous_on_subset greaterThanLessThan_subseteq_atLeastAtMost_iff le_numeral_extra(3) le_numeral_extra(4) path_def simple_path_imp_path) by auto lemma nonempty_simple_path_endless: "simple_path c \ path_image c - {pathstart c,pathfinish c} \ {}" by (simp add: simple_path_endless) subsection\<^marker>\tag unimportant\\The operations on paths\ lemma path_image_subset_reversepath: "path_image(reversepath g) \ path_image g" by (auto simp: path_image_def reversepath_def) lemma path_imp_reversepath: "path g \ path(reversepath g)" apply (auto simp: path_def reversepath_def) using continuous_on_compose [of "{0..1}" "\x. 1 - x" g] apply (auto simp: continuous_on_op_minus) done lemma half_bounded_equal: "1 \ x * 2 \ x * 2 \ 1 \ x = (1/2::real)" by simp lemma continuous_on_joinpaths: assumes "continuous_on {0..1} g1" "continuous_on {0..1} g2" "pathfinish g1 = pathstart g2" shows "continuous_on {0..1} (g1 +++ g2)" proof - have *: "{0..1::real} = {0..1/2} \ {1/2..1}" by auto have gg: "g2 0 = g1 1" by (metis assms(3) pathfinish_def pathstart_def) have 1: "continuous_on {0..1/2} (g1 +++ g2)" apply (rule continuous_on_eq [of _ "g1 \ (\x. 2*x)"]) apply (rule continuous_intros | simp add: joinpaths_def assms)+ done have "continuous_on {1/2..1} (g2 \ (\x. 2*x-1))" apply (rule continuous_on_subset [of "{1/2..1}"]) apply (rule continuous_intros | simp add: image_affinity_atLeastAtMost_diff assms)+ done then have 2: "continuous_on {1/2..1} (g1 +++ g2)" apply (rule continuous_on_eq [of "{1/2..1}" "g2 \ (\x. 2*x-1)"]) apply (rule assms continuous_intros | simp add: joinpaths_def mult.commute half_bounded_equal gg)+ done show ?thesis apply (subst *) apply (rule continuous_on_closed_Un) using 1 2 apply auto done qed lemma path_join_imp: "\path g1; path g2; pathfinish g1 = pathstart g2\ \ path(g1 +++ g2)" by (simp) lemma simple_path_join_loop: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "pathfinish g2 = pathstart g1" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" shows "simple_path(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g12: "g1 1 = g2 0" and g21: "g2 1 = g1 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g1 0, g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xyI: "x = 1 \ y \ 0" and xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy apply auto apply (drule inj_onD [OF injg1]) using g21 [symmetric] xyI apply (auto dest: inj_onD [OF injg2]) done } note * = this { fix x and y::real assume xy: "y \ 1" "0 \ x" "\ y * 2 \ 1" "x * 2 \ 1" "g1 (2 * x) = g2 (2 * y - 1)" have g1im: "g1 (2 * x) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x" in image_eqI, auto) done have "x = 0 \ y = 1" using subsetD [OF sb g1im] xy apply auto apply (force dest: inj_onD [OF injg1]) using g21 [symmetric] apply (auto dest: inj_onD [OF injg2]) done } note ** = this show ?thesis using assms apply (simp add: arc_def simple_path_def path_join, clarify) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis **) apply (force dest: inj_onD [OF injg2]) done qed lemma arc_join: assumes "arc g1" "arc g2" "pathfinish g1 = pathstart g2" "path_image g1 \ path_image g2 \ {pathstart g2}" shows "arc(g1 +++ g2)" proof - have injg1: "inj_on g1 {0..1}" using assms by (simp add: arc_def) have injg2: "inj_on g2 {0..1}" using assms by (simp add: arc_def) have g11: "g1 1 = g2 0" and sb: "g1 ` {0..1} \ g2 ` {0..1} \ {g2 0}" using assms by (simp_all add: arc_def pathfinish_def pathstart_def path_image_def) { fix x and y::real assume xy: "x \ 1" "0 \ y" " y * 2 \ 1" "\ x * 2 \ 1" "g2 (2 * x - 1) = g1 (2 * y)" have g1im: "g1 (2 * y) \ g1 ` {0..1} \ g2 ` {0..1}" using xy apply simp apply (rule_tac x="2 * x - 1" in image_eqI, auto) done have False using subsetD [OF sb g1im] xy by (auto dest: inj_onD [OF injg2]) } note * = this show ?thesis apply (simp add: arc_def inj_on_def) apply (clarsimp simp add: arc_imp_path assms) apply (simp add: joinpaths_def split: if_split_asm) apply (force dest: inj_onD [OF injg1]) apply (metis *) apply (metis *) apply (force dest: inj_onD [OF injg2]) done qed lemma reversepath_joinpaths: "pathfinish g1 = pathstart g2 \ reversepath(g1 +++ g2) = reversepath g2 +++ reversepath g1" unfolding reversepath_def pathfinish_def pathstart_def joinpaths_def by (rule ext) (auto simp: mult.commute) subsection\<^marker>\tag unimportant\\Some reversed and "if and only if" versions of joining theorems\ lemma path_join_path_ends: fixes g1 :: "real \ 'a::metric_space" assumes "path(g1 +++ g2)" "path g2" shows "pathfinish g1 = pathstart g2" proof (rule ccontr) define e where "e = dist (g1 1) (g2 0)" assume Neg: "pathfinish g1 \ pathstart g2" then have "0 < dist (pathfinish g1) (pathstart g2)" by auto then have "e > 0" by (metis e_def pathfinish_def pathstart_def) then obtain d1 where "d1 > 0" and d1: "\x'. \x'\{0..1}; norm x' < d1\ \ dist (g2 x') (g2 0) < e/2" using assms(2) unfolding path_def continuous_on_iff apply (drule_tac x=0 in bspec, simp) by (metis half_gt_zero_iff norm_conv_dist) obtain d2 where "d2 > 0" and d2: "\x'. \x'\{0..1}; dist x' (1/2) < d2\ \ dist ((g1 +++ g2) x') (g1 1) < e/2" using assms(1) \e > 0\ unfolding path_def continuous_on_iff apply (drule_tac x="1/2" in bspec, simp) apply (drule_tac x="e/2" in spec) apply (force simp: joinpaths_def) done have int01_1: "min (1/2) (min d1 d2) / 2 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist1: "norm (min (1 / 2) (min d1 d2) / 2) < d1" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have int01_2: "1/2 + min (1/2) (min d1 d2) / 4 \ {0..1}" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have dist2: "dist (1 / 2 + min (1 / 2) (min d1 d2) / 4) (1 / 2) < d2" using \d1 > 0\ \d2 > 0\ by (simp add: min_def dist_norm) have [simp]: "\ min (1 / 2) (min d1 d2) \ 0" using \d1 > 0\ \d2 > 0\ by (simp add: min_def) have "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g1 1) < e/2" "dist (g2 (min (1 / 2) (min d1 d2) / 2)) (g2 0) < e/2" using d1 [OF int01_1 dist1] d2 [OF int01_2 dist2] by (simp_all add: joinpaths_def) then have "dist (g1 1) (g2 0) < e/2 + e/2" using dist_triangle_half_r e_def by blast then show False by (simp add: e_def [symmetric]) qed lemma path_join_eq [simp]: fixes g1 :: "real \ 'a::metric_space" assumes "path g1" "path g2" shows "path(g1 +++ g2) \ pathfinish g1 = pathstart g2" using assms by (metis path_join_path_ends path_join_imp) lemma simple_path_joinE: assumes "simple_path(g1 +++ g2)" and "pathfinish g1 = pathstart g2" obtains "arc g1" "arc g2" "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" proof - have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have "path g1" using assms path_join simple_path_imp_path by blast moreover have "inj_on g1 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g1 x = g1 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "x/2" "y/2"] by (simp add: joinpaths_def split_ifs) qed ultimately have "arc g1" using assms by (simp add: arc_def) have [simp]: "g2 0 = g1 1" using assms by (metis pathfinish_def pathstart_def) have "path g2" using assms path_join simple_path_imp_path by blast moreover have "inj_on g2 {0..1}" proof (clarsimp simp: inj_on_def) fix x y assume "g2 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then show "x = y" using * [of "(x + 1) / 2" "(y + 1) / 2"] by (force simp: joinpaths_def split_ifs field_split_simps) qed ultimately have "arc g2" using assms by (simp add: arc_def) have "g2 y = g1 0 \ g2 y = g1 1" if "g1 x = g2 y" "0 \ x" "x \ 1" "0 \ y" "y \ 1" for x y using * [of "x / 2" "(y + 1) / 2"] that by (auto simp: joinpaths_def split_ifs field_split_simps) then have "path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (fastforce simp: pathstart_def pathfinish_def path_image_def) with \arc g1\ \arc g2\ show ?thesis using that by blast qed lemma simple_path_join_loop_eq: assumes "pathfinish g2 = pathstart g1" "pathfinish g1 = pathstart g2" shows "simple_path(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g1, pathstart g2}" by (metis assms simple_path_joinE simple_path_join_loop) lemma arc_join_eq: assumes "pathfinish g1 = pathstart g2" shows "arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 \ {pathstart g2}" (is "?lhs = ?rhs") proof assume ?lhs then have "simple_path(g1 +++ g2)" by (rule arc_imp_simple_path) then have *: "\x y. \0 \ x; x \ 1; 0 \ y; y \ 1; (g1 +++ g2) x = (g1 +++ g2) y\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) have False if "g1 0 = g2 u" "0 \ u" "u \ 1" for u using * [of 0 "(u + 1) / 2"] that assms arc_distinct_ends [OF \?lhs\] by (auto simp: joinpaths_def pathstart_def pathfinish_def split_ifs field_split_simps) then have n1: "pathstart g1 \ path_image g2" unfolding pathstart_def path_image_def using atLeastAtMost_iff by blast show ?rhs using \?lhs\ apply (rule simple_path_joinE [OF arc_imp_simple_path assms]) using n1 by force next assume ?rhs then show ?lhs using assms by (fastforce simp: pathfinish_def pathstart_def intro!: arc_join) qed lemma arc_join_eq_alt: "pathfinish g1 = pathstart g2 \ (arc(g1 +++ g2) \ arc g1 \ arc g2 \ path_image g1 \ path_image g2 = {pathstart g2})" using pathfinish_in_path_image by (fastforce simp: arc_join_eq) subsection\<^marker>\tag unimportant\\The joining of paths is associative\ lemma path_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ path(p +++ (q +++ r)) \ path((p +++ q) +++ r)" by simp lemma simple_path_assoc: assumes "pathfinish p = pathstart q" "pathfinish q = pathstart r" shows "simple_path (p +++ (q +++ r)) \ simple_path ((p +++ q) +++ r)" proof (cases "pathstart p = pathfinish r") case True show ?thesis proof assume "simple_path (p +++ q +++ r)" with assms True show "simple_path ((p +++ q) +++ r)" by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join dest: arc_distinct_ends [of r]) next assume 0: "simple_path ((p +++ q) +++ r)" with assms True have q: "pathfinish r \ path_image q" using arc_distinct_ends by (fastforce simp add: simple_path_join_loop_eq arc_join_eq path_image_join) have "pathstart r \ path_image p" using assms by (metis 0 IntI arc_distinct_ends arc_join_eq_alt empty_iff insert_iff pathfinish_in_path_image pathfinish_join simple_path_joinE) with assms 0 q True show "simple_path (p +++ q +++ r)" by (auto simp: simple_path_join_loop_eq arc_join_eq path_image_join dest!: subsetD [OF _ IntI]) qed next case False { fix x :: 'a assume a: "path_image p \ path_image q \ {pathstart q}" "(path_image p \ path_image q) \ path_image r \ {pathstart r}" "x \ path_image p" "x \ path_image r" have "pathstart r \ path_image q" by (metis assms(2) pathfinish_in_path_image) with a have "x = pathstart q" by blast } with False assms show ?thesis by (auto simp: simple_path_eq_arc simple_path_join_loop_eq arc_join_eq path_image_join) qed lemma arc_assoc: "\pathfinish p = pathstart q; pathfinish q = pathstart r\ \ arc(p +++ (q +++ r)) \ arc((p +++ q) +++ r)" by (simp add: arc_simple_path simple_path_assoc) subsubsection\<^marker>\tag unimportant\\Symmetry and loops\ lemma path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path(p +++ q) \ path(q +++ p)" by auto lemma simple_path_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ simple_path(p +++ q) \ simple_path(q +++ p)" by (metis (full_types) inf_commute insert_commute simple_path_joinE simple_path_join_loop) lemma path_image_sym: "\pathfinish p = pathstart q; pathfinish q = pathstart p\ \ path_image(p +++ q) = path_image(q +++ p)" by (simp add: path_image_join sup_commute) subsection\Subpath\ definition\<^marker>\tag important\ subpath :: "real \ real \ (real \ 'a) \ real \ 'a::real_normed_vector" where "subpath a b g \ \x. g((b - a) * x + a)" lemma path_image_subpath_gen: fixes g :: "_ \ 'a::real_normed_vector" shows "path_image(subpath u v g) = g ` (closed_segment u v)" by (auto simp add: closed_segment_real_eq path_image_def subpath_def) lemma path_image_subpath: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = (if u \ v then g ` {u..v} else g ` {v..u})" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_image_subpath_commute: fixes g :: "real \ 'a::real_normed_vector" shows "path_image(subpath u v g) = path_image(subpath v u g)" by (simp add: path_image_subpath_gen closed_segment_eq_real_ivl) lemma path_subpath [simp]: fixes g :: "real \ 'a::real_normed_vector" assumes "path g" "u \ {0..1}" "v \ {0..1}" shows "path(subpath u v g)" proof - have "continuous_on {0..1} (g \ (\x. ((v-u) * x+ u)))" apply (rule continuous_intros | simp)+ apply (simp add: image_affinity_atLeastAtMost [where c=u]) using assms apply (auto simp: path_def continuous_on_subset) done then show ?thesis by (simp add: path_def subpath_def) qed lemma pathstart_subpath [simp]: "pathstart(subpath u v g) = g(u)" by (simp add: pathstart_def subpath_def) lemma pathfinish_subpath [simp]: "pathfinish(subpath u v g) = g(v)" by (simp add: pathfinish_def subpath_def) lemma subpath_trivial [simp]: "subpath 0 1 g = g" by (simp add: subpath_def) lemma subpath_reversepath: "subpath 1 0 g = reversepath g" by (simp add: reversepath_def subpath_def) lemma reversepath_subpath: "reversepath(subpath u v g) = subpath v u g" by (simp add: reversepath_def subpath_def algebra_simps) lemma subpath_translation: "subpath u v ((\x. a + x) \ g) = (\x. a + x) \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma subpath_image: "subpath u v (f \ g) = f \ subpath u v g" by (rule ext) (simp add: subpath_def) lemma affine_ineq: fixes x :: "'a::linordered_idom" assumes "x \ 1" "v \ u" shows "v + x * u \ u + x * v" proof - have "(1-x)*(u-v) \ 0" using assms by auto then show ?thesis by (simp add: algebra_simps) qed lemma sum_le_prod1: fixes a::real shows "\a \ 1; b \ 1\ \ a + b \ 1 + a * b" by (metis add.commute affine_ineq mult.right_neutral) lemma simple_path_subpath_eq: "simple_path(subpath u v g) \ path(subpath u v g) \ u\v \ (\x y. x \ closed_segment u v \ y \ closed_segment u v \ g x = g y \ x = y \ x = u \ y = v \ x = v \ y = u)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0)" by (auto simp: simple_path_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y \ x = u \ y = v \ x = v \ y = u" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (auto split: if_split_asm simp add: closed_segment_real_eq image_affinity_atLeastAtMost) (simp_all add: field_split_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y \ x = u \ y = v \ x = v \ y = u" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y \ x = u \ y = v \ x = v \ y = u" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: closed_segment_real_eq image_affinity_atLeastAtMost) have [simp]: "\x. u + x * v = v + x * u \ u=v \ x=1" by algebra show ?lhs using psp ne unfolding simple_path_def subpath_def by (fastforce simp add: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma arc_subpath_eq: "arc(subpath u v g) \ path(subpath u v g) \ u\v \ inj_on g (closed_segment u v)" (is "?lhs = ?rhs") proof (rule iffI) assume ?lhs then have p: "path (\x. g ((v - u) * x + u))" and sim: "(\x y. \x\{0..1}; y\{0..1}; g ((v - u) * x + u) = g ((v - u) * y + u)\ \ x = y)" by (auto simp: arc_def inj_on_def subpath_def) { fix x y assume "x \ closed_segment u v" "y \ closed_segment u v" "g x = g y" then have "x = y" using sim [of "(x-u)/(v-u)" "(y-u)/(v-u)"] p by (cases "v = u") (simp_all split: if_split_asm add: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost, simp add: field_simps) } moreover have "path(subpath u v g) \ u\v" using sim [of "1/3" "2/3"] p by (auto simp: subpath_def) ultimately show ?rhs unfolding inj_on_def by metis next assume ?rhs then have d1: "\x y. \g x = g y; u \ x; x \ v; u \ y; y \ v\ \ x = y" and d2: "\x y. \g x = g y; v \ x; x \ u; v \ y; y \ u\ \ x = y" and ne: "u < v \ v < u" and psp: "path (subpath u v g)" by (auto simp: inj_on_def closed_segment_real_eq image_affinity_atLeastAtMost) show ?lhs using psp ne unfolding arc_def subpath_def inj_on_def by (auto simp: algebra_simps affine_ineq mult_left_mono crossproduct_eq dest: d1 d2) qed lemma simple_path_subpath: assumes "simple_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "simple_path(subpath u v g)" using assms apply (simp add: simple_path_subpath_eq simple_path_imp_path) apply (simp add: simple_path_def closed_segment_real_eq image_affinity_atLeastAtMost, fastforce) done lemma arc_simple_path_subpath: "\simple_path g; u \ {0..1}; v \ {0..1}; g u \ g v\ \ arc(subpath u v g)" by (force intro: simple_path_subpath simple_path_imp_arc) lemma arc_subpath_arc: "\arc g; u \ {0..1}; v \ {0..1}; u \ v\ \ arc(subpath u v g)" by (meson arc_def arc_imp_simple_path arc_simple_path_subpath inj_onD) lemma arc_simple_path_subpath_interior: "\simple_path g; u \ {0..1}; v \ {0..1}; u \ v; \u-v\ < 1\ \ arc(subpath u v g)" apply (rule arc_simple_path_subpath) apply (force simp: simple_path_def)+ done lemma path_image_subpath_subset: "\u \ {0..1}; v \ {0..1}\ \ path_image(subpath u v g) \ path_image g" apply (simp add: closed_segment_real_eq image_affinity_atLeastAtMost path_image_subpath) apply (auto simp: path_image_def) done lemma join_subpaths_middle: "subpath (0) ((1 / 2)) p +++ subpath ((1 / 2)) 1 p = p" by (rule ext) (simp add: joinpaths_def subpath_def field_split_simps) subsection\<^marker>\tag unimportant\\There is a subpath to the frontier\ lemma subpath_to_frontier_explicit: fixes S :: "'a::metric_space set" assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "\x. 0 \ x \ x < u \ g x \ interior S" "(g u \ interior S)" "(u = 0 \ g u \ closure S)" proof - have gcon: "continuous_on {0..1} g" using g by (simp add: path_def) then have com: "compact ({0..1} \ {u. g u \ closure (- S)})" apply (simp add: Int_commute [of "{0..1}"] compact_eq_bounded_closed closed_vimage_Int [unfolded vimage_def]) using compact_eq_bounded_closed apply fastforce done have "1 \ {u. g u \ closure (- S)}" using assms by (simp add: pathfinish_def closure_def) then have dis: "{0..1} \ {u. g u \ closure (- S)} \ {}" using atLeastAtMost_iff zero_le_one by blast then obtain u where "0 \ u" "u \ 1" and gu: "g u \ closure (- S)" and umin: "\t. \0 \ t; t \ 1; g t \ closure (- S)\ \ u \ t" using compact_attains_inf [OF com dis] by fastforce then have umin': "\t. \0 \ t; t \ 1; t < u\ \ g t \ S" using closure_def by fastforce { assume "u \ 0" then have "u > 0" using \0 \ u\ by auto { fix e::real assume "e > 0" obtain d where "d>0" and d: "\x'. \x' \ {0..1}; dist x' u \ d\ \ dist (g x') (g u) < e" using continuous_onE [OF gcon _ \e > 0\] \0 \ _\ \_ \ 1\ atLeastAtMost_iff by auto have *: "dist (max 0 (u - d / 2)) u \ d" using \0 \ u\ \u \ 1\ \d > 0\ by (simp add: dist_real_def) have "\y\S. dist y (g u) < e" using \0 < u\ \u \ 1\ \d > 0\ by (force intro: d [OF _ *] umin') } then have "g u \ closure S" by (simp add: frontier_def closure_approachable) } then show ?thesis apply (rule_tac u=u in that) apply (auto simp: \0 \ u\ \u \ 1\ gu interior_closure umin) using \_ \ 1\ interior_closure umin apply fastforce done qed lemma subpath_to_frontier_strong: assumes g: "path g" and "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ interior S" "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" proof - obtain u where "0 \ u" "u \ 1" and gxin: "\x. 0 \ x \ x < u \ g x \ interior S" and gunot: "(g u \ interior S)" and u0: "(u = 0 \ g u \ closure S)" using subpath_to_frontier_explicit [OF assms] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (simp add: gunot) using \0 \ u\ u0 by (force simp: subpath_def gxin) qed lemma subpath_to_frontier: assumes g: "path g" and g0: "pathstart g \ closure S" and g1: "pathfinish g \ S" obtains u where "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" proof - obtain u where "0 \ u" "u \ 1" and notin: "g u \ interior S" and disj: "u = 0 \ (\x. 0 \ x \ x < 1 \ subpath 0 u g x \ interior S) \ g u \ closure S" using subpath_to_frontier_strong [OF g g1] by blast show ?thesis apply (rule that [OF \0 \ u\ \u \ 1\]) apply (metis DiffI disj frontier_def g0 notin pathstart_def) using \0 \ u\ g0 disj apply (simp add: path_image_subpath_gen) apply (auto simp: closed_segment_eq_real_ivl pathstart_def pathfinish_def subpath_def) apply (rename_tac y) apply (drule_tac x="y/u" in spec) apply (auto split: if_split_asm) done qed lemma exists_path_subpath_to_frontier: fixes S :: "'a::real_normed_vector set" assumes "path g" "pathstart g \ closure S" "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" proof - obtain u where u: "0 \ u" "u \ 1" "g u \ frontier S" "(path_image(subpath 0 u g) - {g u}) \ interior S" using subpath_to_frontier [OF assms] by blast show ?thesis apply (rule that [of "subpath 0 u g"]) using assms u apply (simp_all add: path_image_subpath) apply (simp add: pathstart_def) apply (force simp: closed_segment_eq_real_ivl path_image_def) done qed lemma exists_path_subpath_to_frontier_closed: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" and g: "path g" and g0: "pathstart g \ S" and g1: "pathfinish g \ S" obtains h where "path h" "pathstart h = pathstart g" "path_image h \ path_image g \ S" "pathfinish h \ frontier S" proof - obtain h where h: "path h" "pathstart h = pathstart g" "path_image h \ path_image g" "path_image h - {pathfinish h} \ interior S" "pathfinish h \ frontier S" using exists_path_subpath_to_frontier [OF g _ g1] closure_closed [OF S] g0 by auto show ?thesis apply (rule that [OF \path h\]) using assms h apply auto apply (metis Diff_single_insert frontier_subset_eq insert_iff interior_subset subset_iff) done qed subsection \Shift Path to Start at Some Given Point\ definition\<^marker>\tag important\ shiftpath :: "real \ (real \ 'a::topological_space) \ real \ 'a" where "shiftpath a f = (\x. if (a + x) \ 1 then f (a + x) else f (a + x - 1))" +lemma 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 pathstart_shiftpath: "a \ 1 \ pathstart (shiftpath a g) = g a" unfolding pathstart_def shiftpath_def by auto lemma pathfinish_shiftpath: assumes "0 \ a" and "pathfinish g = pathstart g" shows "pathfinish (shiftpath a g) = g a" using assms unfolding pathstart_def pathfinish_def shiftpath_def by auto lemma endpoints_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0 .. 1}" shows "pathfinish (shiftpath a g) = g a" and "pathstart (shiftpath a g) = g a" using assms by (auto intro!: pathfinish_shiftpath pathstart_shiftpath) lemma closed_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" shows "pathfinish (shiftpath a g) = pathstart (shiftpath a g)" using endpoints_shiftpath[OF assms] by auto lemma path_shiftpath: assumes "path g" and "pathfinish g = pathstart g" and "a \ {0..1}" shows "path (shiftpath a g)" proof - have *: "{0 .. 1} = {0 .. 1-a} \ {1-a .. 1}" using assms(3) by auto have **: "\x. x + a = 1 \ g (x + a - 1) = g (x + a)" using assms(2)[unfolded pathfinish_def pathstart_def] by auto show ?thesis unfolding path_def shiftpath_def * proof (rule continuous_on_closed_Un) have contg: "continuous_on {0..1} g" using \path g\ path_def by blast show "continuous_on {0..1-a} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {0..1-a} (g \ (+) a)" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed auto show "continuous_on {1-a..1} (\x. if a + x \ 1 then g (a + x) else g (a + x - 1))" proof (rule continuous_on_eq) show "continuous_on {1-a..1} (g \ (+) (a - 1))" by (intro continuous_intros continuous_on_subset [OF contg]) (use \a \ {0..1}\ in auto) qed (auto simp: "**" add.commute add_diff_eq) qed auto qed lemma shiftpath_shiftpath: assumes "pathfinish g = pathstart g" and "a \ {0..1}" and "x \ {0..1}" shows "shiftpath (1 - a) (shiftpath a g) x = g x" using assms unfolding pathfinish_def pathstart_def shiftpath_def by auto lemma path_image_shiftpath: assumes a: "a \ {0..1}" and "pathfinish g = pathstart g" shows "path_image (shiftpath a g) = path_image g" proof - { fix x assume g: "g 1 = g 0" "x \ {0..1::real}" and gne: "\y. y\{0..1} \ {x. \ a + x \ 1} \ g x \ g (a + y - 1)" then have "\y\{0..1} \ {x. a + x \ 1}. g x = g (a + y)" proof (cases "a \ x") case False then show ?thesis apply (rule_tac x="1 + x - a" in bexI) using g gne[of "1 + x - a"] a apply (force simp: field_simps)+ done next case True then show ?thesis using g a by (rule_tac x="x - a" in bexI) (auto simp: field_simps) qed } then show ?thesis using assms unfolding shiftpath_def path_image_def pathfinish_def pathstart_def by (auto simp: image_iff) qed lemma simple_path_shiftpath: assumes "simple_path g" "pathfinish g = pathstart g" and a: "0 \ a" "a \ 1" shows "simple_path (shiftpath a g)" unfolding simple_path_def proof (intro conjI impI ballI) show "path (shiftpath a g)" by (simp add: assms path_shiftpath simple_path_imp_path) have *: "\x y. \g x = g y; x \ {0..1}; y \ {0..1}\ \ x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" using assms by (simp add: simple_path_def) show "x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0" if "x \ {0..1}" "y \ {0..1}" "shiftpath a g x = shiftpath a g y" for x y using that a unfolding shiftpath_def by (force split: if_split_asm dest!: *) qed subsection \Straight-Line Paths\ definition\<^marker>\tag important\ linepath :: "'a::real_normed_vector \ 'a \ real \ 'a" where "linepath a b = (\x. (1 - x) *\<^sub>R a + x *\<^sub>R b)" lemma pathstart_linepath[simp]: "pathstart (linepath a b) = a" unfolding pathstart_def linepath_def by auto lemma pathfinish_linepath[simp]: "pathfinish (linepath a b) = b" unfolding pathfinish_def linepath_def by auto lemma linepath_inner: "linepath a b x \ v = linepath (a \ v) (b \ v) x" by (simp add: linepath_def algebra_simps) lemma Re_linepath': "Re (linepath a b x) = linepath (Re a) (Re b) x" by (simp add: linepath_def) lemma Im_linepath': "Im (linepath a b x) = linepath (Im a) (Im b) x" by (simp add: linepath_def) lemma linepath_0': "linepath a b 0 = a" by (simp add: linepath_def) lemma linepath_1': "linepath a b 1 = b" by (simp add: linepath_def) lemma continuous_linepath_at[intro]: "continuous (at x) (linepath a b)" unfolding linepath_def by (intro continuous_intros) lemma continuous_on_linepath [intro,continuous_intros]: "continuous_on s (linepath a b)" using continuous_linepath_at by (auto intro!: continuous_at_imp_continuous_on) lemma path_linepath[iff]: "path (linepath a b)" unfolding path_def by (rule continuous_on_linepath) lemma path_image_linepath[simp]: "path_image (linepath a b) = closed_segment a b" unfolding path_image_def segment linepath_def by auto lemma reversepath_linepath[simp]: "reversepath (linepath a b) = linepath b a" unfolding reversepath_def linepath_def by auto lemma linepath_0 [simp]: "linepath 0 b x = x *\<^sub>R b" by (simp add: linepath_def) lemma linepath_cnj: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma arc_linepath: assumes "a \ b" shows [simp]: "arc (linepath a b)" proof - { fix x y :: "real" assume "x *\<^sub>R b + y *\<^sub>R a = x *\<^sub>R a + y *\<^sub>R b" then have "(x - y) *\<^sub>R a = (x - y) *\<^sub>R b" by (simp add: algebra_simps) with assms have "x = y" by simp } then show ?thesis unfolding arc_def inj_on_def by (fastforce simp: algebra_simps linepath_def) qed lemma simple_path_linepath[intro]: "a \ b \ simple_path (linepath a b)" by (simp add: arc_imp_simple_path) lemma linepath_trivial [simp]: "linepath a a x = a" by (simp add: linepath_def real_vector.scale_left_diff_distrib) lemma linepath_refl: "linepath a a = (\x. a)" by auto lemma subpath_refl: "subpath a a g = linepath (g a) (g a)" by (simp add: subpath_def linepath_def algebra_simps) lemma linepath_of_real: "(linepath (of_real a) (of_real b) x) = of_real ((1 - x)*a + x*b)" by (simp add: scaleR_conv_of_real linepath_def) lemma of_real_linepath: "of_real (linepath a b x) = linepath (of_real a) (of_real b) x" by (metis linepath_of_real mult.right_neutral of_real_def real_scaleR_def) lemma inj_on_linepath: assumes "a \ b" shows "inj_on (linepath a b) {0..1}" proof (clarsimp simp: inj_on_def linepath_def) fix x y assume "(1 - x) *\<^sub>R a + x *\<^sub>R b = (1 - y) *\<^sub>R a + y *\<^sub>R b" "0 \ x" "x \ 1" "0 \ y" "y \ 1" then have "x *\<^sub>R (a - b) = y *\<^sub>R (a - b)" by (auto simp: algebra_simps) then show "x=y" using assms by auto qed lemma linepath_le_1: fixes a::"'a::linordered_idom" shows "\a \ 1; b \ 1; 0 \ u; u \ 1\ \ (1 - u) * a + u * b \ 1" using mult_left_le [of a "1-u"] mult_left_le [of b u] by auto +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 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 linepath_cnj': "cnj \ linepath a b = linepath (cnj a) (cnj b)" + by (simp add: linepath_def fun_eq_iff) + +lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" + by (auto simp: linepath_def) + +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 + subsection\<^marker>\tag unimportant\\Segments via convex hulls\ lemma segments_subset_convex_hull: "closed_segment a b \ (convex hull {a,b,c})" "closed_segment a c \ (convex hull {a,b,c})" "closed_segment b c \ (convex hull {a,b,c})" "closed_segment b a \ (convex hull {a,b,c})" "closed_segment c a \ (convex hull {a,b,c})" "closed_segment c b \ (convex hull {a,b,c})" by (auto simp: segment_convex_hull linepath_of_real elim!: rev_subsetD [OF _ hull_mono]) lemma midpoints_in_convex_hull: assumes "x \ convex hull s" "y \ convex hull s" shows "midpoint x y \ convex hull s" proof - have "(1 - inverse(2)) *\<^sub>R x + inverse(2) *\<^sub>R y \ convex hull s" by (rule convexD_alt) (use assms in auto) then show ?thesis by (simp add: midpoint_def algebra_simps) qed lemma not_in_interior_convex_hull_3: fixes a :: "complex" shows "a \ interior(convex hull {a,b,c})" "b \ interior(convex hull {a,b,c})" "c \ interior(convex hull {a,b,c})" by (auto simp: card_insert_le_m1 not_in_interior_convex_hull) lemma midpoint_in_closed_segment [simp]: "midpoint a b \ closed_segment a b" using midpoints_in_convex_hull segment_convex_hull by blast lemma midpoint_in_open_segment [simp]: "midpoint a b \ open_segment a b \ a \ b" by (simp add: open_segment_def) lemma continuous_IVT_local_extremum: fixes f :: "'a::euclidean_space \ real" assumes contf: "continuous_on (closed_segment a b) f" and "a \ b" "f a = f b" obtains z where "z \ open_segment a b" "(\w \ closed_segment a b. (f w) \ (f z)) \ (\w \ closed_segment a b. (f z) \ (f w))" proof - obtain c where "c \ closed_segment a b" and c: "\y. y \ closed_segment a b \ f y \ f c" using continuous_attains_sup [of "closed_segment a b" f] contf by auto obtain d where "d \ closed_segment a b" and d: "\y. y \ closed_segment a b \ f d \ f y" using continuous_attains_inf [of "closed_segment a b" f] contf by auto show ?thesis proof (cases "c \ open_segment a b \ d \ open_segment a b") case True then show ?thesis using c d that by blast next case False then have "(c = a \ c = b) \ (d = a \ d = b)" by (simp add: \c \ closed_segment a b\ \d \ closed_segment a b\ open_segment_def) with \a \ b\ \f a = f b\ c d show ?thesis by (rule_tac z = "midpoint a b" in that) (fastforce+) qed qed text\An injective map into R is also an open map w.r.T. the universe, and conversely. \ proposition injective_eq_1d_open_map_UNIV: fixes f :: "real \ real" assumes contf: "continuous_on S f" and S: "is_interval S" shows "inj_on f S \ (\T. open T \ T \ S \ open(f ` T))" (is "?lhs = ?rhs") proof safe fix T assume injf: ?lhs and "open T" and "T \ S" have "\U. open U \ f x \ U \ U \ f ` T" if "x \ T" for x proof - obtain \ where "\ > 0" and \: "cball x \ \ T" using \open T\ \x \ T\ open_contains_cball_eq by blast show ?thesis proof (intro exI conjI) have "closed_segment (x-\) (x+\) = {x-\..x+\}" using \0 < \\ by (auto simp: closed_segment_eq_real_ivl) also have "\ \ S" using \ \T \ S\ by (auto simp: dist_norm subset_eq) finally have "f ` (open_segment (x-\) (x+\)) = open_segment (f (x-\)) (f (x+\))" using continuous_injective_image_open_segment_1 by (metis continuous_on_subset [OF contf] inj_on_subset [OF injf]) then show "open (f ` {x-\<..})" using \0 < \\ by (simp add: open_segment_eq_real_ivl) show "f x \ f ` {x - \<..}" by (auto simp: \\ > 0\) show "f ` {x - \<..} \ f ` T" using \ by (auto simp: dist_norm subset_iff) qed qed with open_subopen show "open (f ` T)" by blast next assume R: ?rhs have False if xy: "x \ S" "y \ S" and "f x = f y" "x \ y" for x y proof - have "open (f ` open_segment x y)" using R by (metis S convex_contains_open_segment is_interval_convex open_greaterThanLessThan open_segment_eq_real_ivl xy) moreover have "continuous_on (closed_segment x y) f" by (meson S closed_segment_subset contf continuous_on_subset is_interval_convex that) then obtain \ where "\ \ open_segment x y" and \: "(\w \ closed_segment x y. (f w) \ (f \)) \ (\w \ closed_segment x y. (f \) \ (f w))" using continuous_IVT_local_extremum [of x y f] \f x = f y\ \x \ y\ by blast ultimately obtain e where "e>0" and e: "\u. dist u (f \) < e \ u \ f ` open_segment x y" using open_dist by (metis image_eqI) have fin: "f \ + (e/2) \ f ` open_segment x y" "f \ - (e/2) \ f ` open_segment x y" using e [of "f \ + (e/2)"] e [of "f \ - (e/2)"] \e > 0\ by (auto simp: dist_norm) show ?thesis using \ \0 < e\ fin open_closed_segment by fastforce qed then show ?lhs by (force simp: inj_on_def) qed subsection\<^marker>\tag unimportant\ \Bounding a point away from a path\ lemma not_on_path_ball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and z: "z \ path_image g" shows "\e > 0. ball z e \ path_image g = {}" proof - have "closed (path_image g)" by (simp add: \path g\ closed_path_image) then obtain a where "a \ path_image g" "\y \ path_image g. dist z a \ dist z y" by (auto intro: distance_attains_inf[OF _ path_image_nonempty, of g z]) then show ?thesis by (rule_tac x="dist z a" in exI) (use dist_commute z in auto) qed lemma not_on_path_cball: fixes g :: "real \ 'a::heine_borel" assumes "path g" and "z \ path_image g" shows "\e>0. cball z e \ (path_image g) = {}" proof - obtain e where "ball z e \ path_image g = {}" "e > 0" using not_on_path_ball[OF assms] by auto moreover have "cball z (e/2) \ ball z e" using \e > 0\ by auto ultimately show ?thesis by (rule_tac x="e/2" in exI) auto qed subsection \Path component\ text \Original formalization by Tom Hales\ definition\<^marker>\tag important\ "path_component s x y \ (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" abbreviation\<^marker>\tag important\ "path_component_set s x \ Collect (path_component s x)" lemmas path_defs = path_def pathstart_def pathfinish_def path_image_def path_component_def lemma path_component_mem: assumes "path_component s x y" shows "x \ s" and "y \ s" using assms unfolding path_defs by auto lemma path_component_refl: assumes "x \ s" shows "path_component s x x" unfolding path_defs apply (rule_tac x="\u. x" in exI) using assms apply (auto intro!: continuous_intros) done lemma path_component_refl_eq: "path_component s x x \ x \ s" by (auto intro!: path_component_mem path_component_refl) lemma path_component_sym: "path_component s x y \ path_component s y x" unfolding path_component_def apply (erule exE) apply (rule_tac x="reversepath g" in exI, auto) done lemma path_component_trans: assumes "path_component s x y" and "path_component s y z" shows "path_component s x z" using assms unfolding path_component_def apply (elim exE) apply (rule_tac x="g +++ ga" in exI) apply (auto simp: path_image_join) done lemma path_component_of_subset: "s \ t \ path_component s x y \ path_component t x y" unfolding path_component_def by auto lemma path_component_linepath: fixes s :: "'a::real_normed_vector set" shows "closed_segment a b \ s \ path_component s a b" unfolding path_component_def by (rule_tac x="linepath a b" in exI, auto) subsubsection\<^marker>\tag unimportant\ \Path components as sets\ lemma path_component_set: "path_component_set s x = {y. (\g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)}" by (auto simp: path_component_def) lemma path_component_subset: "path_component_set s x \ s" by (auto simp: path_component_mem(2)) lemma path_component_eq_empty: "path_component_set s x = {} \ x \ s" using path_component_mem path_component_refl_eq by fastforce lemma path_component_mono: "s \ t \ (path_component_set s x) \ (path_component_set t x)" by (simp add: Collect_mono path_component_of_subset) lemma path_component_eq: "y \ path_component_set s x \ path_component_set s y = path_component_set s x" by (metis (no_types, lifting) Collect_cong mem_Collect_eq path_component_sym path_component_trans) subsection \Path connectedness of a space\ definition\<^marker>\tag important\ "path_connected s \ (\x\s. \y\s. \g. path g \ path_image g \ s \ pathstart g = x \ pathfinish g = y)" lemma path_connectedin_iff_path_connected_real [simp]: "path_connectedin euclideanreal S \ path_connected S" by (simp add: path_connectedin path_connected_def path_defs) lemma path_connected_component: "path_connected s \ (\x\s. \y\s. path_component s x y)" unfolding path_connected_def path_component_def by auto lemma path_connected_component_set: "path_connected s \ (\x\s. path_component_set s x = s)" unfolding path_connected_component path_component_subset using path_component_mem by blast lemma path_component_maximal: "\x \ t; path_connected t; t \ s\ \ t \ (path_component_set s x)" by (metis path_component_mono path_connected_component_set) lemma convex_imp_path_connected: fixes s :: "'a::real_normed_vector set" assumes "convex s" shows "path_connected s" unfolding path_connected_def using assms convex_contains_segment by fastforce lemma path_connected_UNIV [iff]: "path_connected (UNIV :: 'a::real_normed_vector set)" by (simp add: convex_imp_path_connected) lemma path_component_UNIV: "path_component_set UNIV x = (UNIV :: 'a::real_normed_vector set)" using path_connected_component_set by auto lemma path_connected_imp_connected: assumes "path_connected S" shows "connected S" proof (rule connectedI) fix e1 e2 assume as: "open e1" "open e2" "S \ e1 \ e2" "e1 \ e2 \ S = {}" "e1 \ S \ {}" "e2 \ S \ {}" then obtain x1 x2 where obt:"x1 \ e1 \ S" "x2 \ e2 \ S" by auto then obtain g where g: "path g" "path_image g \ S" "pathstart g = x1" "pathfinish g = x2" using assms[unfolded path_connected_def,rule_format,of x1 x2] by auto have *: "connected {0..1::real}" by (auto intro!: convex_connected) have "{0..1} \ {x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2}" using as(3) g(2)[unfolded path_defs] by blast moreover have "{x \ {0..1}. g x \ e1} \ {x \ {0..1}. g x \ e2} = {}" using as(4) g(2)[unfolded path_defs] unfolding subset_eq by auto moreover have "{x \ {0..1}. g x \ e1} \ {} \ {x \ {0..1}. g x \ e2} \ {}" using g(3,4)[unfolded path_defs] using obt by (simp add: ex_in_conv [symmetric], metis zero_le_one order_refl) ultimately show False using *[unfolded connected_local not_ex, rule_format, of "{0..1} \ g -` e1" "{0..1} \ g -` e2"] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(1)] using continuous_openin_preimage_gen[OF g(1)[unfolded path_def] as(2)] by auto qed lemma open_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (path_component_set S x)" unfolding open_contains_ball proof fix y assume as: "y \ path_component_set S x" then have "y \ S" by (simp add: path_component_mem(2)) then obtain e where e: "e > 0" "ball y e \ S" using assms[unfolded open_contains_ball] by auto have "\u. dist y u < e \ path_component S x u" by (metis (full_types) as centre_in_ball convex_ball convex_imp_path_connected e mem_Collect_eq mem_ball path_component_eq path_component_of_subset path_connected_component) then show "\e > 0. ball y e \ path_component_set S x" using \e>0\ by auto qed lemma open_non_path_component: fixes S :: "'a::real_normed_vector set" assumes "open S" shows "open (S - path_component_set S x)" unfolding open_contains_ball proof fix y assume y: "y \ S - path_component_set S x" then obtain e where e: "e > 0" "ball y e \ S" using assms openE by auto show "\e>0. ball y e \ S - path_component_set S x" proof (intro exI conjI subsetI DiffI notI) show "\x. x \ ball y e \ x \ S" using e by blast show False if "z \ ball y e" "z \ path_component_set S x" for z proof - have "y \ path_component_set S z" by (meson assms convex_ball convex_imp_path_connected e open_contains_ball_eq open_path_component path_component_maximal that(1)) then have "y \ path_component_set S x" using path_component_eq that(2) by blast then show False using y by blast qed qed (use e in auto) qed lemma connected_open_path_connected: fixes S :: "'a::real_normed_vector set" assumes "open S" and "connected S" shows "path_connected S" unfolding path_connected_component_set proof (rule, rule, rule path_component_subset, rule) fix x y assume "x \ S" and "y \ S" show "y \ path_component_set S x" proof (rule ccontr) assume "\ ?thesis" moreover have "path_component_set S x \ S \ {}" using \x \ S\ path_component_eq_empty path_component_subset[of S x] by auto ultimately show False using \y \ S\ open_non_path_component[OF assms(1)] open_path_component[OF assms(1)] using assms(2)[unfolded connected_def not_ex, rule_format, of "path_component_set S x" "S - path_component_set S x"] by auto qed qed lemma path_connected_continuous_image: assumes "continuous_on S f" and "path_connected S" shows "path_connected (f ` S)" unfolding path_connected_def proof (rule, rule) fix x' y' assume "x' \ f ` S" "y' \ f ` S" then obtain x y where x: "x \ S" and y: "y \ S" and x': "x' = f x" and y': "y' = f y" by auto from x y obtain g where "path g \ path_image g \ S \ pathstart g = x \ pathfinish g = y" using assms(2)[unfolded path_connected_def] by fast then show "\g. path g \ path_image g \ f ` S \ pathstart g = x' \ pathfinish g = y'" unfolding x' y' apply (rule_tac x="f \ g" in exI) unfolding path_defs apply (intro conjI continuous_on_compose continuous_on_subset[OF assms(1)]) apply auto done qed lemma path_connected_translationI: fixes a :: "'a :: topological_group_add" assumes "path_connected S" shows "path_connected ((\x. a + x) ` S)" by (intro path_connected_continuous_image assms continuous_intros) lemma path_connected_translation: fixes a :: "'a :: topological_group_add" shows "path_connected ((\x. a + x) ` S) = path_connected S" proof - have "\x y. (+) (x::'a) ` (+) (0 - x) ` y = y" by (simp add: image_image) then show ?thesis by (metis (no_types) path_connected_translationI) qed lemma path_connected_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (closed_segment a b)" by (simp add: convex_imp_path_connected) lemma path_connected_open_segment [simp]: fixes a :: "'a::real_normed_vector" shows "path_connected (open_segment a b)" by (simp add: convex_imp_path_connected) lemma homeomorphic_path_connectedness: "S homeomorphic T \ path_connected S \ path_connected T" unfolding homeomorphic_def homeomorphism_def by (metis path_connected_continuous_image) lemma path_connected_empty [simp]: "path_connected {}" unfolding path_connected_def by auto lemma path_connected_singleton [simp]: "path_connected {a}" unfolding path_connected_def pathstart_def pathfinish_def path_image_def apply clarify apply (rule_tac x="\x. a" in exI) apply (simp add: image_constant_conv) apply (simp add: path_def) done lemma path_connected_Un: assumes "path_connected S" and "path_connected T" and "S \ T \ {}" shows "path_connected (S \ T)" unfolding path_connected_component proof (intro ballI) fix x y assume x: "x \ S \ T" and y: "y \ S \ T" from assms obtain z where z: "z \ S" "z \ T" by auto show "path_component (S \ T) x y" using x y proof safe assume "x \ S" "y \ S" then show "path_component (S \ T) x y" by (meson Un_upper1 \path_connected S\ path_component_of_subset path_connected_component) next assume "x \ S" "y \ T" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ S" then show "path_component (S \ T) x y" by (metis z assms(1-2) le_sup_iff order_refl path_component_of_subset path_component_trans path_connected_component) next assume "x \ T" "y \ T" then show "path_component (S \ T) x y" by (metis Un_upper1 assms(2) path_component_of_subset path_connected_component sup_commute) qed qed lemma path_connected_UNION: assumes "\i. i \ A \ path_connected (S i)" and "\i. i \ A \ z \ S i" shows "path_connected (\i\A. S i)" unfolding path_connected_component proof clarify fix x i y j assume *: "i \ A" "x \ S i" "j \ A" "y \ S j" then have "path_component (S i) x z" and "path_component (S j) z y" using assms by (simp_all add: path_connected_component) then have "path_component (\i\A. S i) x z" and "path_component (\i\A. S i) z y" using *(1,3) by (auto elim!: path_component_of_subset [rotated]) then show "path_component (\i\A. S i) x y" by (rule path_component_trans) qed lemma path_component_path_image_pathstart: assumes p: "path p" and x: "x \ path_image p" shows "path_component (path_image p) (pathstart p) x" proof - obtain y where x: "x = p y" and y: "0 \ y" "y \ 1" using x by (auto simp: path_image_def) show ?thesis unfolding path_component_def proof (intro exI conjI) have "continuous_on {0..1} (p \ ((*) y))" apply (rule continuous_intros)+ using p [unfolded path_def] y apply (auto simp: mult_le_one intro: continuous_on_subset [of _ p]) done then show "path (\u. p (y * u))" by (simp add: path_def) show "path_image (\u. p (y * u)) \ path_image p" using y mult_le_one by (fastforce simp: path_image_def image_iff) qed (auto simp: pathstart_def pathfinish_def x) qed lemma path_connected_path_image: "path p \ path_connected(path_image p)" unfolding path_connected_component by (meson path_component_path_image_pathstart path_component_sym path_component_trans) lemma path_connected_path_component [simp]: "path_connected (path_component_set s x)" proof - { fix y z assume pa: "path_component s x y" "path_component s x z" then have pae: "path_component_set s x = path_component_set s y" using path_component_eq by auto have yz: "path_component s y z" using pa path_component_sym path_component_trans by blast then have "\g. path g \ path_image g \ path_component_set s x \ pathstart g = y \ pathfinish g = z" apply (simp add: path_component_def, clarify) apply (rule_tac x=g in exI) by (simp add: pae path_component_maximal path_connected_path_image pathstart_in_path_image) } then show ?thesis by (simp add: path_connected_def) qed lemma path_component: "path_component S x y \ (\t. path_connected t \ t \ S \ x \ t \ y \ t)" apply (intro iffI) apply (metis path_connected_path_image path_defs(5) pathfinish_in_path_image pathstart_in_path_image) using path_component_of_subset path_connected_component by blast lemma path_component_path_component [simp]: "path_component_set (path_component_set S x) x = path_component_set S x" proof (cases "x \ S") case True show ?thesis apply (rule subset_antisym) apply (simp add: path_component_subset) by (simp add: True path_component_maximal path_component_refl) next case False then show ?thesis by (metis False empty_iff path_component_eq_empty) qed lemma path_component_subset_connected_component: "(path_component_set S x) \ (connected_component_set S x)" proof (cases "x \ S") case True show ?thesis apply (rule connected_component_maximal) apply (auto simp: True path_component_subset path_component_refl path_connected_imp_connected) done next case False then show ?thesis using path_component_eq_empty by auto qed subsection\<^marker>\tag unimportant\\Lemmas about path-connectedness\ lemma path_connected_linear_image: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" assumes "path_connected S" "bounded_linear f" shows "path_connected(f ` S)" by (auto simp: linear_continuous_on assms path_connected_continuous_image) lemma is_interval_path_connected: "is_interval S \ path_connected S" by (simp add: convex_imp_path_connected is_interval_convex) lemma path_connected_Ioi[simp]: "path_connected {a<..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Ici[simp]: "path_connected {a..}" for a :: real by (simp add: convex_imp_path_connected) lemma path_connected_Iio[simp]: "path_connected {.. (\x \ topspace X. \y \ topspace X. \S. path_connectedin X S \ x \ S \ y \ S)" unfolding path_connected_space_def Ball_def apply (intro all_cong1 imp_cong refl, safe) using path_connectedin_path_image apply fastforce by (meson path_connectedin) lemma connectedin_path_image: "pathin X g \ connectedin X (g ` ({0..1}))" by (simp add: path_connectedin_imp_connectedin path_connectedin_path_image) lemma compactin_path_image: "pathin X g \ compactin X (g ` ({0..1}))" unfolding pathin_def by (rule image_compactin [of "top_of_set {0..1}"]) auto lemma linear_homeomorphism_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" obtains g where "homeomorphism (f ` S) S g f" using linear_injective_left_inverse [OF assms] apply clarify apply (rule_tac g=g in that) using assms apply (auto simp: homeomorphism_def eq_id_iff [symmetric] image_comp comp_def linear_conv_bounded_linear linear_continuous_on) done lemma linear_homeomorphic_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "S homeomorphic f ` S" by (meson homeomorphic_def homeomorphic_sym linear_homeomorphism_image [OF assms]) lemma path_connected_Times: assumes "path_connected s" "path_connected t" shows "path_connected (s \ t)" proof (simp add: path_connected_def Sigma_def, clarify) fix x1 y1 x2 y2 assume "x1 \ s" "y1 \ t" "x2 \ s" "y2 \ t" obtain g where "path g" and g: "path_image g \ s" and gs: "pathstart g = x1" and gf: "pathfinish g = x2" using \x1 \ s\ \x2 \ s\ assms by (force simp: path_connected_def) obtain h where "path h" and h: "path_image h \ t" and hs: "pathstart h = y1" and hf: "pathfinish h = y2" using \y1 \ t\ \y2 \ t\ assms by (force simp: path_connected_def) have "path (\z. (x1, h z))" using \path h\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = h]) apply (rule continuous_intros | force)+ done moreover have "path (\z. (g z, y2))" using \path g\ apply (simp add: path_def) apply (rule continuous_on_compose2 [where f = g]) apply (rule continuous_intros | force)+ done ultimately have 1: "path ((\z. (x1, h z)) +++ (\z. (g z, y2)))" by (metis hf gs path_join_imp pathstart_def pathfinish_def) have "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ path_image (\z. (x1, h z)) \ path_image (\z. (g z, y2))" by (rule Path_Connected.path_image_join_subset) also have "\ \ (\x\s. \x1\t. {(x, x1)})" using g h \x1 \ s\ \y2 \ t\ by (force simp: path_image_def) finally have 2: "path_image ((\z. (x1, h z)) +++ (\z. (g z, y2))) \ (\x\s. \x1\t. {(x, x1)})" . show "\g. path g \ path_image g \ (\x\s. \x1\t. {(x, x1)}) \ pathstart g = (x1, y1) \ pathfinish g = (x2, y2)" apply (intro exI conjI) apply (rule 1) apply (rule 2) apply (metis hs pathstart_def pathstart_join) by (metis gf pathfinish_def pathfinish_join) qed lemma is_interval_path_connected_1: fixes s :: "real set" shows "is_interval s \ path_connected s" using is_interval_connected_1 is_interval_path_connected path_connected_imp_connected by blast subsection\<^marker>\tag unimportant\\Path components\ lemma Union_path_component [simp]: "Union {path_component_set S x |x. x \ S} = S" apply (rule subset_antisym) using path_component_subset apply force using path_component_refl by auto lemma path_component_disjoint: "disjnt (path_component_set S a) (path_component_set S b) \ (a \ path_component_set S b)" apply (auto simp: disjnt_def) using path_component_eq apply fastforce using path_component_sym path_component_trans by blast lemma path_component_eq_eq: "path_component S x = path_component S y \ (x \ S) \ (y \ S) \ x \ S \ y \ S \ path_component S x y" apply (rule iffI, metis (no_types) path_component_mem(1) path_component_refl) apply (erule disjE, metis Collect_empty_eq_bot path_component_eq_empty) apply (rule ext) apply (metis path_component_trans path_component_sym) done lemma path_component_unique: assumes "x \ c" "c \ S" "path_connected c" "\c'. \x \ c'; c' \ S; path_connected c'\ \ c' \ c" shows "path_component_set S x = c" apply (rule subset_antisym) using assms apply (metis mem_Collect_eq subsetCE path_component_eq_eq path_component_subset path_connected_path_component) by (simp add: assms path_component_maximal) lemma path_component_intermediate_subset: "path_component_set u a \ t \ t \ u \ path_component_set t a = path_component_set u a" by (metis (no_types) path_component_mono path_component_path_component subset_antisym) lemma complement_path_component_Union: fixes x :: "'a :: topological_space" shows "S - path_component_set S x = \({path_component_set S y| y. y \ S} - {path_component_set S x})" proof - have *: "(\x. x \ S - {a} \ disjnt a x) \ \S - a = \(S - {a})" for a::"'a set" and S by (auto simp: disjnt_def) have "\y. y \ {path_component_set S x |x. x \ S} - {path_component_set S x} \ disjnt (path_component_set S x) y" using path_component_disjoint path_component_eq by fastforce then have "\{path_component_set S x |x. x \ S} - path_component_set S x = \({path_component_set S y |y. y \ S} - {path_component_set S x})" by (meson *) then show ?thesis by simp qed subsection\Path components\ definition path_component_of where "path_component_of X x y \ \g. pathin X g \ g 0 = x \ g 1 = y" abbreviation path_component_of_set where "path_component_of_set X x \ Collect (path_component_of X x)" definition path_components_of :: "'a topology \ 'a set set" where "path_components_of X \ path_component_of_set X ` topspace X" lemma pathin_canon_iff: "pathin (top_of_set T) g \ path g \ g ` {0..1} \ T" by (simp add: path_def pathin_def) lemma path_component_of_canon_iff [simp]: "path_component_of (top_of_set T) a b \ path_component T a b" by (simp add: path_component_of_def pathin_canon_iff path_defs) lemma path_component_in_topspace: "path_component_of X x y \ x \ topspace X \ y \ topspace X" by (auto simp: path_component_of_def pathin_def continuous_map_def) lemma path_component_of_refl: "path_component_of X x x \ x \ topspace X" apply (auto simp: path_component_in_topspace) apply (force simp: path_component_of_def pathin_const) done lemma path_component_of_sym: assumes "path_component_of X x y" shows "path_component_of X y x" using assms apply (clarsimp simp: path_component_of_def pathin_def) apply (rule_tac x="g \ (\t. 1 - t)" in exI) apply (auto intro!: continuous_map_compose) apply (force simp: continuous_map_in_subtopology continuous_on_op_minus) done lemma path_component_of_sym_iff: "path_component_of X x y \ path_component_of X y x" by (metis path_component_of_sym) lemma path_component_of_trans: assumes "path_component_of X x y" and "path_component_of X y z" shows "path_component_of X x z" unfolding path_component_of_def pathin_def proof - let ?T01 = "top_of_set {0..1::real}" obtain g1 g2 where g1: "continuous_map ?T01 X g1" "x = g1 0" "y = g1 1" and g2: "continuous_map ?T01 X g2" "g2 0 = g1 1" "z = g2 1" using assms unfolding path_component_of_def pathin_def by blast let ?g = "\x. if x \ 1/2 then (g1 \ (\t. 2 * t)) x else (g2 \ (\t. 2 * t -1)) x" show "\g. continuous_map ?T01 X g \ g 0 = x \ g 1 = z" proof (intro exI conjI) show "continuous_map (subtopology euclideanreal {0..1}) X ?g" proof (intro continuous_map_cases_le continuous_map_compose, force, force) show "continuous_map (subtopology ?T01 {x \ topspace ?T01. x \ 1/2}) ?T01 ((*) 2)" by (auto simp: continuous_map_in_subtopology continuous_map_from_subtopology) have "continuous_map (subtopology (top_of_set {0..1}) {x. 0 \ x \ x \ 1 \ 1 \ x * 2}) euclideanreal (\t. 2 * t - 1)" by (intro continuous_intros) (force intro: continuous_map_from_subtopology) then show "continuous_map (subtopology ?T01 {x \ topspace ?T01. 1/2 \ x}) ?T01 (\t. 2 * t - 1)" by (force simp: continuous_map_in_subtopology) show "(g1 \ (*) 2) x = (g2 \ (\t. 2 * t - 1)) x" if "x \ topspace ?T01" "x = 1/2" for x using that by (simp add: g2(2) mult.commute continuous_map_from_subtopology) qed (auto simp: g1 g2) qed (auto simp: g1 g2) qed lemma path_component_of_mono: "\path_component_of (subtopology X S) x y; S \ T\ \ path_component_of (subtopology X T) x y" unfolding path_component_of_def by (metis subsetD pathin_subtopology) lemma path_component_of: "path_component_of X x y \ (\T. path_connectedin X T \ x \ T \ y \ T)" apply (auto simp: path_component_of_def) using path_connectedin_path_image apply fastforce apply (metis path_connectedin) done lemma path_component_of_set: "path_component_of X x y \ (\g. pathin X g \ g 0 = x \ g 1 = y)" by (auto simp: path_component_of_def) lemma path_component_of_subset_topspace: "Collect(path_component_of X x) \ topspace X" using path_component_in_topspace by fastforce lemma path_component_of_eq_empty: "Collect(path_component_of X x) = {} \ (x \ topspace X)" using path_component_in_topspace path_component_of_refl by fastforce lemma path_connected_space_iff_path_component: "path_connected_space X \ (\x \ topspace X. \y \ topspace X. path_component_of X x y)" by (simp add: path_component_of path_connected_space_subconnected) lemma path_connected_space_imp_path_component_of: "\path_connected_space X; a \ topspace X; b \ topspace X\ \ path_component_of X a b" by (simp add: path_connected_space_iff_path_component) lemma path_connected_space_path_component_set: "path_connected_space X \ (\x \ topspace X. Collect(path_component_of X x) = topspace X)" using path_component_of_subset_topspace path_connected_space_iff_path_component by fastforce lemma path_component_of_maximal: "\path_connectedin X s; x \ s\ \ s \ Collect(path_component_of X x)" using path_component_of by fastforce lemma path_component_of_equiv: "path_component_of X x y \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (simp add: fun_eq_iff path_component_in_topspace) apply (meson path_component_of_sym path_component_of_trans) done qed (simp add: path_component_of_refl) lemma path_component_of_disjoint: "disjnt (Collect (path_component_of X x)) (Collect (path_component_of X y)) \ ~(path_component_of X x y)" by (force simp: disjnt_def path_component_of_eq_empty path_component_of_equiv) lemma path_component_of_eq: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ x \ topspace X \ y \ topspace X \ path_component_of X x y" by (metis Collect_empty_eq_bot path_component_of_eq_empty path_component_of_equiv) lemma path_connectedin_path_component_of: "path_connectedin X (Collect (path_component_of X x))" proof - have "\y. path_component_of X x y \ path_component_of (subtopology X (Collect (path_component_of X x))) x y" by (meson path_component_of path_component_of_maximal path_connectedin_subtopology) then show ?thesis apply (simp add: path_connectedin_def path_component_of_subset_topspace path_connected_space_iff_path_component) by (metis Int_absorb1 mem_Collect_eq path_component_of_equiv path_component_of_subset_topspace topspace_subtopology) qed lemma path_connectedin_euclidean [simp]: "path_connectedin euclidean S \ path_connected S" by (auto simp: path_connectedin_def path_connected_space_iff_path_component path_connected_component) lemma path_connected_space_euclidean_subtopology [simp]: "path_connected_space(subtopology euclidean S) \ path_connected S" using path_connectedin_topspace by force lemma Union_path_components_of: "\(path_components_of X) = topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_components_of_maximal: "\C \ path_components_of X; path_connectedin X S; ~disjnt C S\ \ S \ C" apply (auto simp: path_components_of_def path_component_of_equiv) using path_component_of_maximal path_connectedin_def apply fastforce by (meson disjnt_subset2 path_component_of_disjoint path_component_of_equiv path_component_of_maximal) lemma pairwise_disjoint_path_components_of: "pairwise disjnt (path_components_of X)" by (auto simp: path_components_of_def pairwise_def path_component_of_disjoint path_component_of_equiv) lemma complement_path_components_of_Union: "C \ path_components_of X \ topspace X - C = \(path_components_of X - {C})" by (metis Diff_cancel Diff_subset Union_path_components_of cSup_singleton diff_Union_pairwise_disjoint insert_subset pairwise_disjoint_path_components_of) lemma nonempty_path_components_of: "C \ path_components_of X \ (C \ {})" apply (clarsimp simp: path_components_of_def path_component_of_eq_empty) by (meson path_component_of_refl) lemma path_components_of_subset: "C \ path_components_of X \ C \ topspace X" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_connectedin_path_components_of: "C \ path_components_of X \ path_connectedin X C" by (auto simp: path_components_of_def path_connectedin_path_component_of) lemma path_component_in_path_components_of: "Collect (path_component_of X a) \ path_components_of X \ a \ topspace X" apply (rule iffI) using nonempty_path_components_of path_component_of_eq_empty apply fastforce by (simp add: path_components_of_def) lemma path_connectedin_Union: assumes \: "\S. S \ \ \ path_connectedin X S" "\\ \ {}" shows "path_connectedin X (\\)" proof - obtain a where "\S. S \ \ \ a \ S" using assms by blast then have "\x. x \ topspace (subtopology X (\\)) \ path_component_of (subtopology X (\\)) a x" apply (simp) by (meson Union_upper \ path_component_of path_connectedin_subtopology) then show ?thesis using \ unfolding path_connectedin_def by (metis Sup_le_iff path_component_of_equiv path_connected_space_iff_path_component) qed lemma path_connectedin_Un: "\path_connectedin X S; path_connectedin X T; S \ T \ {}\ \ path_connectedin X (S \ T)" by (blast intro: path_connectedin_Union [of "{S,T}", simplified]) lemma path_connected_space_iff_components_eq: "path_connected_space X \ (\C \ path_components_of X. \C' \ path_components_of X. C = C')" unfolding path_components_of_def proof (intro iffI ballI) assume "\C \ path_component_of_set X ` topspace X. \C' \ path_component_of_set X ` topspace X. C = C'" then show "path_connected_space X" using path_component_of_refl path_connected_space_iff_path_component by fastforce qed (auto simp: path_connected_space_path_component_set) lemma path_components_of_eq_empty: "path_components_of X = {} \ topspace X = {}" using Union_path_components_of nonempty_path_components_of by fastforce lemma path_components_of_empty_space: "topspace X = {} \ path_components_of X = {}" by (simp add: path_components_of_eq_empty) lemma path_components_of_subset_singleton: "path_components_of X \ {S} \ path_connected_space X \ (topspace X = {} \ topspace X = S)" proof (cases "topspace X = {}") case True then show ?thesis by (auto simp: path_components_of_empty_space path_connected_space_topspace_empty) next case False have "(path_components_of X = {S}) \ (path_connected_space X \ topspace X = S)" proof (intro iffI conjI) assume L: "path_components_of X = {S}" then show "path_connected_space X" by (simp add: path_connected_space_iff_components_eq) show "topspace X = S" by (metis L ccpo_Sup_singleton [of S] Union_path_components_of) next assume R: "path_connected_space X \ topspace X = S" then show "path_components_of X = {S}" using ccpo_Sup_singleton [of S] by (metis False all_not_in_conv insert_iff mk_disjoint_insert path_component_in_path_components_of path_connected_space_iff_components_eq path_connected_space_path_component_set) qed with False show ?thesis by (simp add: path_components_of_eq_empty subset_singleton_iff) qed lemma path_connected_space_iff_components_subset_singleton: "path_connected_space X \ (\a. path_components_of X \ {a})" by (simp add: path_components_of_subset_singleton) lemma path_components_of_eq_singleton: "path_components_of X = {S} \ path_connected_space X \ topspace X \ {} \ S = topspace X" by (metis cSup_singleton insert_not_empty path_components_of_subset_singleton subset_singleton_iff) lemma path_components_of_path_connected_space: "path_connected_space X \ path_components_of X = (if topspace X = {} then {} else {topspace X})" by (simp add: path_components_of_eq_empty path_components_of_eq_singleton) lemma path_component_subset_connected_component_of: "path_component_of_set X x \ connected_component_of_set X x" proof (cases "x \ topspace X") case True then show ?thesis by (simp add: connected_component_of_maximal path_component_of_refl path_connectedin_imp_connectedin path_connectedin_path_component_of) next case False then show ?thesis using path_component_of_eq_empty by fastforce qed lemma exists_path_component_of_superset: assumes S: "path_connectedin X S" and ne: "topspace X \ {}" obtains C where "C \ path_components_of X" "S \ C" proof (cases "S = {}") case True then show ?thesis using ne path_components_of_eq_empty that by fastforce next case False then obtain a where "a \ S" by blast show ?thesis proof show "Collect (path_component_of X a) \ path_components_of X" by (meson \a \ S\ S subsetD path_component_in_path_components_of path_connectedin_subset_topspace) show "S \ Collect (path_component_of X a)" by (simp add: S \a \ S\ path_component_of_maximal) qed qed lemma path_component_of_eq_overlap: "path_component_of X x = path_component_of X y \ (x \ topspace X) \ (y \ topspace X) \ Collect (path_component_of X x) \ Collect (path_component_of X y) \ {}" by (metis disjnt_def empty_iff inf_bot_right mem_Collect_eq path_component_of_disjoint path_component_of_eq path_component_of_eq_empty) lemma path_component_of_nonoverlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) = {} \ (x \ topspace X) \ (y \ topspace X) \ path_component_of X x \ path_component_of X y" by (metis inf.idem path_component_of_eq_empty path_component_of_eq_overlap) lemma path_component_of_overlap: "Collect (path_component_of X x) \ Collect (path_component_of X y) \ {} \ x \ topspace X \ y \ topspace X \ path_component_of X x = path_component_of X y" by (meson path_component_of_nonoverlap) lemma path_components_of_disjoint: "\C \ path_components_of X; C' \ path_components_of X\ \ disjnt C C' \ C \ C'" by (auto simp: path_components_of_def path_component_of_disjoint path_component_of_equiv) lemma path_components_of_overlap: "\C \ path_components_of X; C' \ path_components_of X\ \ C \ C' \ {} \ C = C'" by (auto simp: path_components_of_def path_component_of_equiv) lemma path_component_of_unique: "\x \ C; path_connectedin X C; \C'. \x \ C'; path_connectedin X C'\ \ C' \ C\ \ Collect (path_component_of X x) = C" by (meson subsetD eq_iff path_component_of_maximal path_connectedin_path_component_of) lemma path_component_of_discrete_topology [simp]: "Collect (path_component_of (discrete_topology U) x) = (if x \ U then {x} else {})" proof - have "\C'. \x \ C'; path_connectedin (discrete_topology U) C'\ \ C' \ {x}" by (metis path_connectedin_discrete_topology subsetD singletonD) then have "x \ U \ Collect (path_component_of (discrete_topology U) x) = {x}" by (simp add: path_component_of_unique) then show ?thesis using path_component_in_topspace by fastforce qed lemma path_component_of_discrete_topology_iff [simp]: "path_component_of (discrete_topology U) x y \ x \ U \ y=x" by (metis empty_iff insertI1 mem_Collect_eq path_component_of_discrete_topology singletonD) lemma path_components_of_discrete_topology [simp]: "path_components_of (discrete_topology U) = (\x. {x}) ` U" by (auto simp: path_components_of_def image_def fun_eq_iff) lemma homeomorphic_map_path_component_of: assumes f: "homeomorphic_map X Y f" and x: "x \ topspace X" shows "Collect (path_component_of Y (f x)) = f ` Collect(path_component_of X x)" proof - obtain g where g: "homeomorphic_maps X Y f g" using f homeomorphic_map_maps by blast show ?thesis proof have "Collect (path_component_of Y (f x)) \ topspace Y" by (simp add: path_component_of_subset_topspace) moreover have "g ` Collect(path_component_of Y (f x)) \ Collect (path_component_of X (g (f x)))" using g x unfolding homeomorphic_maps_def by (metis f homeomorphic_imp_surjective_map imageI mem_Collect_eq path_component_of_maximal path_component_of_refl path_connectedin_continuous_map_image path_connectedin_path_component_of) ultimately show "Collect (path_component_of Y (f x)) \ f ` Collect (path_component_of X x)" using g x unfolding homeomorphic_maps_def continuous_map_def image_iff subset_iff by metis show "f ` Collect (path_component_of X x) \ Collect (path_component_of Y (f x))" proof (rule path_component_of_maximal) show "path_connectedin Y (f ` Collect (path_component_of X x))" by (meson f homeomorphic_map_path_connectedness_eq path_connectedin_path_component_of) qed (simp add: path_component_of_refl x) qed qed lemma homeomorphic_map_path_components_of: assumes "homeomorphic_map X Y f" shows "path_components_of Y = (image f) ` (path_components_of X)" unfolding path_components_of_def homeomorphic_imp_surjective_map [OF assms, symmetric] apply safe using assms homeomorphic_map_path_component_of apply fastforce by (metis assms homeomorphic_map_path_component_of imageI) subsection \Sphere is path-connected\ lemma path_connected_punctured_universe: assumes "2 \ DIM('a::euclidean_space)" shows "path_connected (- {a::'a})" proof - let ?A = "{x::'a. \i\Basis. x \ i < a \ i}" let ?B = "{x::'a. \i\Basis. a \ i < x \ i}" have A: "path_connected ?A" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i - 1)*\<^sub>R i) \ {x::'a. x \ i < a \ i}" by simp show "path_connected {x. x \ i < a \ i}" using convex_imp_path_connected [OF convex_halfspace_lt, of i "a \ i"] by (simp add: inner_commute) qed have B: "path_connected ?B" unfolding Collect_bex_eq proof (rule path_connected_UNION) fix i :: 'a assume "i \ Basis" then show "(\i\Basis. (a \ i + 1) *\<^sub>R i) \ {x::'a. a \ i < x \ i}" by simp show "path_connected {x. a \ i < x \ i}" using convex_imp_path_connected [OF convex_halfspace_gt, of "a \ i" i] by (simp add: inner_commute) qed obtain S :: "'a set" where "S \ Basis" and "card S = Suc (Suc 0)" using ex_card[OF assms] by auto then obtain b0 b1 :: 'a where "b0 \ Basis" and "b1 \ Basis" and "b0 \ b1" unfolding card_Suc_eq by auto then have "a + b0 - b1 \ ?A \ ?B" by (auto simp: inner_simps inner_Basis) then have "?A \ ?B \ {}" by fast with A B have "path_connected (?A \ ?B)" by (rule path_connected_Un) also have "?A \ ?B = {x. \i\Basis. x \ i \ a \ i}" unfolding neq_iff bex_disj_distrib Collect_disj_eq .. also have "\ = {x. x \ a}" unfolding euclidean_eq_iff [where 'a='a] by (simp add: Bex_def) also have "\ = - {a}" by auto finally show ?thesis . qed corollary connected_punctured_universe: "2 \ DIM('N::euclidean_space) \ connected(- {a::'N})" by (simp add: path_connected_punctured_universe path_connected_imp_connected) proposition path_connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "path_connected(sphere a r)" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (simp) next case equal then show ?thesis by (simp) next case greater then have eq: "(sphere (0::'a) r) = (\x. (r / norm x) *\<^sub>R x) ` (- {0::'a})" by (force simp: image_iff split: if_split_asm) have "continuous_on (- {0::'a}) (\x. (r / norm x) *\<^sub>R x)" by (intro continuous_intros) auto then have "path_connected ((\x. (r / norm x) *\<^sub>R x) ` (- {0::'a}))" by (intro path_connected_continuous_image path_connected_punctured_universe assms) with eq have "path_connected (sphere (0::'a) r)" by auto then have "path_connected((+) a ` (sphere (0::'a) r))" by (simp add: path_connected_translation) then show ?thesis by (metis add.right_neutral sphere_translation) qed lemma connected_sphere: fixes a :: "'a :: euclidean_space" assumes "2 \ DIM('a)" shows "connected(sphere a r)" using path_connected_sphere [OF assms] by (simp add: path_connected_imp_connected) corollary path_connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" and 2: "2 \ DIM('a)" shows "path_connected (- s)" proof (cases "s = {}") case True then show ?thesis using convex_imp_path_connected by auto next case False then obtain a where "a \ s" by auto { fix x y assume "x \ s" "y \ s" then have "x \ a" "y \ a" using \a \ s\ by auto then have bxy: "bounded(insert x (insert y s))" by (simp add: \bounded s\) then obtain B::real where B: "0 < B" and Bx: "norm (a - x) < B" and By: "norm (a - y) < B" and "s \ ball a B" using bounded_subset_ballD [OF bxy, of a] by (auto simp: dist_norm) define C where "C = B / norm(x - a)" { fix u assume u: "(1 - u) *\<^sub>R x + u *\<^sub>R (a + C *\<^sub>R (x - a)) \ s" and "0 \ u" "u \ 1" have CC: "1 \ 1 + (C - 1) * u" using \x \ a\ \0 \ u\ Bx by (auto simp add: C_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R x + u *\<^sub>R (a + v *\<^sub>R (x - a)) = a + (1 + (v - 1) * u) *\<^sub>R (x - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + C * u - u)) *\<^sub>R x + ((u / (1 + C * u - u)) *\<^sub>R a + (C * u / (1 + C * u - u)) *\<^sub>R x)) = (1 + (u / (1 + C * u - u))) *\<^sub>R a + ((1 / (1 + C * u - u)) + (C * u / (1 + C * u - u))) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + C * u - u))) *\<^sub>R a + (1 + (u / (1 + C * u - u))) *\<^sub>R x" using CC by (simp add: field_simps) also have "\ = x + (1 + (u / (1 + C * u - u))) *\<^sub>R a + (u / (1 + C * u - u)) *\<^sub>R x" by (simp add: algebra_simps) also have "\ = x + ((1 / (1 + C * u - u)) *\<^sub>R a + ((u / (1 + C * u - u)) *\<^sub>R x + (C * u / (1 + C * u - u)) *\<^sub>R a))" using CC by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (C - 1) * u)) *\<^sub>R a + (1 / (1 + (C - 1) * u)) *\<^sub>R (a + (1 + (C - 1) * u) *\<^sub>R (x - a)) = x" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (C - 1) * u) *\<^sub>R (x - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (C - 1) * u)" in spec) using \x \ a\ \x \ s\ \0 \ u\ CC apply (auto simp: xeq) done } then have pcx: "path_component (- s) x (a + C *\<^sub>R (x - a))" by (force simp: closed_segment_def intro!: path_component_linepath) define D where "D = B / norm(y - a)" \ \massive duplication with the proof above\ { fix u assume u: "(1 - u) *\<^sub>R y + u *\<^sub>R (a + D *\<^sub>R (y - a)) \ s" and "0 \ u" "u \ 1" have DD: "1 \ 1 + (D - 1) * u" using \y \ a\ \0 \ u\ By by (auto simp add: D_def norm_minus_commute) have *: "\v. (1 - u) *\<^sub>R y + u *\<^sub>R (a + v *\<^sub>R (y - a)) = a + (1 + (v - 1) * u) *\<^sub>R (y - a)" by (simp add: algebra_simps) have "a + ((1 / (1 + D * u - u)) *\<^sub>R y + ((u / (1 + D * u - u)) *\<^sub>R a + (D * u / (1 + D * u - u)) *\<^sub>R y)) = (1 + (u / (1 + D * u - u))) *\<^sub>R a + ((1 / (1 + D * u - u)) + (D * u / (1 + D * u - u))) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = (1 + (u / (1 + D * u - u))) *\<^sub>R a + (1 + (u / (1 + D * u - u))) *\<^sub>R y" using DD by (simp add: field_simps) also have "\ = y + (1 + (u / (1 + D * u - u))) *\<^sub>R a + (u / (1 + D * u - u)) *\<^sub>R y" by (simp add: algebra_simps) also have "\ = y + ((1 / (1 + D * u - u)) *\<^sub>R a + ((u / (1 + D * u - u)) *\<^sub>R y + (D * u / (1 + D * u - u)) *\<^sub>R a))" using DD by (simp add: field_simps) (simp add: add_divide_distrib scaleR_add_left) finally have xeq: "(1 - 1 / (1 + (D - 1) * u)) *\<^sub>R a + (1 / (1 + (D - 1) * u)) *\<^sub>R (a + (1 + (D - 1) * u) *\<^sub>R (y - a)) = y" by (simp add: algebra_simps) have False using \convex s\ apply (simp add: convex_alt) apply (drule_tac x=a in bspec) apply (rule \a \ s\) apply (drule_tac x="a + (1 + (D - 1) * u) *\<^sub>R (y - a)" in bspec) using u apply (simp add: *) apply (drule_tac x="1 / (1 + (D - 1) * u)" in spec) using \y \ a\ \y \ s\ \0 \ u\ DD apply (auto simp: xeq) done } then have pdy: "path_component (- s) y (a + D *\<^sub>R (y - a))" by (force simp: closed_segment_def intro!: path_component_linepath) have pyx: "path_component (- s) (a + D *\<^sub>R (y - a)) (a + C *\<^sub>R (x - a))" apply (rule path_component_of_subset [of "sphere a B"]) using \s \ ball a B\ apply (force simp: ball_def dist_norm norm_minus_commute) apply (rule path_connected_sphere [OF 2, of a B, simplified path_connected_component, rule_format]) using \x \ a\ using \y \ a\ B apply (auto simp: dist_norm C_def D_def) done have "path_component (- s) x y" by (metis path_component_trans path_component_sym pcx pdy pyx) } then show ?thesis by (auto simp: path_connected_component) qed lemma connected_complement_bounded_convex: fixes s :: "'a :: euclidean_space set" assumes "bounded s" "convex s" "2 \ DIM('a)" shows "connected (- s)" using path_connected_complement_bounded_convex [OF assms] path_connected_imp_connected by blast lemma connected_diff_ball: fixes s :: "'a :: euclidean_space set" assumes "connected s" "cball a r \ s" "2 \ DIM('a)" shows "connected (s - ball a r)" apply (rule connected_diff_open_from_closed [OF ball_subset_cball]) using assms connected_sphere apply (auto simp: cball_diff_eq_sphere dist_norm) done proposition connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "connected(S - {a::'N})" proof (cases "a \ S") case True with \open S\ obtain \ where "\ > 0" and \: "cball a \ \ S" using open_contains_cball_eq by blast have "dist a (a + \ *\<^sub>R (SOME i. i \ Basis)) = \" by (simp add: dist_norm SOME_Basis \0 < \\ less_imp_le) with \ have "\{S - ball a r |r. 0 < r \ r < \} \ {} \ False" apply (drule_tac c="a + scaleR (\) ((SOME i. i \ Basis))" in subsetD) by auto then have nonemp: "(\{S - ball a r |r. 0 < r \ r < \}) = {} \ False" by auto have con: "\r. r < \ \ connected (S - ball a r)" using \ by (force intro: connected_diff_ball [OF \connected S\ _ 2]) have "x \ \{S - ball a r |r. 0 < r \ r < \}" if "x \ S - {a}" for x apply (rule UnionI [of "S - ball a (min \ (dist a x) / 2)"]) using that \0 < \\ apply simp_all apply (rule_tac x="min \ (dist a x) / 2" in exI) apply auto done then have "S - {a} = \{S - ball a r | r. 0 < r \ r < \}" by auto then show ?thesis by (auto intro: connected_Union con dest!: nonemp) next case False then show ?thesis by (simp add: \connected S\) qed corollary path_connected_open_delete: assumes "open S" "connected S" and 2: "2 \ DIM('N::euclidean_space)" shows "path_connected(S - {a::'N})" by (simp add: assms connected_open_delete connected_open_path_connected open_delete) corollary path_connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ path_connected(ball a r - {a::'N})" by (simp add: path_connected_open_delete) corollary connected_punctured_ball: "2 \ DIM('N::euclidean_space) \ connected(ball a r - {a::'N})" by (simp add: connected_open_delete) corollary connected_open_delete_finite: fixes S T::"'a::euclidean_space set" assumes S: "open S" "connected S" and 2: "2 \ DIM('a)" and "finite T" shows "connected(S - T)" using \finite T\ S proof (induct T) case empty show ?case using \connected S\ by simp next case (insert x F) then have "connected (S-F)" by auto moreover have "open (S - F)" using finite_imp_closed[OF \finite F\] \open S\ by auto ultimately have "connected (S - F - {x})" using connected_open_delete[OF _ _ 2] by auto thus ?case by (metis Diff_insert) qed lemma sphere_1D_doubleton_zero: assumes 1: "DIM('a) = 1" and "r > 0" obtains x y::"'a::euclidean_space" where "sphere 0 r = {x,y} \ dist x y = 2*r" proof - obtain b::'a where b: "Basis = {b}" using 1 card_1_singletonE by blast show ?thesis proof (intro that conjI) have "x = norm x *\<^sub>R b \ x = - norm x *\<^sub>R b" if "r = norm x" for x proof - have xb: "(x \ b) *\<^sub>R b = x" using euclidean_representation [of x, unfolded b] by force then have "norm ((x \ b) *\<^sub>R b) = norm x" by simp with b have "\x \ b\ = norm x" using norm_Basis by (simp add: b) with xb show ?thesis apply (simp add: abs_if split: if_split_asm) apply (metis add.inverse_inverse real_vector.scale_minus_left xb) done qed with \r > 0\ b show "sphere 0 r = {r *\<^sub>R b, - r *\<^sub>R b}" by (force simp: sphere_def dist_norm) have "dist (r *\<^sub>R b) (- r *\<^sub>R b) = norm (r *\<^sub>R b + r *\<^sub>R b)" by (simp add: dist_norm) also have "\ = norm ((2*r) *\<^sub>R b)" by (metis mult_2 scaleR_add_left) also have "\ = 2*r" using \r > 0\ b norm_Basis by fastforce finally show "dist (r *\<^sub>R b) (- r *\<^sub>R b) = 2*r" . qed qed lemma sphere_1D_doubleton: fixes a :: "'a :: euclidean_space" assumes "DIM('a) = 1" and "r > 0" obtains x y where "sphere a r = {x,y} \ dist x y = 2*r" proof - have "sphere a r = (+) a ` sphere 0 r" by (metis add.right_neutral sphere_translation) then show ?thesis using sphere_1D_doubleton_zero [OF assms] by (metis (mono_tags, lifting) dist_add_cancel image_empty image_insert that) qed lemma psubset_sphere_Compl_connected: fixes S :: "'a::euclidean_space set" assumes S: "S \ sphere a r" and "0 < r" and 2: "2 \ DIM('a)" shows "connected(- S)" proof - have "S \ sphere a r" using S by blast obtain b where "dist a b = r" and "b \ S" using S mem_sphere by blast have CS: "- S = {x. dist a x \ r \ (x \ S)} \ {x. r \ dist a x \ (x \ S)}" by auto have "{x. dist a x \ r \ x \ S} \ {x. r \ dist a x \ x \ S} \ {}" using \b \ S\ \dist a b = r\ by blast moreover have "connected {x. dist a x \ r \ x \ S}" apply (rule connected_intermediate_closure [of "ball a r"]) using assms by auto moreover have "connected {x. r \ dist a x \ x \ S}" apply (rule connected_intermediate_closure [of "- cball a r"]) using assms apply (auto intro: connected_complement_bounded_convex) apply (metis ComplI interior_cball interior_closure mem_ball not_less) done ultimately show ?thesis by (simp add: CS connected_Un) qed subsection\Every annulus is a connected set\ lemma path_connected_2DIM_I: fixes a :: "'N::euclidean_space" assumes 2: "2 \ DIM('N)" and pc: "path_connected {r. 0 \ r \ P r}" shows "path_connected {x. P(norm(x - a))}" proof - have "{x. P(norm(x - a))} = (+) a ` {x. P(norm x)}" by force moreover have "path_connected {x::'N. P(norm x)}" proof - let ?D = "{x. 0 \ x \ P x} \ sphere (0::'N) 1" have "x \ (\z. fst z *\<^sub>R snd z) ` ?D" if "P (norm x)" for x::'N proof (cases "x=0") case True with that show ?thesis apply (simp add: image_iff) apply (rule_tac x=0 in exI, simp) using vector_choose_size zero_le_one by blast next case False with that show ?thesis by (rule_tac x="(norm x, x /\<^sub>R norm x)" in image_eqI) auto qed then have *: "{x::'N. P(norm x)} = (\z. fst z *\<^sub>R snd z) ` ?D" by auto have "continuous_on ?D (\z:: real\'N. fst z *\<^sub>R snd z)" by (intro continuous_intros) moreover have "path_connected ?D" by (metis path_connected_Times [OF pc] path_connected_sphere 2) ultimately show ?thesis apply (subst *) apply (rule path_connected_continuous_image, auto) done qed ultimately show ?thesis using path_connected_translation by metis qed proposition path_connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N)" shows "path_connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "path_connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: is_interval_def intro!: is_interval_convex convex_imp_path_connected path_connected_2DIM_I [OF assms]) proposition connected_annulus: fixes a :: "'N::euclidean_space" assumes "2 \ DIM('N::euclidean_space)" shows "connected {x. r1 < norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 < norm(x - a) \ norm(x - a) \ r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) < r2}" "connected {x. r1 \ norm(x - a) \ norm(x - a) \ r2}" by (auto simp: path_connected_annulus [OF assms] path_connected_imp_connected) subsection\<^marker>\tag unimportant\\Relations between components and path components\ lemma open_connected_component: fixes s :: "'a::real_normed_vector set" shows "open s \ open (connected_component_set s x)" apply (simp add: open_contains_ball, clarify) apply (rename_tac y) apply (drule_tac x=y in bspec) apply (simp add: connected_component_in, clarify) apply (rule_tac x=e in exI) by (metis mem_Collect_eq connected_component_eq connected_component_maximal centre_in_ball connected_ball) corollary open_components: fixes s :: "'a::real_normed_vector set" shows "\open u; s \ components u\ \ open s" by (simp add: components_iff) (metis open_connected_component) lemma in_closure_connected_component: fixes s :: "'a::real_normed_vector set" assumes x: "x \ s" and s: "open s" shows "x \ closure (connected_component_set s y) \ x \ connected_component_set s y" proof - { assume "x \ closure (connected_component_set s y)" moreover have "x \ connected_component_set s x" using x by simp ultimately have "x \ connected_component_set s y" using s by (meson Compl_disjoint closure_iff_nhds_not_empty connected_component_disjoint disjoint_eq_subset_Compl open_connected_component) } then show ?thesis by (auto simp: closure_def) qed lemma connected_disjoint_Union_open_pick: assumes "pairwise disjnt B" "\S. S \ A \ connected S \ S \ {}" "\S. S \ B \ open S" "\A \ \B" "S \ A" obtains T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" proof - have "S \ \B" "connected S" "S \ {}" using assms \S \ A\ by blast+ then obtain T where "T \ B" "S \ T \ {}" by (metis Sup_inf_eq_bot_iff inf.absorb_iff2 inf_commute) have 1: "open T" by (simp add: \T \ B\ assms) have 2: "open (\(B-{T}))" using assms by blast have 3: "S \ T \ \(B - {T})" using \S \ \B\ by blast have "T \ \(B - {T}) = {}" using \T \ B\ \pairwise disjnt B\ by (auto simp: pairwise_def disjnt_def) then have 4: "T \ \(B - {T}) \ S = {}" by auto from connectedD [OF \connected S\ 1 2 3 4] have "S \ \(B-{T}) = {}" by (auto simp: Int_commute \S \ T \ {}\) with \T \ B\ have "S \ T" using "3" by auto show ?thesis using \S \ \(B - {T}) = {}\ \S \ T\ \T \ B\ that by auto qed lemma connected_disjoint_Union_open_subset: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A \ B" proof fix S assume "S \ A" obtain T where "T \ B" "S \ T" "S \ \(B - {T}) = {}" apply (rule connected_disjoint_Union_open_pick [OF B, of A]) using SA SB \S \ A\ by auto moreover obtain S' where "S' \ A" "T \ S'" "T \ \(A - {S'}) = {}" apply (rule connected_disjoint_Union_open_pick [OF A, of B]) using SA SB \T \ B\ by auto ultimately have "S' = S" by (metis A Int_subset_iff SA \S \ A\ disjnt_def inf.orderE pairwise_def) with \T \ S'\ have "T \ S" by simp with \S \ T\ have "S = T" by blast with \T \ B\ show "S \ B" by simp qed lemma connected_disjoint_Union_open_unique: assumes A: "pairwise disjnt A" and B: "pairwise disjnt B" and SA: "\S. S \ A \ open S \ connected S \ S \ {}" and SB: "\S. S \ B \ open S \ connected S \ S \ {}" and eq [simp]: "\A = \B" shows "A = B" by (rule subset_antisym; metis connected_disjoint_Union_open_subset assms) proposition components_open_unique: fixes S :: "'a::real_normed_vector set" assumes "pairwise disjnt A" "\A = S" "\X. X \ A \ open X \ connected X \ X \ {}" shows "components S = A" proof - have "open S" using assms by blast show ?thesis apply (rule connected_disjoint_Union_open_unique) apply (simp add: components_eq disjnt_def pairwise_def) using \open S\ apply (simp_all add: assms open_components in_components_connected in_components_nonempty) done qed subsection\<^marker>\tag unimportant\\Existence of unbounded components\ lemma cobounded_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes "bounded (-s)" shows "\x. x \ s \ \ bounded (connected_component_set s x)" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have unbounded_inner: "\ bounded {x. inner i x \ B}" apply (auto simp: bounded_def dist_norm) apply (rule_tac x="x + (max B e + 1 + \i \ x\) *\<^sub>R i" in exI) apply simp using i apply (auto simp: algebra_simps) done have **: "{x. B \ i \ x} \ connected_component_set s (B *\<^sub>R i)" apply (rule connected_component_maximal) apply (auto simp: i intro: convex_connected convex_halfspace_ge [of B]) apply (rule *) apply (rule order_trans [OF _ Basis_le_norm [OF i]]) by (simp add: inner_commute) have "B *\<^sub>R i \ s" by (rule *) (simp add: norm_Basis [OF i]) then show ?thesis apply (rule_tac x="B *\<^sub>R i" in exI, clarify) apply (frule bounded_subset [of _ "{x. B \ i \ x}", OF _ **]) using unbounded_inner apply blast done qed lemma cobounded_unique_unbounded_component: fixes s :: "'a :: euclidean_space set" assumes bs: "bounded (-s)" and "2 \ DIM('a)" and bo: "\ bounded(connected_component_set s x)" "\ bounded(connected_component_set s y)" shows "connected_component_set s x = connected_component_set s y" proof - obtain i::'a where i: "i \ Basis" using nonempty_Basis by blast obtain B where B: "B>0" "-s \ ball 0 B" using bounded_subset_ballD [OF bs, of 0] by auto then have *: "\x. B \ norm x \ x \ s" by (force simp: ball_def dist_norm) have ccb: "connected (- ball 0 B :: 'a set)" using assms by (auto intro: connected_complement_bounded_convex) obtain x' where x': "connected_component s x x'" "norm x' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) obtain y' where y': "connected_component s y y'" "norm y' > B" using bo [unfolded bounded_def dist_norm, simplified, rule_format] by (metis diff_zero norm_minus_commute not_less) have x'y': "connected_component s x' y'" apply (simp add: connected_component_def) apply (rule_tac x="- ball 0 B" in exI) using x' y' apply (auto simp: ccb dist_norm *) done show ?thesis apply (rule connected_component_eq) using x' y' x'y' by (metis (no_types, lifting) connected_component_eq_empty connected_component_eq_eq connected_component_idemp connected_component_in) qed lemma cobounded_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "bounded (-s) \ \c. c \ components s \ \bounded c" by (metis cobounded_unbounded_component components_def imageI) lemma cobounded_unique_unbounded_components: fixes s :: "'a :: euclidean_space set" shows "\bounded (- s); c \ components s; \ bounded c; c' \ components s; \ bounded c'; 2 \ DIM('a)\ \ c' = c" unfolding components_iff by (metis cobounded_unique_unbounded_component) lemma cobounded_has_bounded_component: fixes S :: "'a :: euclidean_space set" assumes "bounded (- S)" "\ connected S" "2 \ DIM('a)" obtains C where "C \ components S" "bounded C" by (meson cobounded_unique_unbounded_components connected_eq_connected_components_eq assms) subsection\The \inside\ and \outside\ of a Set\ text\<^marker>\tag important\\The inside comprises the points in a bounded connected component of the set's complement. The outside comprises the points in unbounded connected component of the complement.\ definition\<^marker>\tag important\ inside where "inside S \ {x. (x \ S) \ bounded(connected_component_set ( - S) x)}" definition\<^marker>\tag important\ outside where "outside S \ -S \ {x. \ bounded(connected_component_set (- S) x)}" lemma outside: "outside S = {x. \ bounded(connected_component_set (- S) x)}" by (auto simp: outside_def) (metis Compl_iff bounded_empty connected_component_eq_empty) lemma inside_no_overlap [simp]: "inside S \ S = {}" by (auto simp: inside_def) lemma outside_no_overlap [simp]: "outside S \ S = {}" by (auto simp: outside_def) lemma inside_Int_outside [simp]: "inside S \ outside S = {}" by (auto simp: inside_def outside_def) lemma inside_Un_outside [simp]: "inside S \ outside S = (- S)" by (auto simp: inside_def outside_def) lemma inside_eq_outside: "inside S = outside S \ S = UNIV" by (auto simp: inside_def outside_def) lemma inside_outside: "inside S = (- (S \ outside S))" by (force simp: inside_def outside) lemma outside_inside: "outside S = (- (S \ inside S))" by (auto simp: inside_outside) (metis IntI equals0D outside_no_overlap) lemma union_with_inside: "S \ inside S = - outside S" by (auto simp: inside_outside) (simp add: outside_inside) lemma union_with_outside: "S \ outside S = - inside S" by (simp add: inside_outside) lemma outside_mono: "S \ T \ outside T \ outside S" by (auto simp: outside bounded_subset connected_component_mono) lemma inside_mono: "S \ T \ inside S - T \ inside T" by (auto simp: inside_def bounded_subset connected_component_mono) lemma segment_bound_lemma: fixes u::real assumes "x \ B" "y \ B" "0 \ u" "u \ 1" shows "(1 - u) * x + u * y \ B" proof - obtain dx dy where "dx \ 0" "dy \ 0" "x = B + dx" "y = B + dy" using assms by auto (metis add.commute diff_add_cancel) with \0 \ u\ \u \ 1\ show ?thesis by (simp add: add_increasing2 mult_left_le field_simps) qed lemma cobounded_outside: fixes S :: "'a :: real_normed_vector set" assumes "bounded S" shows "bounded (- outside S)" proof - obtain B where B: "B>0" "S \ ball 0 B" using bounded_subset_ballD [OF assms, of 0] by auto { fix x::'a and C::real assume Bno: "B \ norm x" and C: "0 < C" have "\y. connected_component (- S) x y \ norm y > C" proof (cases "x = 0") case True with B Bno show ?thesis by force next case False have "closed_segment x (((B + C) / norm x) *\<^sub>R x) \ - ball 0 B" proof fix w assume "w \ closed_segment x (((B + C) / norm x) *\<^sub>R x)" then obtain u where w: "w = (1 - u + u * (B + C) / norm x) *\<^sub>R x" "0 \ u" "u \ 1" by (auto simp add: closed_segment_def real_vector_class.scaleR_add_left [symmetric]) with False B C have "B \ (1 - u) * norm x + u * (B + C)" using segment_bound_lemma [of B "norm x" "B + C" u] Bno by simp with False B C show "w \ - ball 0 B" using distrib_right [of _ _ "norm x"] by (simp add: ball_def w not_less) qed also have "... \ -S" by (simp add: B) finally have "\T. connected T \ T \ - S \ x \ T \ ((B + C) / norm x) *\<^sub>R x \ T" by (rule_tac x="closed_segment x (((B+C)/norm x) *\<^sub>R x)" in exI) simp with False B show ?thesis by (rule_tac x="((B+C)/norm x) *\<^sub>R x" in exI) (simp add: connected_component_def) qed } then show ?thesis apply (simp add: outside_def assms) apply (rule bounded_subset [OF bounded_ball [of 0 B]]) apply (force simp: dist_norm not_less bounded_pos) done qed lemma unbounded_outside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ \ bounded(outside S)" using cobounded_imp_unbounded cobounded_outside by blast lemma bounded_inside: fixes S :: "'a::{real_normed_vector, perfect_space} set" shows "bounded S \ bounded(inside S)" by (simp add: bounded_Int cobounded_outside inside_outside) lemma connected_outside: fixes S :: "'a::euclidean_space set" assumes "bounded S" "2 \ DIM('a)" shows "connected(outside S)" apply (clarsimp simp add: connected_iff_connected_component outside) apply (rule_tac s="connected_component_set (- S) x" in connected_component_of_subset) apply (metis (no_types) assms cobounded_unbounded_component cobounded_unique_unbounded_component connected_component_eq_eq connected_component_idemp double_complement mem_Collect_eq) apply clarify apply (metis connected_component_eq_eq connected_component_in) done lemma outside_connected_component_lt: "outside S = {x. \B. \y. B < norm(y) \ connected_component (- S) x y}" apply (auto simp: outside bounded_def dist_norm) apply (metis diff_0 norm_minus_cancel not_less) by (metis less_diff_eq norm_minus_commute norm_triangle_ineq2 order.trans pinf(6)) lemma outside_connected_component_le: "outside S = {x. \B. \y. B \ norm(y) \ connected_component (- S) x y}" apply (simp add: outside_connected_component_lt) apply (simp add: Set.set_eq_iff) by (meson gt_ex leD le_less_linear less_imp_le order.trans) lemma not_outside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" and "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B < norm(y) \ \ connected_component (- S) x y}" proof - obtain B::real where B: "0 < B" and Bno: "\x. x \ S \ norm x \ B" using S [simplified bounded_pos] by auto { fix y::'a and z::'a assume yz: "B < norm z" "B < norm y" have "connected_component (- cball 0 B) y z" apply (rule connected_componentI [OF _ subset_refl]) apply (rule connected_complement_bounded_convex) using assms yz by (auto simp: dist_norm) then have "connected_component (- S) y z" apply (rule connected_component_of_subset) apply (metis Bno Compl_anti_mono mem_cball_0 subset_iff) done } note cyz = this show ?thesis apply (auto simp: outside) apply (metis Compl_iff bounded_iff cobounded_imp_unbounded mem_Collect_eq not_le) apply (simp add: bounded_pos) by (metis B connected_component_trans cyz not_le) qed lemma not_outside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "- (outside S) = {x. \B. \y. B \ norm(y) \ \ connected_component (- S) x y}" apply (auto intro: less_imp_le simp: not_outside_connected_component_lt [OF assms]) by (meson gt_ex less_le_trans) lemma inside_connected_component_lt: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B < norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_lt [OF assms]) lemma inside_connected_component_le: fixes S :: "'a::euclidean_space set" assumes S: "bounded S" "2 \ DIM('a)" shows "inside S = {x. (x \ S) \ (\B. \y. B \ norm(y) \ \ connected_component (- S) x y)}" by (auto simp: inside_outside not_outside_connected_component_le [OF assms]) lemma inside_subset: assumes "connected U" and "\ bounded U" and "T \ U = - S" shows "inside S \ T" apply (auto simp: inside_def) by (metis bounded_subset [of "connected_component_set (- S) _"] connected_component_maximal Compl_iff Un_iff assms subsetI) lemma frontier_not_empty: fixes S :: "'a :: real_normed_vector set" shows "\S \ {}; S \ UNIV\ \ frontier S \ {}" using connected_Int_frontier [of UNIV S] by auto lemma frontier_eq_empty: fixes S :: "'a :: real_normed_vector set" shows "frontier S = {} \ S = {} \ S = UNIV" using frontier_UNIV frontier_empty frontier_not_empty by blast lemma frontier_of_connected_component_subset: fixes S :: "'a::real_normed_vector set" shows "frontier(connected_component_set S x) \ frontier S" proof - { fix y assume y1: "y \ closure (connected_component_set S x)" and y2: "y \ interior (connected_component_set S x)" have "y \ closure S" using y1 closure_mono connected_component_subset by blast moreover have "z \ interior (connected_component_set S x)" if "0 < e" "ball y e \ interior S" "dist y z < e" for e z proof - have "ball y e \ connected_component_set S y" apply (rule connected_component_maximal) using that interior_subset mem_ball apply auto done then show ?thesis using y1 apply (simp add: closure_approachable open_contains_ball_eq [OF open_interior]) by (metis connected_component_eq dist_commute mem_Collect_eq mem_ball mem_interior subsetD \0 < e\ y2) qed then have "y \ interior S" using y2 by (force simp: open_contains_ball_eq [OF open_interior]) ultimately have "y \ frontier S" by (auto simp: frontier_def) } then show ?thesis by (auto simp: frontier_def) qed lemma frontier_Union_subset_closure: fixes F :: "'a::real_normed_vector set set" shows "frontier(\F) \ closure(\t \ F. frontier t)" proof - have "\y\F. \y\frontier y. dist y x < e" if "T \ F" "y \ T" "dist y x < e" "x \ interior (\F)" "0 < e" for x y e T proof (cases "x \ T") case True with that show ?thesis by (metis Diff_iff Sup_upper closure_subset contra_subsetD dist_self frontier_def interior_mono) next case False have 1: "closed_segment x y \ T \ {}" using \y \ T\ by blast have 2: "closed_segment x y - T \ {}" using False by blast obtain c where "c \ closed_segment x y" "c \ frontier T" using False connected_Int_frontier [OF connected_segment 1 2] by auto then show ?thesis proof - have "norm (y - x) < e" by (metis dist_norm \dist y x < e\) moreover have "norm (c - x) \ norm (y - x)" by (simp add: \c \ closed_segment x y\ segment_bound(1)) ultimately have "norm (c - x) < e" by linarith then show ?thesis by (metis (no_types) \c \ frontier T\ dist_norm that(1)) qed qed then show ?thesis by (fastforce simp add: frontier_def closure_approachable) qed lemma frontier_Union_subset: fixes F :: "'a::real_normed_vector set set" shows "finite F \ frontier(\F) \ (\t \ F. frontier t)" by (rule order_trans [OF frontier_Union_subset_closure]) (auto simp: closure_subset_eq) lemma frontier_of_components_subset: fixes S :: "'a::real_normed_vector set" shows "C \ components S \ frontier C \ frontier S" by (metis Path_Connected.frontier_of_connected_component_subset components_iff) lemma frontier_of_components_closed_complement: fixes S :: "'a::real_normed_vector set" shows "\closed S; C \ components (- S)\ \ frontier C \ S" using frontier_complement frontier_of_components_subset frontier_subset_eq by blast lemma frontier_minimal_separating_closed: fixes S :: "'a::real_normed_vector set" assumes "closed S" and nconn: "\ connected(- S)" and C: "C \ components (- S)" and conn: "\T. \closed T; T \ S\ \ connected(- T)" shows "frontier C = S" proof (rule ccontr) assume "frontier C \ S" then have "frontier C \ S" using frontier_of_components_closed_complement [OF \closed S\ C] by blast then have "connected(- (frontier C))" by (simp add: conn) have "\ connected(- (frontier C))" unfolding connected_def not_not proof (intro exI conjI) show "open C" using C \closed S\ open_components by blast show "open (- closure C)" by blast show "C \ - closure C \ - frontier C = {}" using closure_subset by blast show "C \ - frontier C \ {}" using C \open C\ components_eq frontier_disjoint_eq by fastforce show "- frontier C \ C \ - closure C" by (simp add: \open C\ closed_Compl frontier_closures) then show "- closure C \ - frontier C \ {}" by (metis (no_types, lifting) C Compl_subset_Compl_iff \frontier C \ S\ compl_sup frontier_closures in_components_subset psubsetE sup.absorb_iff2 sup.boundedE sup_bot.right_neutral sup_inf_absorb) qed then show False using \connected (- frontier C)\ by blast qed lemma connected_component_UNIV [simp]: fixes x :: "'a::real_normed_vector" shows "connected_component_set UNIV x = UNIV" using connected_iff_eq_connected_component_set [of "UNIV::'a set"] connected_UNIV by auto lemma connected_component_eq_UNIV: fixes x :: "'a::real_normed_vector" shows "connected_component_set s x = UNIV \ s = UNIV" using connected_component_in connected_component_UNIV by blast lemma components_UNIV [simp]: "components UNIV = {UNIV :: 'a::real_normed_vector set}" by (auto simp: components_eq_sing_iff) lemma interior_inside_frontier: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "interior s \ inside (frontier s)" proof - { fix x y assume x: "x \ interior s" and y: "y \ s" and cc: "connected_component (- frontier s) x y" have "connected_component_set (- frontier s) x \ frontier s \ {}" apply (rule connected_Int_frontier, simp) apply (metis IntI cc connected_component_in connected_component_refl empty_iff interiorE mem_Collect_eq rev_subsetD x) using y cc by blast then have "bounded (connected_component_set (- frontier s) x)" using connected_component_in by auto } then show ?thesis apply (auto simp: inside_def frontier_def) apply (rule classical) apply (rule bounded_subset [OF assms], blast) done qed lemma inside_empty [simp]: "inside {} = ({} :: 'a :: {real_normed_vector, perfect_space} set)" by (simp add: inside_def) lemma outside_empty [simp]: "outside {} = (UNIV :: 'a :: {real_normed_vector, perfect_space} set)" using inside_empty inside_Un_outside by blast lemma inside_same_component: "\connected_component (- s) x y; x \ inside s\ \ y \ inside s" using connected_component_eq connected_component_in by (fastforce simp add: inside_def) lemma outside_same_component: "\connected_component (- s) x y; x \ outside s\ \ y \ outside s" using connected_component_eq connected_component_in by (fastforce simp add: outside_def) lemma convex_in_outside: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "convex s" and z: "z \ s" shows "z \ outside s" proof (cases "s={}") case True then show ?thesis by simp next case False then obtain a where "a \ s" by blast with z have zna: "z \ a" by auto { assume "bounded (connected_component_set (- s) z)" with bounded_pos_less obtain B where "B>0" and B: "\x. connected_component (- s) z x \ norm x < B" by (metis mem_Collect_eq) define C where "C = (B + 1 + norm z) / norm (z-a)" have "C > 0" using \0 < B\ zna by (simp add: C_def field_split_simps add_strict_increasing) have "\norm (z + C *\<^sub>R (z-a)) - norm (C *\<^sub>R (z-a))\ \ norm z" by (metis add_diff_cancel norm_triangle_ineq3) moreover have "norm (C *\<^sub>R (z-a)) > norm z + B" using zna \B>0\ by (simp add: C_def le_max_iff_disj) ultimately have C: "norm (z + C *\<^sub>R (z-a)) > B" by linarith { fix u::real assume u: "0\u" "u\1" and ins: "(1 - u) *\<^sub>R z + u *\<^sub>R (z + C *\<^sub>R (z - a)) \ s" then have Cpos: "1 + u * C > 0" by (meson \0 < C\ add_pos_nonneg less_eq_real_def zero_le_mult_iff zero_less_one) then have *: "(1 / (1 + u * C)) *\<^sub>R z + (u * C / (1 + u * C)) *\<^sub>R z = z" by (simp add: scaleR_add_left [symmetric] field_split_simps) then have False using convexD_alt [OF s \a \ s\ ins, of "1/(u*C + 1)"] \C>0\ \z \ s\ Cpos u by (simp add: * field_split_simps) } note contra = this have "connected_component (- s) z (z + C *\<^sub>R (z-a))" apply (rule connected_componentI [OF connected_segment [of z "z + C *\<^sub>R (z-a)"]]) apply (simp add: closed_segment_def) using contra apply auto done then have False using zna B [of "z + C *\<^sub>R (z-a)"] C by (auto simp: field_split_simps max_mult_distrib_right) } then show ?thesis by (auto simp: outside_def z) qed lemma outside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "convex s" shows "outside s = - s" by (metis ComplD assms convex_in_outside equalityI inside_Un_outside subsetI sup.cobounded2) lemma outside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "outside {x} = -{x}" by (auto simp: outside_convex) lemma inside_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "convex s \ inside s = {}" by (simp add: inside_outside outside_convex) lemma inside_singleton [simp]: fixes x :: "'a :: {real_normed_vector, perfect_space}" shows "inside {x} = {}" by (auto simp: inside_convex) lemma outside_subset_convex: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\convex t; s \ t\ \ - t \ outside s" using outside_convex outside_mono by blast lemma outside_Un_outside_Un: fixes S :: "'a::real_normed_vector set" assumes "S \ outside(T \ U) = {}" shows "outside(T \ U) \ outside(T \ S)" proof fix x assume x: "x \ outside (T \ U)" have "Y \ - S" if "connected Y" "Y \ - T" "Y \ - U" "x \ Y" "u \ Y" for u Y proof - have "Y \ connected_component_set (- (T \ U)) x" by (simp add: connected_component_maximal that) also have "\ \ outside(T \ U)" by (metis (mono_tags, lifting) Collect_mono mem_Collect_eq outside outside_same_component x) finally have "Y \ outside(T \ U)" . with assms show ?thesis by auto qed with x show "x \ outside (T \ S)" by (simp add: outside_connected_component_lt connected_component_def) meson qed lemma outside_frontier_misses_closure: fixes s :: "'a::real_normed_vector set" assumes "bounded s" shows "outside(frontier s) \ - closure s" unfolding outside_inside Lattices.boolean_algebra_class.compl_le_compl_iff proof - { assume "interior s \ inside (frontier s)" hence "interior s \ inside (frontier s) = inside (frontier s)" by (simp add: subset_Un_eq) then have "closure s \ frontier s \ inside (frontier s)" using frontier_def by auto } then show "closure s \ frontier s \ inside (frontier s)" using interior_inside_frontier [OF assms] by blast qed lemma outside_frontier_eq_complement_closure: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" "convex s" shows "outside(frontier s) = - closure s" by (metis Diff_subset assms convex_closure frontier_def outside_frontier_misses_closure outside_subset_convex subset_antisym) lemma inside_frontier_eq_interior: fixes s :: "'a :: {real_normed_vector, perfect_space} set" shows "\bounded s; convex s\ \ inside(frontier s) = interior s" apply (simp add: inside_outside outside_frontier_eq_complement_closure) using closure_subset interior_subset apply (auto simp: frontier_def) done lemma open_inside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (inside s)" proof - { fix x assume x: "x \ inside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis (no_types, lifting) Compl_iff connected_component_refl_eq inside_def mem_Collect_eq x) then have "\e>0. ball x e \ inside s" by (metis e dist_commute inside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma open_outside: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "open (outside s)" proof - { fix x assume x: "x \ outside s" have "open (connected_component_set (- s) x)" using assms open_connected_component by blast then obtain e where e: "e>0" and e: "\y. dist y x < e \ connected_component (- s) x y" using dist_not_less_zero apply (simp add: open_dist) by (metis Int_iff outside_def connected_component_refl_eq x) then have "\e>0. ball x e \ outside s" by (metis e dist_commute outside_same_component mem_ball subsetI x) } then show ?thesis by (simp add: open_contains_ball) qed lemma closure_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(inside s) \ s \ inside s" by (metis assms closure_minimal open_closed open_outside sup.cobounded2 union_with_inside) lemma frontier_inside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(inside s) \ s" proof - have "closure (inside s) \ - inside s = closure (inside s) - interior (inside s)" by (metis (no_types) Diff_Compl assms closure_closed interior_closure open_closed open_inside) moreover have "- inside s \ - outside s = s" by (metis (no_types) compl_sup double_compl inside_Un_outside) moreover have "closure (inside s) \ - outside s" by (metis (no_types) assms closure_inside_subset union_with_inside) ultimately have "closure (inside s) - interior (inside s) \ s" by blast then show ?thesis by (simp add: frontier_def open_inside interior_open) qed lemma closure_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "closure(outside s) \ s \ outside s" apply (rule closure_minimal, simp) by (metis assms closed_open inside_outside open_inside) lemma frontier_outside_subset: fixes s :: "'a::real_normed_vector set" assumes "closed s" shows "frontier(outside s) \ s" apply (simp add: frontier_def open_outside interior_open) by (metis Diff_subset_conv assms closure_outside_subset interior_eq open_outside sup.commute) lemma inside_complement_unbounded_connected_empty: "\connected (- s); \ bounded (- s)\ \ inside s = {}" apply (simp add: inside_def) by (meson Compl_iff bounded_subset connected_component_maximal order_refl) lemma inside_bounded_complement_connected_empty: fixes s :: "'a::{real_normed_vector, perfect_space} set" shows "\connected (- s); bounded s\ \ inside s = {}" by (metis inside_complement_unbounded_connected_empty cobounded_imp_unbounded) lemma inside_inside: assumes "s \ inside t" shows "inside s - t \ inside t" unfolding inside_def proof clarify fix x assume x: "x \ t" "x \ s" and bo: "bounded (connected_component_set (- s) x)" show "bounded (connected_component_set (- t) x)" proof (cases "s \ connected_component_set (- t) x = {}") case True show ?thesis apply (rule bounded_subset [OF bo]) apply (rule connected_component_maximal) using x True apply auto done next case False then show ?thesis using assms [unfolded inside_def] x apply (simp add: disjoint_iff_not_equal, clarify) apply (drule subsetD, assumption, auto) by (metis (no_types, hide_lams) ComplI connected_component_eq_eq) qed qed lemma inside_inside_subset: "inside(inside s) \ s" using inside_inside union_with_outside by fastforce lemma inside_outside_intersect_connected: "\connected t; inside s \ t \ {}; outside s \ t \ {}\ \ s \ t \ {}" apply (simp add: inside_def outside_def ex_in_conv [symmetric] disjoint_eq_subset_Compl, clarify) by (metis (no_types, hide_lams) Compl_anti_mono connected_component_eq connected_component_maximal contra_subsetD double_compl) lemma outside_bounded_nonempty: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes "bounded s" shows "outside s \ {}" by (metis (no_types, lifting) Collect_empty_eq Collect_mem_eq Compl_eq_Diff_UNIV Diff_cancel Diff_disjoint UNIV_I assms ball_eq_empty bounded_diff cobounded_outside convex_ball double_complement order_refl outside_convex outside_def) lemma outside_compact_in_open: fixes s :: "'a :: {real_normed_vector,perfect_space} set" assumes s: "compact s" and t: "open t" and "s \ t" "t \ {}" shows "outside s \ t \ {}" proof - have "outside s \ {}" by (simp add: compact_imp_bounded outside_bounded_nonempty s) with assms obtain a b where a: "a \ outside s" and b: "b \ t" by auto show ?thesis proof (cases "a \ t") case True with a show ?thesis by blast next case False have front: "frontier t \ - s" using \s \ t\ frontier_disjoint_eq t by auto { fix \ assume "path \" and pimg_sbs: "path_image \ - {pathfinish \} \ interior (- t)" and pf: "pathfinish \ \ frontier t" and ps: "pathstart \ = a" define c where "c = pathfinish \" have "c \ -s" unfolding c_def using front pf by blast moreover have "open (-s)" using s compact_imp_closed by blast ultimately obtain \::real where "\ > 0" and \: "cball c \ \ -s" using open_contains_cball[of "-s"] s by blast then obtain d where "d \ t" and d: "dist d c < \" using closure_approachable [of c t] pf unfolding c_def by (metis Diff_iff frontier_def) then have "d \ -s" using \ using dist_commute by (metis contra_subsetD mem_cball not_le not_less_iff_gr_or_eq) have pimg_sbs_cos: "path_image \ \ -s" using pimg_sbs apply (auto simp: path_image_def) apply (drule subsetD) using \c \ - s\ \s \ t\ interior_subset apply (auto simp: c_def) done have "closed_segment c d \ cball c \" apply (simp add: segment_convex_hull) apply (rule hull_minimal) using \\ > 0\ d apply (auto simp: dist_commute) done with \ have "closed_segment c d \ -s" by blast moreover have con_gcd: "connected (path_image \ \ closed_segment c d)" by (rule connected_Un) (auto simp: c_def \path \\ connected_path_image) ultimately have "connected_component (- s) a d" unfolding connected_component_def using pimg_sbs_cos ps by blast then have "outside s \ t \ {}" using outside_same_component [OF _ a] by (metis IntI \d \ t\ empty_iff) } note * = this have pal: "pathstart (linepath a b) \ closure (- t)" by (auto simp: False closure_def) show ?thesis by (rule exists_path_subpath_to_frontier [OF path_linepath pal _ *]) (auto simp: b) qed qed lemma inside_inside_compact_connected: fixes s :: "'a :: euclidean_space set" assumes s: "closed s" and t: "compact t" and "connected t" "s \ inside t" shows "inside s \ inside t" proof (cases "inside t = {}") case True with assms show ?thesis by auto next case False consider "DIM('a) = 1" | "DIM('a) \ 2" using antisym not_less_eq_eq by fastforce then show ?thesis proof cases case 1 then show ?thesis using connected_convex_1_gen assms False inside_convex by blast next case 2 have coms: "compact s" using assms apply (simp add: s compact_eq_bounded_closed) by (meson bounded_inside bounded_subset compact_imp_bounded) then have bst: "bounded (s \ t)" by (simp add: compact_imp_bounded t) then obtain r where "0 < r" and r: "s \ t \ ball 0 r" using bounded_subset_ballD by blast have outst: "outside s \ outside t \ {}" proof - have "- ball 0 r \ outside s" apply (rule outside_subset_convex) using r by auto moreover have "- ball 0 r \ outside t" apply (rule outside_subset_convex) using r by auto ultimately show ?thesis by (metis Compl_subset_Compl_iff Int_subset_iff bounded_ball inf.orderE outside_bounded_nonempty outside_no_overlap) qed have "s \ t = {}" using assms by (metis disjoint_iff_not_equal inside_no_overlap subsetCE) moreover have "outside s \ inside t \ {}" by (meson False assms(4) compact_eq_bounded_closed coms open_inside outside_compact_in_open t) ultimately have "inside s \ t = {}" using inside_outside_intersect_connected [OF \connected t\, of s] by (metis "2" compact_eq_bounded_closed coms connected_outside inf.commute inside_outside_intersect_connected outst) then show ?thesis using inside_inside [OF \s \ inside t\] by blast qed qed lemma connected_with_inside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ inside s)" proof (cases "s \ inside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ inside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ inside s)" if "a \ (s \ inside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ inside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "inside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_inside_subset s apply fastforce by (metis (no_types, lifting) frontier_inside_subset insertE insert_Diff interior_eq open_inside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed text\The proof is virtually the same as that above.\ lemma connected_with_outside: fixes s :: "'a :: real_normed_vector set" assumes s: "closed s" and cons: "connected s" shows "connected(s \ outside s)" proof (cases "s \ outside s = UNIV") case True with assms show ?thesis by auto next case False then obtain b where b: "b \ s" "b \ outside s" by blast have *: "\y t. y \ s \ connected t \ a \ t \ y \ t \ t \ (s \ outside s)" if "a \ (s \ outside s)" for a using that proof assume "a \ s" then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x="{a}" in exI, simp) done next assume a: "a \ outside s" show ?thesis apply (rule exists_path_subpath_to_frontier [OF path_linepath [of a b], of "outside s"]) using a apply (simp add: closure_def) apply (simp add: b) apply (rule_tac x="pathfinish h" in exI) apply (rule_tac x="path_image h" in exI) apply (simp add: pathfinish_in_path_image connected_path_image, auto) using frontier_outside_subset s apply fastforce by (metis (no_types, lifting) frontier_outside_subset insertE insert_Diff interior_eq open_outside pathfinish_in_path_image s subsetCE) qed show ?thesis apply (simp add: connected_iff_connected_component) apply (simp add: connected_component_def) apply (clarify dest!: *) apply (rename_tac u u' t t') apply (rule_tac x="(s \ t \ t')" in exI) apply (auto simp: intro!: connected_Un cons) done qed lemma inside_inside_eq_empty [simp]: fixes s :: "'a :: {real_normed_vector, perfect_space} set" assumes s: "closed s" and cons: "connected s" shows "inside (inside s) = {}" by (metis (no_types) unbounded_outside connected_with_outside [OF assms] bounded_Un inside_complement_unbounded_connected_empty unbounded_outside union_with_outside) lemma inside_in_components: "inside s \ components (- s) \ connected(inside s) \ inside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: inside_same_component connected_componentI) apply (metis IntI empty_iff inside_no_overlap) done text\The proof is virtually the same as that above.\ lemma outside_in_components: "outside s \ components (- s) \ connected(outside s) \ outside s \ {}" apply (simp add: in_components_maximal) apply (auto intro: outside_same_component connected_componentI) apply (metis IntI empty_iff outside_no_overlap) done lemma bounded_unique_outside: fixes s :: "'a :: euclidean_space set" shows "\bounded s; DIM('a) \ 2\ \ (c \ components (- s) \ \ bounded c \ c = outside s)" apply (rule iffI) apply (metis cobounded_unique_unbounded_components connected_outside double_compl outside_bounded_nonempty outside_in_components unbounded_outside) by (simp add: connected_outside outside_bounded_nonempty outside_in_components unbounded_outside) subsection\Condition for an open map's image to contain a ball\ proposition ball_subset_open_map_image: fixes f :: "'a::heine_borel \ 'b :: {real_normed_vector,heine_borel}" assumes contf: "continuous_on (closure S) f" and oint: "open (f ` interior S)" and le_no: "\z. z \ frontier S \ r \ norm(f z - f a)" and "bounded S" "a \ S" "0 < r" shows "ball (f a) r \ f ` S" proof (cases "f ` S = UNIV") case True then show ?thesis by simp next case False obtain w where w: "w \ frontier (f ` S)" and dw_le: "\y. y \ frontier (f ` S) \ norm (f a - w) \ norm (f a - y)" apply (rule distance_attains_inf [of "frontier(f ` S)" "f a"]) using \a \ S\ by (auto simp: frontier_eq_empty dist_norm False) then obtain \ where \: "\n. \ n \ f ` S" and tendsw: "\ \ w" by (metis Diff_iff frontier_def closure_sequential) then have "\n. \x \ S. \ n = f x" by force then obtain z where zs: "\n. z n \ S" and fz: "\n. \ n = f (z n)" by metis then obtain y K where y: "y \ closure S" and "strict_mono (K :: nat \ nat)" and Klim: "(z \ K) \ y" using \bounded S\ apply (simp add: compact_closure [symmetric] compact_def) apply (drule_tac x=z in spec) using closure_subset apply force done then have ftendsw: "((\n. f (z n)) \ K) \ w" by (metis LIMSEQ_subseq_LIMSEQ fun.map_cong0 fz tendsw) have zKs: "\n. (z \ K) n \ S" by (simp add: zs) have fz: "f \ z = \" "(\n. f (z n)) = \" using fz by auto then have "(\ \ K) \ f y" by (metis (no_types) Klim zKs y contf comp_assoc continuous_on_closure_sequentially) with fz have wy: "w = f y" using fz LIMSEQ_unique ftendsw by auto have rle: "r \ norm (f y - f a)" apply (rule le_no) using w wy oint by (force simp: imageI image_mono interiorI interior_subset frontier_def y) have **: "(b \ (- S) \ {} \ b - (- S) \ {} \ b \ f \ {}) \ (b \ S \ {}) \ b \ f = {} \ b \ S" for b f and S :: "'b set" by blast show ?thesis apply (rule **) (*such a horrible mess*) apply (rule connected_Int_frontier [where t = "f`S", OF connected_ball]) using \a \ S\ \0 < r\ apply (auto simp: disjoint_iff_not_equal dist_norm) by (metis dw_le norm_minus_commute not_less order_trans rle wy) qed subsubsection\Special characterizations of classes of functions into and out of R.\ proposition embedding_map_into_euclideanreal: assumes "path_connected_space X" shows "embedding_map X euclideanreal f \ continuous_map X euclideanreal f \ inj_on f (topspace X)" proof safe show "continuous_map X euclideanreal f" if "embedding_map X euclideanreal f" using continuous_map_in_subtopology homeomorphic_imp_continuous_map that unfolding embedding_map_def by blast show "inj_on f (topspace X)" if "embedding_map X euclideanreal f" using that homeomorphic_imp_injective_map unfolding embedding_map_def by blast show "embedding_map X euclideanreal f" if cont: "continuous_map X euclideanreal f" and inj: "inj_on f (topspace X)" proof - obtain g where gf: "\x. x \ topspace X \ g (f x) = x" using inv_into_f_f [OF inj] by auto show ?thesis unfolding embedding_map_def homeomorphic_map_maps homeomorphic_maps_def proof (intro exI conjI) show "continuous_map X (top_of_set (f ` topspace X)) f" by (simp add: cont continuous_map_in_subtopology) let ?S = "f ` topspace X" have eq: "{x \ ?S. g x \ U} = f ` U" if "openin X U" for U using openin_subset [OF that] by (auto simp: gf) have 1: "g ` ?S \ topspace X" using eq by blast have "openin (top_of_set ?S) {x \ ?S. g x \ T}" if "openin X T" for T proof - have "T \ topspace X" by (simp add: openin_subset that) have RR: "\x \ ?S \ g -` T. \d>0. \x' \ ?S \ ball x d. g x' \ T" proof (clarsimp simp add: gf) have pcS: "path_connectedin euclidean ?S" using assms cont path_connectedin_continuous_map_image path_connectedin_topspace by blast show "\d>0. \x'\f ` topspace X \ ball (f x) d. g x' \ T" if "x \ T" for x proof - have x: "x \ topspace X" using \T \ topspace X\ \x \ T\ by blast obtain u v d where "0 < d" "u \ topspace X" "v \ topspace X" and sub_fuv: "?S \ {f x - d .. f x + d} \ {f u..f v}" proof (cases "\u \ topspace X. f u < f x") case True then obtain u where u: "u \ topspace X" "f u < f x" .. show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "min (f x - f u) (f v - f x)" show "0 < ?d" by (simp add: \f u < f x\ \f x < f v\) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f v}" by fastforce qed (auto simp: u v) next case False show ?thesis proof let ?d = "f x - f u" show "0 < ?d" by (simp add: u) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f u..f x}" using x u False by auto qed (auto simp: x u) qed next case False note no_u = False show ?thesis proof (cases "\v \ topspace X. f x < f v") case True then obtain v where v: "v \ topspace X" "f x < f v" .. show ?thesis proof let ?d = "f v - f x" show "0 < ?d" by (simp add: v) show "f ` topspace X \ {f x - ?d..f x + ?d} \ {f x..f v}" using False by auto qed (auto simp: x v) next case False show ?thesis proof show "f ` topspace X \ {f x - 1..f x + 1} \ {f x..f x}" using False no_u by fastforce qed (auto simp: x) qed qed then obtain h where "pathin X h" "h 0 = u" "h 1 = v" using assms unfolding path_connected_space_def by blast obtain C where "compactin X C" "connectedin X C" "u \ C" "v \ C" proof show "compactin X (h ` {0..1})" using that by (simp add: \pathin X h\ compactin_path_image) show "connectedin X (h ` {0..1})" using \pathin X h\ connectedin_path_image by blast qed (use \h 0 = u\ \h 1 = v\ in auto) have "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) (subtopology X C) g" proof (rule continuous_inverse_map) show "compact_space (subtopology X C)" using \compactin X C\ compactin_subspace by blast show "continuous_map (subtopology X C) euclideanreal f" by (simp add: cont continuous_map_from_subtopology) have "{f u .. f v} \ f ` topspace (subtopology X C)" proof (rule connected_contains_Icc) show "connected (f ` topspace (subtopology X C))" using connectedin_continuous_map_image [OF cont] by (simp add: \compactin X C\ \connectedin X C\ compactin_subset_topspace inf_absorb2) show "f u \ f ` topspace (subtopology X C)" by (simp add: \u \ C\ \u \ topspace X\) show "f v \ f ` topspace (subtopology X C)" by (simp add: \v \ C\ \v \ topspace X\) qed then show "f ` topspace X \ {f x - d..f x + d} \ f ` topspace (subtopology X C)" using sub_fuv by blast qed (auto simp: gf) then have contg: "continuous_map (subtopology euclideanreal (?S \ {f x - d .. f x + d})) X g" using continuous_map_in_subtopology by blast have "\e>0. \x \ ?S \ {f x - d .. f x + d} \ ball (f x) e. g x \ T" using openin_continuous_map_preimage [OF contg \openin X T\] x \x \ T\ \0 < d\ unfolding openin_euclidean_subtopology_iff by (force simp: gf dist_commute) then obtain e where "e > 0 \ (\x\f ` topspace X \ {f x - d..f x + d} \ ball (f x) e. g x \ T)" by metis with \0 < d\ have "min d e > 0" "\u. u \ topspace X \ \f x - f u\ < min d e \ u \ T" using dist_real_def gf by force+ then show ?thesis by (metis (full_types) Int_iff dist_real_def image_iff mem_ball gf) qed qed then obtain d where d: "\r. r \ ?S \ g -` T \ d r > 0 \ (\x \ ?S \ ball r (d r). g x \ T)" by metis show ?thesis unfolding openin_subtopology proof (intro exI conjI) show "{x \ ?S. g x \ T} = (\r \ ?S \ g -` T. ball r (d r)) \ f ` topspace X" using d by (auto simp: gf) qed auto qed then show "continuous_map (top_of_set ?S) X g" by (simp add: continuous_map_def gf) qed (auto simp: gf) qed qed subsubsection \An injective function into R is a homeomorphism and so an open map.\ lemma injective_into_1d_eq_homeomorphism: fixes f :: "'a::topological_space \ real" assumes f: "continuous_on S f" and S: "path_connected S" shows "inj_on f S \ (\g. homeomorphism S (f ` S) f g)" proof show "\g. homeomorphism S (f ` S) f g" if "inj_on f S" proof - have "embedding_map (top_of_set S) euclideanreal f" using that embedding_map_into_euclideanreal [of "top_of_set S" f] assms by auto then show ?thesis by (simp add: embedding_map_def) (metis all_closedin_homeomorphic_image f homeomorphism_injective_closed_map that) qed qed (metis homeomorphism_def inj_onI) lemma injective_into_1d_imp_open_map: fixes f :: "'a::topological_space \ real" assumes "continuous_on S f" "path_connected S" "inj_on f S" "openin (subtopology euclidean S) T" shows "openin (subtopology euclidean (f ` S)) (f ` T)" using assms homeomorphism_imp_open_map injective_into_1d_eq_homeomorphism by blast lemma homeomorphism_into_1d: fixes f :: "'a::topological_space \ real" assumes "path_connected S" "continuous_on S f" "f ` S = T" "inj_on f S" shows "\g. homeomorphism S T f g" using assms injective_into_1d_eq_homeomorphism by blast + +subsection\<^marker>\tag unimportant\ \Rectangular paths\ + +definition\<^marker>\tag unimportant\ 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 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) + end diff --git a/src/HOL/Analysis/Smooth_Paths.thy b/src/HOL/Analysis/Smooth_Paths.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Analysis/Smooth_Paths.thy @@ -0,0 +1,490 @@ +(* + Material originally from HOL Light, ported by Larry Paulson, moved here by Manuel Eberl +*) +section\<^marker>\tag unimportant\ \Smooth paths\ +theory Smooth_Paths + imports + Retracts +begin + +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 \Piecewise differentiability of paths\ + +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 + +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) + 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 + +lemma valid_path_imp_reverse: + assumes "valid_path g" + shows "valid_path(reversepath g)" +proof - + obtain S where "finite S" and S: "g C1_differentiable_on ({0..1} - S)" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + then have "finite ((-) 1 ` S)" + by auto + moreover have "(reversepath g C1_differentiable_on ({0..1} - (-) 1 ` S))" + unfolding reversepath_def + apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) + using S + by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq elim!: continuous_on_subset)+ + ultimately show ?thesis using assms + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) +qed + +lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" + using valid_path_imp_reverse by force + +lemma 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 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 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 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 valid_path_rectpath [simp, intro]: "valid_path (rectpath a b)" + by (simp add: Let_def rectpath_def) + +end \ No newline at end of file diff --git a/src/HOL/Analysis/Vitali_Covering_Theorem.thy b/src/HOL/Analysis/Vitali_Covering_Theorem.thy --- a/src/HOL/Analysis/Vitali_Covering_Theorem.thy +++ b/src/HOL/Analysis/Vitali_Covering_Theorem.thy @@ -1,657 +1,657 @@ (* Title: HOL/Analysis/Vitali_Covering_Theorem.thy Authors: LC Paulson, based on material from HOL Light *) section \Vitali Covering Theorem and an Application to Negligibility\ theory Vitali_Covering_Theorem imports Ball_Volume "HOL-Library.Permutations" begin lemma stretch_Galois: fixes x :: "real^'n" shows "(\k. m k \ 0) \ ((y = (\ k. m k * x$k)) \ (\ k. y$k / m k) = x)" by auto lemma lambda_swap_Galois: "(x = (\ i. y $ Fun.swap m n id i) \ (\ i. x $ Fun.swap m n id i) = y)" by (auto; simp add: pointfree_idE vec_eq_iff) lemma lambda_add_Galois: fixes x :: "real^'n" shows "m \ n \ (x = (\ i. if i = m then y$m + y$n else y$i) \ (\ i. if i = m then x$m - x$n else x$i) = y)" by (safe; simp add: vec_eq_iff) lemma Vitali_covering_lemma_cballs_balls: fixes a :: "'a \ 'b::euclidean_space" assumes "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" proof (cases "K = {}") case True with that show ?thesis by auto next case False then have "B > 0" using assms less_le_trans by auto have rgt0[simp]: "\i. i \ K \ 0 < r i" using assms by auto let ?djnt = "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j)))" have "\C. \n. (C n \ K \ (\i \ C n. B/2 ^ n \ r i) \ ?djnt (C n) \ (\i \ K. B/2 ^ n < r i \ (\j. j \ C n \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)))) \ (C n \ C(Suc n))" proof (rule dependent_nat_choice, safe) fix C n define D where "D \ {i \ K. B/2 ^ Suc n < r i \ (\j\C. disjnt (cball(a i)(r i)) (cball (a j) (r j)))}" let ?cover_ar = "\i j. \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" assume "C \ K" and Ble: "\i\C. B/2 ^ n \ r i" and djntC: "?djnt C" and cov_n: "\i\K. B/2 ^ n < r i \ (\j. j \ C \ ?cover_ar i j)" have *: "\C\chains {C. C \ D \ ?djnt C}. \C \ {C. C \ D \ ?djnt C}" proof (clarsimp simp: chains_def) fix C assume C: "C \ {C. C \ D \ ?djnt C}" and "chain\<^sub>\ C" show "\C \ D \ ?djnt (\C)" unfolding pairwise_def proof (intro ballI conjI impI) show "\C \ D" using C by blast next fix x y assume "x \ \C" and "y \ \C" and "x \ y" then obtain X Y where XY: "x \ X" "X \ C" "y \ Y" "Y \ C" by blast then consider "X \ Y" | "Y \ X" by (meson \chain\<^sub>\ C\ chain_subset_def) then show "disjnt (cball (a x) (r x)) (cball (a y) (r y))" proof cases case 1 with C XY \x \ y\ show ?thesis unfolding pairwise_def by blast next case 2 with C XY \x \ y\ show ?thesis unfolding pairwise_def by blast qed qed qed obtain E where "E \ D" and djntE: "?djnt E" and maximalE: "\X. \X \ D; ?djnt X; E \ X\ \ X = E" using Zorn_Lemma [OF *] by safe blast show "\L. (L \ K \ (\i\L. B/2 ^ Suc n \ r i) \ ?djnt L \ (\i\K. B/2 ^ Suc n < r i \ (\j. j \ L \ ?cover_ar i j))) \ C \ L" proof (intro exI conjI ballI) show "C \ E \ K" using D_def \C \ K\ \E \ D\ by blast show "B/2 ^ Suc n \ r i" if i: "i \ C \ E" for i using i proof assume "i \ C" have "B/2 ^ Suc n \ B/2 ^ n" using \B > 0\ by (simp add: field_split_simps) also have "\ \ r i" using Ble \i \ C\ by blast finally show ?thesis . qed (use D_def \E \ D\ in auto) show "?djnt (C \ E)" using D_def \C \ K\ \E \ D\ djntC djntE unfolding pairwise_def disjnt_def by blast next fix i assume "i \ K" show "B/2 ^ Suc n < r i \ (\j. j \ C \ E \ ?cover_ar i j)" proof (cases "r i \ B/2^n") case False then show ?thesis using cov_n \i \ K\ by auto next case True have "cball (a i) (r i) \ ball (a j) (5 * r j)" if less: "B/2 ^ Suc n < r i" and j: "j \ C \ E" and nondis: "\ disjnt (cball (a i) (r i)) (cball (a j) (r j))" for j proof - obtain x where x: "dist (a i) x \ r i" "dist (a j) x \ r j" using nondis by (force simp: disjnt_def) have "dist (a i) (a j) \ dist (a i) x + dist x (a j)" by (simp add: dist_triangle) also have "\ \ r i + r j" by (metis add_mono_thms_linordered_semiring(1) dist_commute x) finally have aij: "dist (a i) (a j) + r i < 5 * r j" if "r i < 2 * r j" using that by auto show ?thesis using j proof assume "j \ C" have "B/2^n < 2 * r j" using Ble True \j \ C\ less by auto with aij True show "cball (a i) (r i) \ ball (a j) (5 * r j)" by (simp add: cball_subset_ball_iff) next assume "j \ E" then have "B/2 ^ n < 2 * r j" using D_def \E \ D\ by auto with True have "r i < 2 * r j" by auto with aij show "cball (a i) (r i) \ ball (a j) (5 * r j)" by (simp add: cball_subset_ball_iff) qed qed moreover have "\j. j \ C \ E \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j))" if "B/2 ^ Suc n < r i" proof (rule classical) assume NON: "\ ?thesis" show ?thesis proof (cases "i \ D") case True have "insert i E = E" proof (rule maximalE) show "insert i E \ D" by (simp add: True \E \ D\) show "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (insert i E)" using False NON by (auto simp: pairwise_insert djntE disjnt_sym) qed auto then show ?thesis using \i \ K\ assms by fastforce next case False with that show ?thesis by (auto simp: D_def disjnt_def \i \ K\) qed qed ultimately show "B/2 ^ Suc n < r i \ (\j. j \ C \ E \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j))" by blast qed qed auto qed (use assms in force) then obtain F where FK: "\n. F n \ K" and Fle: "\n i. i \ F n \ B/2 ^ n \ r i" and Fdjnt: "\n. ?djnt (F n)" and FF: "\n i. \i \ K; B/2 ^ n < r i\ \ \j. j \ F n \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" and inc: "\n. F n \ F(Suc n)" by (force simp: all_conj_distrib) show thesis proof have *: "countable I" if "I \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) I" for I proof - show ?thesis proof (rule countable_image_inj_on [of "\i. cball(a i)(r i)"]) show "countable ((\i. cball (a i) (r i)) ` I)" proof (rule countable_disjoint_nonempty_interior_subsets) show "disjoint ((\i. cball (a i) (r i)) ` I)" by (auto simp: dest: pairwiseD [OF pw] intro: pairwise_imageI) show "\S. \S \ (\i. cball (a i) (r i)) ` I; interior S = {}\ \ S = {}" using \I \ K\ by (auto simp: not_less [symmetric]) qed next have "\x y. \x \ I; y \ I; a x = a y; r x = r y\ \ x = y" using pw \I \ K\ assms apply (clarsimp simp: pairwise_def disjnt_def) by (metis assms centre_in_cball subsetD empty_iff inf.idem less_eq_real_def) then show "inj_on (\i. cball (a i) (r i)) I" using \I \ K\ by (fastforce simp: inj_on_def cball_eq_cball_iff dest: assms) qed qed show "(Union(range F)) \ K" using FK by blast moreover show "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) (Union(range F))" proof (rule pairwise_chain_Union) show "chain\<^sub>\ (range F)" unfolding chain_subset_def by clarify (meson inc lift_Suc_mono_le linear subsetCE) qed (use Fdjnt in blast) ultimately show "countable (Union(range F))" by (blast intro: *) next fix i assume "i \ K" then obtain n where "(1/2) ^ n < r i / B" using \B > 0\ assms real_arch_pow_inv by fastforce then have B2: "B/2 ^ n < r i" using \B > 0\ by (simp add: field_split_simps) have "0 < r i" "r i \ B" by (auto simp: \i \ K\ assms) show "\j. j \ (Union(range F)) \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" using FF [OF \i \ K\ B2] by auto qed qed subsection\Vitali covering theorem\ lemma Vitali_covering_lemma_cballs: fixes a :: "'a \ 'b::euclidean_space" assumes S: "S \ (\i\K. cball (a i) (r i))" and r: "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "S \ (\i\C. cball (a i) (5 * r i))" proof - obtain C where C: "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ show ?thesis proof have "(\i\K. cball (a i) (r i)) \ (\i\C. cball (a i) (5 * r i))" using cov subset_iff by fastforce with S show "S \ (\i\C. cball (a i) (5 * r i))" by blast qed (use C in auto) qed lemma Vitali_covering_lemma_balls: fixes a :: "'a \ 'b::euclidean_space" assumes S: "S \ (\i\K. ball (a i) (r i))" and r: "\i. i \ K \ 0 < r i \ r i \ B" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" "S \ (\i\C. ball (a i) (5 * r i))" proof - obtain C where C: "countable C" "C \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [OF r, where a=a]) (blast intro: that)+ show ?thesis proof have "(\i\K. ball (a i) (r i)) \ (\i\C. ball (a i) (5 * r i))" using cov subset_iff by clarsimp (meson less_imp_le mem_ball mem_cball subset_eq) with S show "S \ (\i\C. ball (a i) (5 * r i))" by blast show "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" using pw by (clarsimp simp: pairwise_def) (meson ball_subset_cball disjnt_subset1 disjnt_subset2) qed (use C in auto) qed theorem Vitali_covering_theorem_cballs: fixes a :: "'a \ 'n::euclidean_space" assumes r: "\i. i \ K \ 0 < r i" and S: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ cball (a i) (r i) \ r i < d" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" "negligible(S - (\i \ C. cball (a i) (r i)))" proof - let ?\ = "measure lebesgue" have *: "\C. countable C \ C \ K \ pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C \ negligible(S - (\i \ C. cball (a i) (r i)))" if r01: "\i. i \ K \ 0 < r i \ r i \ 1" and Sd: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ cball (a i) (r i) \ r i < d" for K r and a :: "'a \ 'n" proof - obtain C where C: "countable C" "C \ K" and pwC: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and cov: "\i. i \ K \ \j. j \ C \ \ disjnt (cball (a i) (r i)) (cball (a j) (r j)) \ cball (a i) (r i) \ ball (a j) (5 * r j)" by (rule Vitali_covering_lemma_cballs_balls [of K r 1 a]) (auto simp: r01) have ar_injective: "\x y. \x \ C; y \ C; a x = a y; r x = r y\ \ x = y" using \C \ K\ pwC cov by (force simp: pairwise_def disjnt_def) show ?thesis proof (intro exI conjI) show "negligible (S - (\i\C. cball (a i) (r i)))" proof (clarsimp simp: negligible_on_intervals [of "S-T" for T]) fix l u show "negligible ((S - (\i\C. cball (a i) (r i))) \ cbox l u)" unfolding negligible_outer_le proof (intro allI impI) fix e::real assume "e > 0" define D where "D \ {i \ C. \ disjnt (ball(a i) (5 * r i)) (cbox l u)}" then have "D \ C" by auto have "countable D" unfolding D_def using \countable C\ by simp have UD: "(\i\D. cball (a i) (r i)) \ lmeasurable" proof (rule fmeasurableI2) show "cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One) \ lmeasurable" by blast have "y \ cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" if "i \ C" and x: "x \ cbox l u" and ai: "dist (a i) y \ r i" "dist (a i) x < 5 * r i" for i x y proof - have d6: "dist y x < 6 * r i" using dist_triangle3 [of y x "a i"] that by linarith show ?thesis proof (clarsimp simp: mem_box algebra_simps) fix j::'n assume j: "j \ Basis" then have xyj: "\x \ j - y \ j\ \ dist y x" by (metis Basis_le_norm dist_commute dist_norm inner_diff_left) have "l \ j \ x \ j" using \j \ Basis\ mem_box \x \ cbox l u\ by blast also have "\ \ y \ j + 6 * r i" using d6 xyj by (auto simp: algebra_simps) also have "\ \ y \ j + 6" using r01 [of i] \C \ K\ \i \ C\ by auto finally have l: "l \ j \ y \ j + 6" . have "y \ j \ x \ j + 6 * r i" using d6 xyj by (auto simp: algebra_simps) also have "\ \ u \ j + 6 * r i" using j x by (auto simp: mem_box) also have "\ \ u \ j + 6" using r01 [of i] \C \ K\ \i \ C\ by auto finally have u: "y \ j \ u \ j + 6" . show "l \ j \ y \ j + 6 \ y \ j \ u \ j + 6" using l u by blast qed qed then show "(\i\D. cball (a i) (r i)) \ cbox (l - 6 *\<^sub>R One) (u + 6 *\<^sub>R One)" by (force simp: D_def disjnt_def) show "(\i\D. cball (a i) (r i)) \ sets lebesgue" using \countable D\ by auto qed obtain D1 where "D1 \ D" "finite D1" and measD1: "?\ (\i\D. cball (a i) (r i)) - e / 5 ^ DIM('n) < ?\ (\i\D1. cball (a i) (r i))" proof (rule measure_countable_Union_approachable [where e = "e / 5 ^ (DIM('n))"]) show "countable ((\i. cball (a i) (r i)) ` D)" using \countable D\ by auto show "\d. d \ (\i. cball (a i) (r i)) ` D \ d \ lmeasurable" by auto show "\D'. \D' \ (\i. cball (a i) (r i)) ` D; finite D'\ \ ?\ (\D') \ ?\ (\i\D. cball (a i) (r i))" by (fastforce simp add: intro!: measure_mono_fmeasurable UD) qed (use \e > 0\ in \auto dest: finite_subset_image\) show "\T. (S - (\i\C. cball (a i) (r i))) \ cbox l u \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "(S - (\i\C. cball (a i) (r i))) \ cbox l u \ (\i\D - D1. ball (a i) (5 * r i))" proof clarify fix x assume x: "x \ cbox l u" "x \ S" "x \ (\i\C. cball (a i) (r i))" have "closed (\i\D1. cball (a i) (r i))" using \finite D1\ by blast moreover have "x \ (\j\D1. cball (a j) (r j))" using x \D1 \ D\ unfolding D_def by blast ultimately obtain q where "q > 0" and q: "ball x q \ - (\i\D1. cball (a i) (r i))" by (metis (no_types, lifting) ComplI open_contains_ball closed_def) obtain i where "i \ K" and xi: "x \ cball (a i) (r i)" and ri: "r i < q/2" using Sd [OF \x \ S\] \q > 0\ half_gt_zero by blast then obtain j where "j \ C" and nondisj: "\ disjnt (cball (a i) (r i)) (cball (a j) (r j))" and sub5j: "cball (a i) (r i) \ ball (a j) (5 * r j)" using cov [OF \i \ K\] by metis show "x \ (\i\D - D1. ball (a i) (5 * r i))" proof show "j \ D - D1" proof show "j \ D" using \j \ C\ sub5j \x \ cbox l u\ xi by (auto simp: D_def disjnt_def) obtain y where yi: "dist (a i) y \ r i" and yj: "dist (a j) y \ r j" using disjnt_def nondisj by fastforce have "dist x y \ r i + r i" by (metis add_mono dist_commute dist_triangle_le mem_cball xi yi) also have "\ < q" using ri by linarith finally have "y \ ball x q" by simp with yj q show "j \ D1" by (auto simp: disjoint_UN_iff) qed show "x \ ball (a j) (5 * r j)" using xi sub5j by blast qed qed have 3: "?\ (\i\D2. ball (a i) (5 * r i)) \ e" if D2: "D2 \ D - D1" and "finite D2" for D2 proof - have rgt0: "0 < r i" if "i \ D2" for i using \C \ K\ D_def \i \ D2\ D2 r01 by (simp add: subset_iff) then have inj: "inj_on (\i. ball (a i) (5 * r i)) D2" using \C \ K\ D2 by (fastforce simp: inj_on_def D_def ball_eq_ball_iff intro: ar_injective) have "?\ (\i\D2. ball (a i) (5 * r i)) \ sum (?\) ((\i. ball (a i) (5 * r i)) ` D2)" using that by (force intro: measure_Union_le) also have "\ = (\i\D2. ?\ (ball (a i) (5 * r i)))" by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) also have "\ = (\i\D2. 5 ^ DIM('n) * ?\ (ball (a i) (r i)))" proof (rule sum.cong [OF refl]) fix i assume "i \ D2" show "?\ (ball (a i) (5 * r i)) = 5 ^ DIM('n) * ?\ (ball (a i) (r i))" using rgt0 [OF \i \ D2\] by (simp add: content_ball) qed also have "\ = (\i\D2. ?\ (ball (a i) (r i))) * 5 ^ DIM('n)" by (simp add: sum_distrib_left mult.commute) finally have "?\ (\i\D2. ball (a i) (5 * r i)) \ (\i\D2. ?\ (ball (a i) (r i))) * 5 ^ DIM('n)" . moreover have "(\i\D2. ?\ (ball (a i) (r i))) \ e / 5 ^ DIM('n)" proof - have D12_dis: "((\x\D1. cball (a x) (r x)) \ (\x\D2. cball (a x) (r x))) \ {}" proof clarify fix w d1 d2 assume "d1 \ D1" "w d1 d2 \ cball (a d1) (r d1)" "d2 \ D2" "w d1 d2 \ cball (a d2) (r d2)" then show "w d1 d2 \ {}" by (metis DiffE disjnt_iff subsetCE D2 \D1 \ D\ \D \ C\ pairwiseD [OF pwC, of d1 d2]) qed have inj: "inj_on (\i. cball (a i) (r i)) D2" using rgt0 D2 \D \ C\ by (force simp: inj_on_def cball_eq_cball_iff intro!: ar_injective) have ds: "disjoint ((\i. cball (a i) (r i)) ` D2)" using D2 \D \ C\ by (auto intro: pairwiseI pairwiseD [OF pwC]) have "(\i\D2. ?\ (ball (a i) (r i))) = (\i\D2. ?\ (cball (a i) (r i)))" using rgt0 by (simp add: content_ball content_cball less_eq_real_def) also have "\ = sum ?\ ((\i. cball (a i) (r i)) ` D2)" by (simp add: comm_monoid_add_class.sum.reindex [OF inj]) also have "\ = ?\ (\i\D2. cball (a i) (r i))" by (auto intro: measure_Union' [symmetric] ds simp add: \finite D2\) finally have "?\ (\i\D1. cball (a i) (r i)) + (\i\D2. ?\ (ball (a i) (r i))) = ?\ (\i\D1. cball (a i) (r i)) + ?\ (\i\D2. cball (a i) (r i))" by simp also have "\ = ?\ (\i \ D1 \ D2. cball (a i) (r i))" using D12_dis by (simp add: measure_Un3 \finite D1\ \finite D2\ fmeasurable.finite_UN) also have "\ \ ?\ (\i\D. cball (a i) (r i))" using D2 \D1 \ D\ by (fastforce intro!: measure_mono_fmeasurable [OF _ _ UD] \finite D1\ \finite D2\) finally have "?\ (\i\D1. cball (a i) (r i)) + (\i\D2. ?\ (ball (a i) (r i))) \ ?\ (\i\D. cball (a i) (r i))" . with measD1 show ?thesis by simp qed ultimately show ?thesis by (simp add: field_split_simps) qed have co: "countable (D - D1)" by (simp add: \countable D\) show "(\i\D - D1. ball (a i) (5 * r i)) \ lmeasurable" using \e > 0\ by (auto simp: fmeasurable_UN_bound [OF co _ 3]) show "?\ (\i\D - D1. ball (a i) (5 * r i)) \ e" using \e > 0\ by (auto simp: measure_UN_bound [OF co _ 3]) qed qed qed qed (use C pwC in auto) qed define K' where "K' \ {i \ K. r i \ 1}" have 1: "\i. i \ K' \ 0 < r i \ r i \ 1" using K'_def r by auto have 2: "\i. i \ K' \ x \ cball (a i) (r i) \ r i < d" if "x \ S \ 0 < d" for x d using that by (auto simp: K'_def dest!: S [where d = "min d 1"]) have "K' \ K" using K'_def by auto then show thesis using * [OF 1 2] that by fastforce qed theorem Vitali_covering_theorem_balls: fixes a :: "'a \ 'b::euclidean_space" assumes S: "\x d. \x \ S; 0 < d\ \ \i. i \ K \ x \ ball (a i) (r i) \ r i < d" obtains C where "countable C" "C \ K" "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" "negligible(S - (\i \ C. ball (a i) (r i)))" proof - have 1: "\i. i \ {i \ K. 0 < r i} \ x \ cball (a i) (r i) \ r i < d" if xd: "x \ S" "d > 0" for x d by (metis (mono_tags, lifting) assms ball_eq_empty less_eq_real_def mem_Collect_eq mem_ball mem_cball not_le xd(1) xd(2)) obtain C where C: "countable C" "C \ K" and pw: "pairwise (\i j. disjnt (cball (a i) (r i)) (cball (a j) (r j))) C" and neg: "negligible(S - (\i \ C. cball (a i) (r i)))" by (rule Vitali_covering_theorem_cballs [of "{i \ K. 0 < r i}" r S a, OF _ 1]) auto show thesis proof show "pairwise (\i j. disjnt (ball (a i) (r i)) (ball (a j) (r j))) C" apply (rule pairwise_mono [OF pw]) apply (auto simp: disjnt_def) by (meson disjoint_iff_not_equal less_imp_le mem_cball) have "negligible (\i\C. sphere (a i) (r i))" by (auto intro: negligible_sphere \countable C\) then have "negligible (S - (\i \ C. cball(a i)(r i)) \ (\i \ C. sphere (a i) (r i)))" by (rule negligible_Un [OF neg]) then show "negligible (S - (\i\C. ball (a i) (r i)))" by (rule negligible_subset) force qed (use C in auto) qed lemma negligible_eq_zero_density_alt: "negligible S \ (\x \ S. \e > 0. \d U. 0 < d \ d \ e \ S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d))" (is "_ = (\x \ S. \e > 0. ?Q x e)") proof (intro iffI ballI allI impI) fix x and e :: real assume "negligible S" and "x \ S" and "e > 0" then show "\d U. 0 < d \ d \ e \ S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d)" apply (rule_tac x=e in exI) apply (rule_tac x="S \ ball x e" in exI) apply (auto simp: negligible_imp_measurable negligible_Int negligible_imp_measure0 zero_less_measure_iff) done next assume R [rule_format]: "\x \ S. \e > 0. ?Q x e" let ?\ = "measure lebesgue" have "\U. openin (top_of_set S) U \ z \ U \ negligible U" if "z \ S" for z proof (intro exI conjI) show "openin (top_of_set S) (S \ ball z 1)" by (simp add: openin_open_Int) show "z \ S \ ball z 1" using \z \ S\ by auto show "negligible (S \ ball z 1)" proof (clarsimp simp: negligible_outer_le) fix e :: "real" assume "e > 0" let ?K = "{(x,d). x \ S \ 0 < d \ ball x d \ ball z 1 \ (\U. S \ ball x d \ U \ U \ lmeasurable \ ?\ U < e / ?\ (ball z 1) * ?\ (ball x d))}" obtain C where "countable C" and Csub: "C \ ?K" and pwC: "pairwise (\i j. disjnt (ball (fst i) (snd i)) (ball (fst j) (snd j))) C" and negC: "negligible((S \ ball z 1) - (\i \ C. ball (fst i) (snd i)))" proof (rule Vitali_covering_theorem_balls [of "S \ ball z 1" ?K fst snd]) fix x and d :: "real" assume x: "x \ S \ ball z 1" and "d > 0" obtain k where "k > 0" and k: "ball x k \ ball z 1" by (meson Int_iff open_ball openE x) let ?\ = "min (e / ?\ (ball z 1) / 2) (min (d / 2) k)" obtain r U where r: "r > 0" "r \ ?\" and U: "S \ ball x r \ U" "U \ lmeasurable" and mU: "?\ U < ?\ * ?\ (ball x r)" using R [of x ?\] \d > 0\ \e > 0\ \k > 0\ x by auto show "\i. i \ ?K \ x \ ball (fst i) (snd i) \ snd i < d" proof (rule exI [of _ "(x,r)"], simp, intro conjI exI) have "ball x r \ ball x k" using r by (simp add: ball_subset_ball_iff) also have "\ \ ball z 1" using ball_subset_ball_iff k by auto finally show "ball x r \ ball z 1" . have "?\ * ?\ (ball x r) \ e * content (ball x r) / content (ball z 1)" using r \e > 0\ by (simp add: ord_class.min_def field_split_simps) with mU show "?\ U < e * content (ball x r) / content (ball z 1)" by auto qed (use r U x in auto) qed have "\U. case p of (x,d) \ S \ ball x d \ U \ U \ lmeasurable \ ?\ U < e / ?\ (ball z 1) * ?\ (ball x d)" if "p \ C" for p - using that Csub by auto + using that Csub unfolding case_prod_unfold by blast then obtain U where U: "\p. p \ C \ case p of (x,d) \ S \ ball x d \ U p \ U p \ lmeasurable \ ?\ (U p) < e / ?\ (ball z 1) * ?\ (ball x d)" by (rule that [OF someI_ex]) let ?T = "((S \ ball z 1) - (\(x,d)\C. ball x d)) \ \(U ` C)" show "\T. S \ ball z 1 \ T \ T \ lmeasurable \ ?\ T \ e" proof (intro exI conjI) show "S \ ball z 1 \ ?T" using U by fastforce { have Um: "U i \ lmeasurable" if "i \ C" for i using that U by blast have lee: "?\ (\i\I. U i) \ e" if "I \ C" "finite I" for I proof - have "?\ (\(x,d)\I. ball x d) \ ?\ (ball z 1)" apply (rule measure_mono_fmeasurable) using \I \ C\ \finite I\ Csub by (force simp: prod.case_eq_if sets.finite_UN)+ then have le1: "(?\ (\(x,d)\I. ball x d) / ?\ (ball z 1)) \ 1" by simp have "?\ (\i\I. U i) \ (\i\I. ?\ (U i))" using that U by (blast intro: measure_UNION_le) also have "\ \ (\(x,r)\I. e / ?\ (ball z 1) * ?\ (ball x r))" by (rule sum_mono) (use \I \ C\ U in force) also have "\ = (e / ?\ (ball z 1)) * (\(x,r)\I. ?\ (ball x r))" by (simp add: case_prod_app prod.case_distrib sum_distrib_left) also have "\ = e * (?\ (\(x,r)\I. ball x r) / ?\ (ball z 1))" apply (subst measure_UNION') using that pwC by (auto simp: case_prod_unfold elim: pairwise_mono) also have "\ \ e" by (metis mult.commute mult.left_neutral real_mult_le_cancel_iff1 \e > 0\ le1) finally show ?thesis . qed have "\(U ` C) \ lmeasurable" "?\ (\(U ` C)) \ e" using \e > 0\ Um lee by(auto intro!: fmeasurable_UN_bound [OF \countable C\] measure_UN_bound [OF \countable C\]) } moreover have "?\ ?T = ?\ (\(U ` C))" proof (rule measure_negligible_symdiff [OF \\(U ` C) \ lmeasurable\]) show "negligible((\(U ` C) - ?T) \ (?T - \(U ` C)))" by (force intro!: negligible_subset [OF negC]) qed ultimately show "?T \ lmeasurable" "?\ ?T \ e" by (simp_all add: fmeasurable.Un negC negligible_imp_measurable split_def) qed qed qed with locally_negligible_alt show "negligible S" by metis qed proposition negligible_eq_zero_density: "negligible S \ (\x\S. \r>0. \e>0. \d. 0 < d \ d \ r \ (\U. S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * measure lebesgue (ball x d)))" proof - let ?Q = "\x d e. \U. S \ ball x d \ U \ U \ lmeasurable \ measure lebesgue U < e * content (ball x d)" have "(\e>0. \d>0. d \ e \ ?Q x d e) = (\r>0. \e>0. \d>0. d \ r \ ?Q x d e)" if "x \ S" for x proof (intro iffI allI impI) fix r :: "real" and e :: "real" assume L [rule_format]: "\e>0. \d>0. d \ e \ ?Q x d e" and "r > 0" "e > 0" show "\d>0. d \ r \ ?Q x d e" using L [of "min r e"] apply (rule ex_forward) using \r > 0\ \e > 0\ by (auto intro: less_le_trans elim!: ex_forward) qed auto then show ?thesis by (force simp: negligible_eq_zero_density_alt) qed end diff --git a/src/HOL/Analysis/Winding_Numbers.thy b/src/HOL/Analysis/Winding_Numbers.thy deleted file mode 100644 --- a/src/HOL/Analysis/Winding_Numbers.thy +++ /dev/null @@ -1,1211 +0,0 @@ -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 -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 - diff --git a/src/HOL/Complex_Analysis/Cauchy_Integral_Theorem.thy b/src/HOL/Complex_Analysis/Cauchy_Integral_Theorem.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/Cauchy_Integral_Theorem.thy @@ -0,0 +1,7159 @@ +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 + "HOL-Analysis.Analysis" +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\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 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 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 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 Henstock_Kurzweil_Integration.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_contour_integral_linepath: + shows "(f has_contour_integral i) (linepath a b) \ + ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" + by (simp add: has_contour_integral) + +lemma 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 + + +subsection\Relation to subpath construction\ + +lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" + by (simp add: has_contour_integral subpath_def) + +lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" + using has_contour_integral_subpath_refl contour_integrable_on_def by blast + +lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" + by (simp add: contour_integral_unique) + +lemma has_contour_integral_subpath: + assumes f: "f contour_integrable_on g" and g: "valid_path g" + and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) + (subpath u v g)" +proof (cases "v=u") + case True + then show ?thesis + using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) +next + case False + obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" + using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast + have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) + has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) + {0..1}" + using f uv + apply (simp add: contour_integrable_on subpath_def has_contour_integral) + apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) + apply (simp_all add: has_integral_integral) + apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) + apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) + apply (simp add: divide_simps False) + done + { fix x + have "x \ {0..1} \ + x \ (\t. (v-u) *\<^sub>R t + u) -` s \ + vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" + apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) + apply (intro derivative_eq_intros | simp)+ + apply (cut_tac s [of "(v - u) * x + u"]) + using uv mult_left_le [of x "v-u"] + apply (auto simp: vector_derivative_works) + done + } note vd = this + show ?thesis + apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) + using fs assms + apply (simp add: False subpath_def has_contour_integral) + apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) + apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) + done +qed + +lemma contour_integrable_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" + shows "f contour_integrable_on (subpath u v g)" + apply (cases u v rule: linorder_class.le_cases) + apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) + apply (subst reversepath_subpath [symmetric]) + apply (rule contour_integrable_reversepath) + using assms apply (blast intro: valid_path_subpath) + apply (simp add: contour_integrable_on_def) + using assms apply (blast intro: has_contour_integral_subpath) + done + +lemma has_integral_contour_integral_subpath: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "(((\x. f(g x) * vector_derivative g (at x))) + has_integral contour_integral (subpath u v g) f) {u..v}" + using assms + apply (auto simp: has_integral_integrable_integral) + apply (rule integrable_on_subcbox [where a=u and b=v and S = "{0..1}", simplified]) + apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) + done + +lemma contour_integral_subcontour_integral: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" + shows "contour_integral (subpath u v g) f = + integral {u..v} (\x. f(g x) * vector_derivative g (at x))" + using assms has_contour_integral_subpath contour_integral_unique by blast + +lemma contour_integral_subpath_combine_less: + assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" + "u {0..1}" "v \ {0..1}" "w \ {0..1}" + shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = + contour_integral (subpath u w g) f" +proof (cases "u\v \ v\w \ u\w") + case True + have *: "subpath v u g = reversepath(subpath u v g) \ + subpath w u g = reversepath(subpath u w g) \ + subpath w v g = reversepath(subpath v w g)" + by (auto simp: reversepath_subpath) + have "u < v \ v < w \ + u < w \ w < v \ + v < u \ u < w \ + v < w \ w < u \ + w < u \ u < v \ + w < v \ v < u" + using True assms by linarith + with assms show ?thesis + using contour_integral_subpath_combine_less [of f g u v w] + contour_integral_subpath_combine_less [of f g u w v] + contour_integral_subpath_combine_less [of f g v u w] + contour_integral_subpath_combine_less [of f g v w u] + contour_integral_subpath_combine_less [of f g w u v] + contour_integral_subpath_combine_less [of f g w v u] + apply simp + apply (elim disjE) + apply (auto simp: * contour_integral_reversepath contour_integrable_subpath + valid_path_subpath algebra_simps) + done +next + case False + then show ?thesis + apply (auto) + using assms + by (metis eq_neg_iff_add_eq_0 contour_integral_reversepath reversepath_subpath valid_path_subpath) +qed + +lemma contour_integral_integral: + "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" + by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) + +lemma contour_integral_cong: + assumes "g = g'" "\x. x \ path_image g \ f x = f' x" + shows "contour_integral g f = contour_integral g' f'" + unfolding contour_integral_integral using assms + by (intro integral_cong) (auto simp: path_image_def) + + +text \Contour integral along a segment on the real axis\ + +lemma has_contour_integral_linepath_Reals_iff: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "(f has_contour_integral I) (linepath a b) \ + ((\x. f (of_real x)) has_integral I) {Re a..Re b}" +proof - + from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" + by (simp_all add: complex_eq_iff) + from assms have "a \ b" by auto + have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ + ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" + by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) + (insert assms, simp_all add: field_simps scaleR_conv_of_real) + also have "(\x. f (a + b * of_real x - a * of_real x)) = + (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" + using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) + also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ + ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms + by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) + also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def + by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) + finally show ?thesis by simp +qed + +lemma contour_integrable_linepath_Reals_iff: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "(f contour_integrable_on linepath a b) \ + (\x. f (of_real x)) integrable_on {Re a..Re b}" + using has_contour_integral_linepath_Reals_iff[OF assms, of f] + by (auto simp: contour_integrable_on_def integrable_on_def) + +lemma contour_integral_linepath_Reals_eq: + fixes a b :: complex and f :: "complex \ complex" + assumes "a \ Reals" "b \ Reals" "Re a < Re b" + shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" +proof (cases "f contour_integrable_on linepath a b") + case True + thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] + using has_contour_integral_integral has_contour_integral_unique by blast +next + case False + thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] + by (simp add: not_integrable_contour_integral not_integrable_integral) +qed + + + +text\Cauchy's theorem where there's a primitive\ + +lemma contour_integral_primitive_lemma: + fixes f :: "complex \ complex" and g :: "real \ complex" + assumes "a \ b" + and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" + shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) + has_integral (f(g b) - f(g a))) {a..b}" +proof - + obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" + using assms by (auto simp: piecewise_differentiable_on_def) + have cfg: "continuous_on {a..b} (\x. f (g x))" + apply (rule continuous_on_compose [OF cg, unfolded o_def]) + using assms + apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) + done + { fix x::real + assume a: "a < x" and b: "x < b" and xk: "x \ k" + then have "g differentiable at x within {a..b}" + using k by (simp add: differentiable_at_withinI) + then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" + by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) + then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" + by (simp add: has_vector_derivative_def scaleR_conv_of_real) + have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" + using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) + then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" + by (simp add: has_field_derivative_def) + have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" + using diff_chain_within [OF gdiff fdiff] + by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) + } note * = this + show ?thesis + apply (rule fundamental_theorem_of_calculus_interior_strong) + using k assms cfg * + apply (auto simp: at_within_Icc_at) + done +qed + +lemma contour_integral_primitive: + assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \ s" + shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" + using assms + apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) + apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) + done + +corollary Cauchy_theorem_primitive: + assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" + and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" + shows "(f' has_contour_integral 0) g" + using assms + by (metis diff_self contour_integral_primitive) + +text\Existence of path integral for continuous function\ +lemma contour_integrable_continuous_linepath: + assumes "continuous_on (closed_segment a b) f" + shows "f contour_integrable_on (linepath a b)" +proof - + have "continuous_on {0..1} ((\x. f x * (b - a)) \ linepath a b)" + apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) + apply (rule continuous_intros | simp add: assms)+ + done + then show ?thesis + apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) + apply (rule integrable_continuous [of 0 "1::real", simplified]) + apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) + apply (auto simp: vector_derivative_linepath_within) + done +qed + +lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" + by (rule has_derivative_imp_has_field_derivative) + (rule derivative_intros | simp)+ + +lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" + apply (rule contour_integral_unique) + using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] + apply (auto simp: field_simps has_field_der_id) + done + +lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" + by (simp add: contour_integrable_continuous_linepath) + +lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" + by (simp add: contour_integrable_continuous_linepath) + +subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ + +lemma has_contour_integral_neg: + "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" + by (simp add: has_integral_neg has_contour_integral_def) + +lemma has_contour_integral_add: + "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ + \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" + by (simp add: has_integral_add has_contour_integral_def algebra_simps) + +lemma has_contour_integral_diff: + "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ + \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" + by (simp add: has_integral_diff has_contour_integral_def algebra_simps) + +lemma has_contour_integral_lmul: + "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" +apply (simp add: has_contour_integral_def) +apply (drule has_integral_mult_right) +apply (simp add: algebra_simps) +done + +lemma has_contour_integral_rmul: + "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" +apply (drule has_contour_integral_lmul) +apply (simp add: mult.commute) +done + +lemma has_contour_integral_div: + "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" + by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) + +lemma has_contour_integral_eq: + "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" +apply (simp add: path_image_def has_contour_integral_def) +by (metis (no_types, lifting) image_eqI has_integral_eq) + +lemma has_contour_integral_bound_linepath: + assumes "(f has_contour_integral i) (linepath a b)" + "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" + shows "norm i \ B * norm(b - a)" +proof - + { fix x::real + assume x: "0 \ x" "x \ 1" + have "norm (f (linepath a b x)) * + norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" + by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) + } note * = this + have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" + apply (rule has_integral_bound + [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) + using assms * unfolding has_contour_integral_def + apply (auto simp: norm_mult) + done + then show ?thesis + by (auto simp: content_real) +qed + +(*UNUSED +lemma has_contour_integral_bound_linepath_strong: + fixes a :: real and f :: "complex \ real" + assumes "(f has_contour_integral i) (linepath a b)" + "finite k" + "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" + shows "norm i \ B*norm(b - a)" +*) + +lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" + unfolding has_contour_integral_linepath + by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) + +lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" + by (simp add: has_contour_integral_def) + +lemma has_contour_integral_is_0: + "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" + by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto + +lemma has_contour_integral_sum: + "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ + \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" + by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) + +subsection\<^marker>\tag unimportant\ \Operations on path integrals\ + +lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" + by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) + +lemma contour_integral_neg: + "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) + +lemma contour_integral_add: + "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = + contour_integral g f1 + contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) + +lemma contour_integral_diff: + "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = + contour_integral g f1 - contour_integral g f2" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) + +lemma contour_integral_lmul: + shows "f contour_integrable_on g + \ contour_integral g (\x. c * f x) = c*contour_integral g f" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) + +lemma contour_integral_rmul: + shows "f contour_integrable_on g + \ contour_integral g (\x. f x * c) = contour_integral g f * c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) + +lemma contour_integral_div: + shows "f contour_integrable_on g + \ contour_integral g (\x. f x / c) = contour_integral g f / c" + by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) + +lemma contour_integral_eq: + "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" + apply (simp add: contour_integral_def) + using has_contour_integral_eq + by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) + +lemma contour_integral_eq_0: + "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" + by (simp add: has_contour_integral_is_0 contour_integral_unique) + +lemma contour_integral_bound_linepath: + shows + "\f contour_integrable_on (linepath a b); + 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ + \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" + apply (rule has_contour_integral_bound_linepath [of f]) + apply (auto simp: has_contour_integral_integral) + done + +lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" + by (simp add: contour_integral_unique has_contour_integral_0) + +lemma contour_integral_sum: + "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ + \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" + by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral) + +lemma contour_integrable_eq: + "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_eq) + + +subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ + +lemma contour_integrable_neg: + "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" + using has_contour_integral_neg contour_integrable_on_def by blast + +lemma contour_integrable_add: + "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" + using has_contour_integral_add contour_integrable_on_def + by fastforce + +lemma contour_integrable_diff: + "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" + using has_contour_integral_diff contour_integrable_on_def + by fastforce + +lemma contour_integrable_lmul: + "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" + using has_contour_integral_lmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_rmul: + "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" + using has_contour_integral_rmul contour_integrable_on_def + by fastforce + +lemma contour_integrable_div: + "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" + using has_contour_integral_div contour_integrable_on_def + by fastforce + +lemma contour_integrable_sum: + "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ + \ (\x. sum (\a. f a x) s) contour_integrable_on p" + unfolding contour_integrable_on_def + by (metis has_contour_integral_sum) + + +subsection\<^marker>\tag unimportant\ \Reversing a path integral\ + +lemma has_contour_integral_reverse_linepath: + "(f has_contour_integral i) (linepath a b) + \ (f has_contour_integral (-i)) (linepath b a)" + using has_contour_integral_reversepath valid_path_linepath by fastforce + +lemma contour_integral_reverse_linepath: + "continuous_on (closed_segment a b) f + \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" +apply (rule contour_integral_unique) +apply (rule has_contour_integral_reverse_linepath) +by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) + + +(* Splitting a path integral in a flat way.*) + +lemma has_contour_integral_split: + assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "(f has_contour_integral (i + j)) (linepath a b)" +proof (cases "k = 0 \ k = 1") + case True + then show ?thesis + using assms by auto +next + case False + then have k: "0 < k" "k < 1" "complex_of_real k \ 1" + using assms by auto + have c': "c = k *\<^sub>R (b - a) + a" + by (metis diff_add_cancel c) + have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" + by (simp add: algebra_simps c') + { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" + have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" + using False apply (simp add: c' algebra_simps) + apply (simp add: real_vector.scale_left_distrib [symmetric] field_split_simps) + done + have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" + using k has_integral_affinity01 [OF *, of "inverse k" "0"] + apply (simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) + apply (auto dest: has_integral_cmul [where c = "inverse k"]) + done + } note fi = this + { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" + have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" + using k + apply (simp add: c' field_simps) + apply (simp add: scaleR_conv_of_real divide_simps) + apply (simp add: field_simps) + done + have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" + using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] + apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) + apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) + done + } note fj = this + show ?thesis + using f k + apply (simp add: has_contour_integral_linepath) + apply (simp add: linepath_def) + apply (rule has_integral_combine [OF _ _ fi fj], simp_all) + done +qed + +lemma continuous_on_closed_segment_transform: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "continuous_on (closed_segment a c) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + have "closed_segment a c \ closed_segment a b" + by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) + then show "continuous_on (closed_segment a c) f" + by (rule continuous_on_subset [OF f]) +qed + +lemma contour_integral_split: + assumes f: "continuous_on (closed_segment a b) f" + and k: "0 \ k" "k \ 1" + and c: "c - a = k *\<^sub>R (b - a)" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" +proof - + have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" + using c by (simp add: algebra_simps) + have "closed_segment a c \ closed_segment a b" + by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) + moreover have "closed_segment c b \ closed_segment a b" + by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) + ultimately + have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" + by (auto intro: continuous_on_subset [OF f]) + show ?thesis + by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) +qed + +lemma contour_integral_split_linepath: + assumes f: "continuous_on (closed_segment a b) f" + and c: "c \ closed_segment a b" + shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" + using c by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) + +text\The special case of midpoints used in the main quadrisection\ + +lemma has_contour_integral_midpoint: + assumes "(f has_contour_integral i) (linepath a (midpoint a b))" + "(f has_contour_integral j) (linepath (midpoint a b) b)" + shows "(f has_contour_integral (i + j)) (linepath a b)" + apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) + using assms + apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) + done + +lemma contour_integral_midpoint: + "continuous_on (closed_segment a b) f + \ contour_integral (linepath a b) f = + contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" + apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) + apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) + done + + +text\A couple of special case lemmas that are useful below\ + +lemma triangle_linear_has_chain_integral: + "((\x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" + apply (rule Cauchy_theorem_primitive [of UNIV "\x. m/2 * x^2 + d*x"]) + apply (auto intro!: derivative_eq_intros) + done + +lemma has_chain_integral_chain_integral3: + "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) + \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" + apply (subst contour_integral_unique [symmetric], assumption) + apply (drule has_contour_integral_integrable) + apply (simp add: valid_path_join) + done + +lemma has_chain_integral_chain_integral4: + "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) + \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" + apply (subst contour_integral_unique [symmetric], assumption) + apply (drule has_contour_integral_integrable) + apply (simp add: valid_path_join) + done + +subsection\Reversing the order in a double path integral\ + +text\The condition is stronger than needed but it's often true in typical situations\ + +lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" + by (auto simp: cbox_Pair_eq) + +lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" + by (auto simp: cbox_Pair_eq) + +proposition contour_integral_swap: + assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" + and vp: "valid_path g" "valid_path h" + and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" + and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" + shows "contour_integral g (\w. contour_integral h (f w)) = + contour_integral h (\z. contour_integral g (\w. f w z))" +proof - + have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" + using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) + have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) \ (\t. (g x, h t))" + by (rule ext) simp + have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) \ (\t. (g t, h x))" + by (rule ext) simp + have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" + by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) + have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" + by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) + have "\y. y \ {0..1} \ continuous_on {0..1} (\x. f (g x) (h y))" + by (subst fgh2) (rule fcon_im2 gcon continuous_intros | simp)+ + then have vdg: "\y. y \ {0..1} \ (\x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" + using continuous_on_mult gvcon integrable_continuous_real by blast + have "(\z. vector_derivative g (at (fst z))) = (\x. vector_derivative g (at x)) \ fst" + by auto + then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\x. vector_derivative g (at (fst x)))" + apply (rule ssubst) + apply (rule continuous_intros | simp add: gvcon)+ + done + have "(\z. vector_derivative h (at (snd z))) = (\x. vector_derivative h (at x)) \ snd" + by auto + then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" + apply (rule ssubst) + apply (rule continuous_intros | simp add: hvcon)+ + done + have "(\x. f (g (fst x)) (h (snd x))) = (\(y1,y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w))" + by auto + then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" + apply (rule ssubst) + apply (rule gcon hcon continuous_intros | simp)+ + apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) + done + have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = + integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" + proof (rule integral_cong [OF contour_integral_rmul [symmetric]]) + show "\x. x \ {0..1} \ f (g x) contour_integrable_on h" + unfolding contour_integrable_on + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF _ hvcon]) + apply (subst fgh1) + apply (rule fcon_im1 hcon continuous_intros | simp)+ + done + qed + also have "\ = integral {0..1} + (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" + unfolding contour_integral_integral + apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) + apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ + unfolding integral_mult_left [symmetric] + apply (simp only: mult_ac) + done + also have "\ = contour_integral h (\z. contour_integral g (\w. f w z))" + unfolding contour_integral_integral + apply (rule integral_cong) + unfolding integral_mult_left [symmetric] + apply (simp add: algebra_simps) + done + finally show ?thesis + by (simp add: contour_integral_integral) +qed + + +subsection\<^marker>\tag unimportant\ \The key quadrisection step\ + +lemma norm_sum_half: + assumes "norm(a + b) \ e" + shows "norm a \ e/2 \ norm b \ e/2" +proof - + have "e \ norm (- a - b)" + by (simp add: add.commute assms norm_minus_commute) + thus ?thesis + using norm_triangle_ineq4 order_trans by fastforce +qed + +lemma norm_sum_lemma: + assumes "e \ norm (a + b + c + d)" + shows "e / 4 \ norm a \ e / 4 \ norm b \ e / 4 \ norm c \ e / 4 \ norm d" +proof - + have "e \ norm ((a + b) + (c + d))" using assms + by (simp add: algebra_simps) + then show ?thesis + by (auto dest!: norm_sum_half) +qed + +lemma Cauchy_theorem_quadrisection: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and dist: "dist a b \ K" "dist b c \ K" "dist c a \ K" + and e: "e * K^2 \ + norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" + shows "\a' b' c'. + a' \ convex hull {a,b,c} \ b' \ convex hull {a,b,c} \ c' \ convex hull {a,b,c} \ + dist a' b' \ K/2 \ dist b' c' \ K/2 \ dist c' a' \ K/2 \ + e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" + (is "\x y z. ?\ x y z") +proof - + note divide_le_eq_numeral1 [simp del] + define a' where "a' = midpoint b c" + define b' where "b' = midpoint c a" + define c' where "c' = midpoint a b" + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + have fcont': "continuous_on (closed_segment c' b') f" + "continuous_on (closed_segment a' c') f" + "continuous_on (closed_segment b' a') f" + unfolding a'_def b'_def c'_def + by (rule continuous_on_subset [OF f], + metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ + let ?pathint = "\x y. contour_integral(linepath x y) f" + have *: "?pathint a b + ?pathint b c + ?pathint c a = + (?pathint a c' + ?pathint c' b' + ?pathint b' a) + + (?pathint a' c' + ?pathint c' b + ?pathint b a') + + (?pathint a' c + ?pathint c b' + ?pathint b' a') + + (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" + by (simp add: fcont' contour_integral_reverse_linepath) (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) + have [simp]: "\x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" + by (metis left_diff_distrib mult.commute norm_mult_numeral1) + have [simp]: "\x y. cmod (x - y) = cmod (y - x)" + by (simp add: norm_minus_commute) + consider "e * K\<^sup>2 / 4 \ cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | + "e * K\<^sup>2 / 4 \ cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | + "e * K\<^sup>2 / 4 \ cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | + "e * K\<^sup>2 / 4 \ cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" + using assms unfolding * by (blast intro: that dest!: norm_sum_lemma) + then show ?thesis + proof cases + case 1 then have "?\ a c' b'" + using assms + apply (clarsimp simp: c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) + done + then show ?thesis by blast + next + case 2 then have "?\ a' c' b" + using assms + apply (clarsimp simp: a'_def c'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) + done + then show ?thesis by blast + next + case 3 then have "?\ a' c b'" + using assms + apply (clarsimp simp: a'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) + done + then show ?thesis by blast + next + case 4 then have "?\ a' b' c'" + using assms + apply (clarsimp simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) + apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) + done + then show ?thesis by blast + qed +qed + +subsection\<^marker>\tag unimportant\ \Cauchy's theorem for triangles\ + +lemma triangle_points_closer: + fixes a::complex + shows "\x \ convex hull {a,b,c}; y \ convex hull {a,b,c}\ + \ norm(x - y) \ norm(a - b) \ + norm(x - y) \ norm(b - c) \ + norm(x - y) \ norm(c - a)" + using simplex_extremal_le [of "{a,b,c}"] + by (auto simp: norm_minus_commute) + +lemma holomorphic_point_small_triangle: + assumes x: "x \ S" + and f: "continuous_on S f" + and cd: "f field_differentiable (at x within S)" + and e: "0 < e" + shows "\k>0. \a b c. dist a b \ k \ dist b c \ k \ dist c a \ k \ + x \ convex hull {a,b,c} \ convex hull {a,b,c} \ S + \ norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + + contour_integral(linepath c a) f) + \ e*(dist a b + dist b c + dist c a)^2" + (is "\k>0. \a b c. _ \ ?normle a b c") +proof - + have le_of_3: "\a x y z. \0 \ x*y; 0 \ x*z; 0 \ y*z; a \ (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\ + \ a \ e*(x + y + z)^2" + by (simp add: algebra_simps power2_eq_square) + have disj_le: "\x \ a \ x \ b \ x \ c; 0 \ a; 0 \ b; 0 \ c\ \ x \ a + b + c" + for x::real and a b c + by linarith + have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" + if "convex hull {a, b, c} \ S" for a b c + using segments_subset_convex_hull that + by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ + note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] + { fix f' a b c d + assume d: "0 < d" + and f': "\y. \cmod (y - x) \ d; y \ S\ \ cmod (f y - f x - f' * (y - x)) \ e * cmod (y - x)" + and le: "cmod (a - b) \ d" "cmod (b - c) \ d" "cmod (c - a) \ d" + and xc: "x \ convex hull {a, b, c}" + and S: "convex hull {a, b, c} \ S" + have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = + contour_integral (linepath a b) (\y. f y - f x - f'*(y - x)) + + contour_integral (linepath b c) (\y. f y - f x - f'*(y - x)) + + contour_integral (linepath c a) (\y. f y - f x - f'*(y - x))" + apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF S]) + apply (simp add: field_simps) + done + { fix y + assume yc: "y \ convex hull {a,b,c}" + have "cmod (f y - f x - f' * (y - x)) \ e*norm(y - x)" + proof (rule f') + show "cmod (y - x) \ d" + by (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) + qed (use S yc in blast) + also have "\ \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" + by (simp add: yc e xc disj_le [OF triangle_points_closer]) + finally have "cmod (f y - f x - f' * (y - x)) \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . + } note cm_le = this + have "?normle a b c" + unfolding dist_norm pa + apply (rule le_of_3) + using f' xc S e + apply simp_all + apply (intro norm_triangle_le add_mono path_bound) + apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) + apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ + done + } note * = this + show ?thesis + using cd e + apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) + apply (clarify dest!: spec mp) + using * unfolding dist_norm + apply blast + done +qed + + +text\Hence the most basic theorem for a triangle.\ + +locale Chain = + fixes x0 At Follows + assumes At0: "At x0 0" + and AtSuc: "\x n. At x n \ \x'. At x' (Suc n) \ Follows x' x" +begin + primrec f where + "f 0 = x0" + | "f (Suc n) = (SOME x. At x (Suc n) \ Follows x (f n))" + + lemma At: "At (f n) n" + proof (induct n) + case 0 show ?case + by (simp add: At0) + next + case (Suc n) show ?case + by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) + qed + + lemma Follows: "Follows (f(Suc n)) (f n)" + by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) + + declare f.simps(2) [simp del] +end + +lemma Chain3: + assumes At0: "At x0 y0 z0 0" + and AtSuc: "\x y z n. At x y z n \ \x' y' z'. At x' y' z' (Suc n) \ Follows x' y' z' x y z" + obtains f g h where + "f 0 = x0" "g 0 = y0" "h 0 = z0" + "\n. At (f n) (g n) (h n) n" + "\n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" +proof - + interpret three: Chain "(x0,y0,z0)" "\(x,y,z). At x y z" "\(x',y',z'). \(x,y,z). Follows x' y' z' x y z" + apply unfold_locales + using At0 AtSuc by auto + show ?thesis + apply (rule that [of "\n. fst (three.f n)" "\n. fst (snd (three.f n))" "\n. snd (snd (three.f n))"]) + using three.At three.Follows + apply simp_all + apply (simp_all add: split_beta') + done +qed + +proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle: + assumes "f holomorphic_on (convex hull {a,b,c})" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +proof - + have contf: "continuous_on (convex hull {a,b,c}) f" + by (metis assms holomorphic_on_imp_continuous_on) + let ?pathint = "\x y. contour_integral(linepath x y) f" + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \ 0" + define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" + define e where "e = norm y / K^2" + have K1: "K \ 1" by (simp add: K_def max.coboundedI1) + then have K: "K > 0" by linarith + have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" + by (simp_all add: K_def) + have e: "e > 0" + unfolding e_def using ynz K1 by simp + define At where "At x y z n \ + convex hull {x,y,z} \ convex hull {a,b,c} \ + dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ + norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" + for x y z n + have At0: "At a b c 0" + using fy + by (simp add: At_def e_def has_chain_integral_chain_integral3) + { fix x y z n + assume At: "At x y z n" + then have contf': "continuous_on (convex hull {x,y,z}) f" + using contf At_def continuous_on_subset by metis + have "\x' y' z'. At x' y' z' (Suc n) \ convex hull {x',y',z'} \ convex hull {x,y,z}" + using At Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] + apply (simp add: At_def algebra_simps) + apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) + done + } note AtSuc = this + obtain fa fb fc + where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" + and cosb: "\n. convex hull {fa n, fb n, fc n} \ convex hull {a,b,c}" + and dist: "\n. dist (fa n) (fb n) \ K/2^n" + "\n. dist (fb n) (fc n) \ K/2^n" + "\n. dist (fc n) (fa n) \ K/2^n" + and no: "\n. norm(?pathint (fa n) (fb n) + + ?pathint (fb n) (fc n) + + ?pathint (fc n) (fa n)) \ e * (K/2^n)^2" + and conv_le: "\n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \ convex hull {fa n, fb n, fc n}" + apply (rule Chain3 [of At, OF At0 AtSuc]) + apply (auto simp: At_def) + done + obtain x where x: "\n. x \ convex hull {fa n, fb n, fc n}" + proof (rule bounded_closed_nest) + show "\n. closed (convex hull {fa n, fb n, fc n})" + by (simp add: compact_imp_closed finite_imp_compact_convex_hull) + show "\m n. m \ n \ convex hull {fa n, fb n, fc n} \ convex hull {fa m, fb m, fc m}" + by (erule transitive_stepwise_le) (auto simp: conv_le) + qed (fastforce intro: finite_imp_bounded_convex_hull)+ + then have xin: "x \ convex hull {a,b,c}" + using assms f0 by blast + then have fx: "f field_differentiable at x within (convex hull {a,b,c})" + using assms holomorphic_on_def by blast + { fix k n + assume k: "0 < k" + and le: + "\x' y' z'. + \dist x' y' \ k; dist y' z' \ k; dist z' x' \ k; + x \ convex hull {x',y',z'}; + convex hull {x',y',z'} \ convex hull {a,b,c}\ + \ + cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 + \ e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" + and Kk: "K / k < 2 ^ n" + have "K / 2 ^ n < k" using Kk k + by (auto simp: field_simps) + then have DD: "dist (fa n) (fb n) \ k" "dist (fb n) (fc n) \ k" "dist (fc n) (fa n) \ k" + using dist [of n] k + by linarith+ + have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 + \ (3 * K / 2 ^ n)\<^sup>2" + using dist [of n] e K + by (simp add: abs_le_square_iff [symmetric]) + have less10: "\x y::real. 0 < x \ y \ 9*x \ y < x*10" + by linarith + have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ e * (3 * K / 2 ^ n)\<^sup>2" + using ynz dle e mult_le_cancel_left_pos by blast + also have "\ < + cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" + using no [of n] e K + apply (simp add: e_def field_simps) + apply (simp only: zero_less_norm_iff [symmetric]) + done + finally have False + using le [OF DD x cosb] by auto + } then + have ?thesis + using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e + apply clarsimp + apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]], force+) + done + } + moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" + by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) + segments_subset_convex_hull(3) segments_subset_convex_hull(5)) + ultimately show ?thesis + using has_contour_integral_integral by fastforce +qed + +subsection\<^marker>\tag unimportant\ \Version needing function holomorphic in interior only\ + +lemma Cauchy_theorem_flat_lemma: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and c: "c - a = k *\<^sub>R (b - a)" + and k: "0 \ k" + shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof - + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + show ?thesis + proof (cases "k \ 1") + case True show ?thesis + by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) + next + case False then show ?thesis + using fabc c + apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) + apply (metis closed_segment_commute fabc(3)) + apply (auto simp: k contour_integral_reverse_linepath) + done + qed +qed + +lemma Cauchy_theorem_flat: + assumes f: "continuous_on (convex hull {a,b,c}) f" + and c: "c - a = k *\<^sub>R (b - a)" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "0 \ k") + case True with assms show ?thesis + by (blast intro: Cauchy_theorem_flat_lemma) +next + case False + have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using f continuous_on_subset segments_subset_convex_hull by metis+ + moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + + contour_integral (linepath c b) f = 0" + apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) + using False c + apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) + done + ultimately show ?thesis + apply (auto simp: contour_integral_reverse_linepath) + using add_eq_0_iff by force +qed + +lemma Cauchy_theorem_triangle_interior: + assumes contf: "continuous_on (convex hull {a,b,c}) f" + and holf: "f holomorphic_on interior (convex hull {a,b,c})" + shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" +proof - + have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" + using contf continuous_on_subset segments_subset_convex_hull by metis+ + have "bounded (f ` (convex hull {a,b,c}))" + by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) + then obtain B where "0 < B" and Bnf: "\x. x \ convex hull {a,b,c} \ norm (f x) \ B" + by (auto simp: dest!: bounded_pos [THEN iffD1]) + have "bounded (convex hull {a,b,c})" + by (simp add: bounded_convex_hull) + then obtain C where C: "0 < C" and Cno: "\y. y \ convex hull {a,b,c} \ norm y < C" + using bounded_pos_less by blast + then have diff_2C: "norm(x - y) \ 2*C" + if x: "x \ convex hull {a, b, c}" and y: "y \ convex hull {a, b, c}" for x y + proof - + have "cmod x \ C" + using x by (meson Cno not_le not_less_iff_gr_or_eq) + hence "cmod (x - y) \ C + C" + using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) + thus "cmod (x - y) \ 2 * C" + by (metis mult_2) + qed + have contf': "continuous_on (convex hull {b,a,c}) f" + using contf by (simp add: insert_commute) + { fix y::complex + assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" + and ynz: "y \ 0" + have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" + by (rule has_chain_integral_chain_integral3 [OF fy]) + have ?thesis + proof (cases "c=a \ a=b \ b=c") + case True then show ?thesis + using Cauchy_theorem_flat [OF contf, of 0] + using has_chain_integral_chain_integral3 [OF fy] ynz + by (force simp: fabc contour_integral_reverse_linepath) + next + case False + then have car3: "card {a, b, c} = Suc (DIM(complex))" + by auto + { assume "interior(convex hull {a,b,c}) = {}" + then have "collinear{a,b,c}" + using interior_convex_hull_eq_empty [OF car3] + by (simp add: collinear_3_eq_affine_dependent) + with False obtain d where "c \ a" "a \ b" "b \ c" "c - b = d *\<^sub>R (a - b)" + by (auto simp: collinear_3 collinear_lemma) + then have "False" + using False Cauchy_theorem_flat [OF contf'] pi_eq_y ynz + by (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) + } + then obtain d where d: "d \ interior (convex hull {a, b, c})" + by blast + { fix d1 + assume d1_pos: "0 < d1" + and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ + \ cmod (f x' - f x) < cmod y / (24 * C)" + define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" + define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x + let ?pathint = "\x y. contour_integral(linepath x y) f" + have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" + using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) + then have eCB: "24 * e * C * B \ cmod y" + using \C>0\ \B>0\ by (simp add: field_simps) + have e_le_d1: "e * (4 * C) \ d1" + using e \C>0\ by (simp add: field_simps) + have "shrink a \ interior(convex hull {a,b,c})" + "shrink b \ interior(convex hull {a,b,c})" + "shrink c \ interior(convex hull {a,b,c})" + using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) + then have fhp0: "(f has_contour_integral 0) + (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" + by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal) + then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" + by (simp add: has_chain_integral_chain_integral3) + have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" + "f contour_integrable_on linepath (shrink b) (shrink c)" + "f contour_integrable_on linepath (shrink c) (shrink a)" + using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) + have cmod_shr: "\x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" + using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) + have sh_eq: "\a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" + by (simp add: algebra_simps) + have "cmod y / (24 * C) \ cmod y / cmod (b - a) / 12" + using False \C>0\ diff_2C [of b a] ynz + by (auto simp: field_split_simps hull_inc) + have less_C: "\u \ convex hull {a, b, c}; 0 \ x; x \ 1\ \ x * cmod u < C" for x u + apply (cases "x=0", simp add: \0) + using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast + { fix u v + assume uv: "u \ convex hull {a, b, c}" "v \ convex hull {a, b, c}" "u\v" + and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" + have shr_uv: "shrink u \ interior(convex hull {a,b,c})" + "shrink v \ interior(convex hull {a,b,c})" + using d e uv + by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) + have cmod_fuv: "\x. 0\x \ x\1 \ cmod (f (linepath (shrink u) (shrink v) x)) \ B" + using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) + have By_uv: "B * (12 * (e * cmod (u - v))) \ cmod y" + apply (rule order_trans [OF _ eCB]) + using e \B>0\ diff_2C [of u v] uv + by (auto simp: field_simps) + { fix x::real assume x: "0\x" "x\1" + have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" + apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) + using uv x d interior_subset + apply (auto simp: hull_inc intro!: less_C) + done + have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" + by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) + have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" + apply (simp only: ll norm_mult scaleR_diff_right) + using \e>0\ cmod_less_4C apply (force intro: norm_triangle_lt less_le_trans [OF _ e_le_d1]) + done + have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" + using x uv shr_uv cmod_less_dt + by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) + also have "\ \ cmod y / cmod (v - u) / 12" + using False uv \C>0\ diff_2C [of v u] ynz + by (auto simp: field_split_simps hull_inc) + finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / cmod (v - u) / 12" + by simp + then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \ cmod y" + using uv False by (auto simp: field_simps) + have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) + \ B * (cmod y / 24 / C / B * 2 * C) + 2 * C * (cmod y / 24 / C)" + apply (rule add_mono [OF mult_mono]) + using By_uv e \0 < B\ \0 < C\ x apply (simp_all add: cmod_fuv cmod_shr cmod_12_le) + apply (simp add: field_simps) + done + also have "\ \ cmod y / 6" + by simp + finally have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) + \ cmod y / 6" . + } note cmod_diff_le = this + have f_uv: "continuous_on (closed_segment u v) f" + by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) + have **: "\f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" + by (simp add: algebra_simps) + have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) + \ (B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)) * content (cbox 0 (1::real))" + apply (rule has_integral_bound + [of _ "\x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" + _ 0 1]) + using ynz \0 < B\ \0 < C\ + apply (simp_all del: le_divide_eq_numeral1) + apply (simp add: has_integral_diff has_contour_integral_linepath [symmetric] has_contour_integral_integral + fpi_uv f_uv contour_integrable_continuous_linepath) + apply (auto simp: ** norm_triangle_le norm_mult cmod_diff_le simp del: le_divide_eq_numeral1) + done + also have "\ \ norm y / 6" + by simp + finally have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ norm y / 6" . + } note * = this + have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \ norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + moreover + have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \ norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + moreover + have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \ norm y / 6" + using False fpi_abc by (rule_tac *) (auto simp: hull_inc) + ultimately + have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + + (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) + \ norm y / 6 + norm y / 6 + norm y / 6" + by (metis norm_triangle_le add_mono) + also have "\ = norm y / 2" + by simp + finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - + (?pathint a b + ?pathint b c + ?pathint c a)) + \ norm y / 2" + by (simp add: algebra_simps) + then + have "norm(?pathint a b + ?pathint b c + ?pathint c a) \ norm y / 2" + by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) + then have "False" + using pi_eq_y ynz by auto + } + note * = this + have "uniformly_continuous_on (convex hull {a,b,c}) f" + by (simp add: contf compact_convex_hull compact_uniformly_continuous) + moreover have "norm y / (24 * C) > 0" + using ynz \C > 0\ by auto + ultimately obtain \ where "\ > 0" and + "\x\convex hull {a, b, c}. \x'\convex hull {a, b, c}. + dist x' x < \ \ dist (f x') (f x) < cmod y / (24 * C)" + using \C > 0\ ynz unfolding uniformly_continuous_on_def dist_norm by blast + hence False using *[of \] by (auto simp: 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 using \L > 0\ + by (intro exI[of _ L]) auto +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 blast + { fix e::real and w::complex + assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" + then have [simp]: "w \ path_image p" + using cbp p(2) \0 < pe\ + by (force simp: dist_norm norm_minus_commute path_image_def cball_def) + have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = + contour_integral p (\x. 1/(x - w) - 1/(x - z))" + by (simp add: p contour_integrable_inversediff contour_integral_diff) + { fix x + assume pe: "3/4 * pe < cmod (z - x)" + have "cmod (w - x) < pe/4 + cmod (z - x)" + by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) + then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp + have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" + using norm_diff_triangle_le by blast + also have "\ < pe/4 + cmod (w - x)" + using w by (simp add: norm_minus_commute) + finally have "pe/2 < cmod (w - x)" + using pe by auto + then have "(pe/2)^2 < cmod (w - x) ^ 2" + apply (rule power_strict_mono) + using \pe>0\ by auto + then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" + by (simp add: power_divide) + have "8 * L * cmod (w - z) < e * pe\<^sup>2" + using w \L>0\ by (simp add: field_simps) + also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" + using pe2 \e>0\ by (simp add: power2_eq_square) + also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" + using wx + apply (rule mult_strict_left_mono) + using pe2 e not_less_iff_gr_or_eq by fastforce + finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" + by simp + also have "\ \ e * cmod (w - x) * cmod (z - x)" + using e by simp + finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . + have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" + apply (cases "x=z \ x=w") + using pe \pe>0\ w \L>0\ + apply (force simp: norm_minus_commute) + using wx w(2) \L>0\ pe pe2 Lwz + apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) + done + } note L_cmod_le = this + have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" + apply (rule L) + using \pe>0\ w + apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) + using \pe>0\ w \L>0\ + apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) + done + have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" + apply simp + apply (rule le_less_trans [OF *]) + using \L>0\ e + apply (force simp: field_simps) + done + then have "cmod (winding_number p w - winding_number p z) < e" + using pi_ge_two e + by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) + } note cmod_wn_diff = this + then have "isCont (winding_number p) z" + apply (simp add: continuous_at_eps_delta, clarify) + apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) + using \pe>0\ \L>0\ + apply (simp add: dist_norm cmod_wn_diff) + done + then show ?thesis + apply (rule continuous_transform_within [where d = "min d e / 2"]) + apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) + done +qed + +corollary continuous_on_winding_number: + "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" + by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) + +subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ + +lemma winding_number_constant: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" + shows "winding_number \ constant_on S" +proof - + have *: "1 \ cmod (winding_number \ y - winding_number \ z)" + if ne: "winding_number \ y \ winding_number \ z" and "y \ S" "z \ S" for y z + proof - + have "winding_number \ y \ \" "winding_number \ z \ \" + using that integer_winding_number [OF \ loop] sg \y \ S\ by auto + with ne show ?thesis + by (auto simp: Ints_def simp flip: of_int_diff) + qed + have cont: "continuous_on S (\w. winding_number \ w)" + using continuous_on_winding_number [OF \] sg + by (meson continuous_on_subset disjoint_eq_subset_Compl) + show ?thesis + using "*" zero_less_one + by (blast intro: continuous_discrete_range_constant [OF cs cont]) +qed + +lemma winding_number_eq: + "\path \; pathfinish \ = pathstart \; w \ S; z \ S; connected S; S \ path_image \ = {}\ + \ winding_number \ w = winding_number \ z" + using winding_number_constant by (metis constant_on_def) + +lemma open_winding_number_levelsets: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + shows "open {z. z \ path_image \ \ winding_number \ z = k}" +proof - + have opn: "open (- path_image \)" + by (simp add: closed_path_image \ open_Compl) + { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" + obtain e where e: "e>0" "ball z e \ - path_image \" + using open_contains_ball [of "- path_image \"] opn z + by blast + have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" + apply (rule_tac x=e in exI) + using e apply (simp add: dist_norm ball_def norm_minus_commute) + apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where S = "ball z e"]) + done + } then + show ?thesis + by (auto simp: open_dist) +qed + +subsection\Winding number is zero "outside" a curve\ + +proposition winding_number_zero_in_outside: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" + shows "winding_number \ z = 0" +proof - + obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto + obtain w::complex where w: "w \ ball 0 (B + 1)" + by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) + have "- ball 0 (B + 1) \ outside (path_image \)" + apply (rule outside_subset_convex) + using B subset_ball by auto + then have wout: "w \ outside (path_image \)" + using w by blast + moreover have "winding_number \ constant_on outside (path_image \)" + using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside + by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) + ultimately have "winding_number \ z = winding_number \ w" + by (metis (no_types, hide_lams) constant_on_def z) + also have "\ = 0" + proof - + have wnot: "w \ path_image \" using wout by (simp add: outside_def) + { fix e::real assume "0" "pathfinish p = pathfinish \" + and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" + and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" + using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force + have pip: "path_image p \ ball 0 (B + 1)" + using B + apply (clarsimp simp add: path_image_def dist_norm ball_def) + apply (frule (1) pg1) + apply (fastforce dest: norm_add_less) + done + then have "w \ path_image p" using w by blast + then have "\p. valid_path p \ w \ path_image p \ + pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ + (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" + apply (rule_tac x=p in exI) + apply (simp add: p valid_path_polynomial_function) + apply (intro conjI) + using pge apply (simp add: norm_minus_commute) + apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) + apply (rule holomorphic_intros | simp add: dist_norm)+ + using mem_ball_0 w apply blast + using p apply (simp_all add: valid_path_polynomial_function loop pip) + done + } + then show ?thesis + by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) + qed + finally show ?thesis . +qed + +corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" + by (rule winding_number_zero_in_outside) + (auto simp: pathfinish_def pathstart_def path_polynomial_function) + +corollary\<^marker>\tag unimportant\ winding_number_zero_outside: + "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" + by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) + +lemma winding_number_zero_at_infinity: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + shows "\B. \z. B \ norm z \ winding_number \ z = 0" +proof - + obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" + using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto + then show ?thesis + apply (rule_tac x="B+1" in exI, clarify) + apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) + apply (meson less_add_one mem_cball_0 not_le order_trans) + using ball_subset_cball by blast +qed + +lemma winding_number_zero_point: + "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ + \ \z. z \ s \ winding_number \ z = 0" + using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside + by (fastforce simp add: compact_path_image) + + +text\If a path winds round a set, it winds rounds its inside.\ +lemma winding_number_around_inside: + assumes \: "path \" and loop: "pathfinish \ = pathstart \" + and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" + and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" + shows "winding_number \ w = winding_number \ z" +proof - + have ssb: "s \ inside(path_image \)" + proof + fix x :: complex + assume "x \ s" + hence "x \ path_image \" + by (meson disjoint_iff_not_equal s_disj) + thus "x \ inside (path_image \)" + using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) +qed + show ?thesis + apply (rule winding_number_eq [OF \ loop w]) + using z apply blast + apply (simp add: cls connected_with_inside cos) + apply (simp add: Int_Un_distrib2 s_disj, safe) + by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) + qed + + +text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ +lemma winding_number_subpath_continuous: + assumes \: "valid_path \" and z: "z \ path_image \" + shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" +proof - + have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = + winding_number (subpath 0 x \) z" + if x: "0 \ x" "x \ 1" for x + proof - + have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = + 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" + using assms x + apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) + done + also have "\ = winding_number (subpath 0 x \) z" + apply (subst winding_number_valid_path) + using assms x + apply (simp_all add: path_image_subpath valid_path_subpath) + by (force simp: path_image_def) + finally show ?thesis . + qed + show ?thesis + apply (rule continuous_on_eq + [where f = "\x. 1 / (2*pi*\) * + integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) + apply (rule continuous_intros)+ + apply (rule indefinite_integral_continuous_1) + apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) + using assms + apply (simp add: *) + done +qed + +lemma winding_number_ivt_pos: + assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" + shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" + apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) + apply (rule winding_number_subpath_continuous [OF \ z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_neg: + assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" + shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" + apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) + apply (rule winding_number_subpath_continuous [OF \ z]) + using assms + apply (auto simp: path_image_def image_def) + done + +lemma winding_number_ivt_abs: + assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" + shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" + using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] + by force + +lemma winding_number_lt_half_lemma: + assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" + shows "Re(winding_number \ z) < 1/2" +proof - + { assume "Re(winding_number \ z) \ 1/2" + then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" + using winding_number_ivt_pos [OF \ z, of "1/2"] by auto + have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" + using winding_number_exp_2pi [of "subpath 0 t \" z] + apply (simp add: t \ valid_path_imp_path) + using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) + have "b < a \ \ 0" + proof - + have "\ 0 \ {c. b < a \ c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) + thus ?thesis + by blast + qed + moreover have "b < a \ \ t" + proof - + have "\ t \ {c. b < a \ c}" + by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) + thus ?thesis + by blast + qed + ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az + by (simp add: inner_diff_right)+ + then have False + by (simp add: gt inner_mult_right mult_less_0_iff) + } + then show ?thesis by force +qed + +lemma winding_number_lt_half: + assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" + shows "\Re (winding_number \ z)\ < 1/2" +proof - + have "z \ path_image \" using assms by auto + with assms show ?thesis + apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) + apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] + winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) + done +qed + +lemma winding_number_le_half: + assumes \: "valid_path \" and z: "z \ path_image \" + and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" + shows "\Re (winding_number \ z)\ \ 1/2" +proof - + { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" + have "isCont (winding_number \) z" + by (metis continuous_at_winding_number valid_path_imp_path \ z) + then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" + using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast + define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" + have *: "a \ z' \ b - d / 3 * cmod a" + unfolding z'_def inner_mult_right' divide_inverse + apply (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square anz) + apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) + done + have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" + using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) + then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" + by simp + then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" + using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp + then have wnz_12': "\Re (winding_number \ z')\ > 1/2" + by linarith + moreover have "\Re (winding_number \ z')\ < 1/2" + apply (rule winding_number_lt_half [OF \ *]) + using azb \d>0\ pag + apply (auto simp: add_strict_increasing anz field_split_simps dest!: subsetD) + done + ultimately have False + by simp + } + then show ?thesis by force +qed + +lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" + using separating_hyperplane_closed_point [of "closed_segment a b" z] + apply auto + apply (simp add: closed_segment_def) + apply (drule less_imp_le) + apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) + apply (auto simp: segment) + done + + +text\ Positivity of WN for a linepath.\ +lemma winding_number_linepath_pos_lt: + assumes "0 < Im ((b - a) * cnj (b - z))" + shows "0 < Re(winding_number(linepath a b) z)" +proof - + have z: "z \ path_image (linepath a b)" + using assms + by (simp add: closed_segment_def) (force simp: algebra_simps) + show ?thesis + apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) + apply (simp add: linepath_def algebra_simps) + done +qed + + +subsection\Cauchy's integral formula, again for a convex enclosing set\ + +lemma Cauchy_integral_formula_weak: + assumes s: "convex s" and "finite k" and conf: "continuous_on s f" + and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" + and z: "z \ interior s - k" and vpg: "valid_path \" + and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + obtain f' where f': "(f has_field_derivative f') (at z)" + using fcd [OF z] by (auto simp: field_differentiable_def) + have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ + have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x + proof (cases "x = z") + case True then show ?thesis + apply (simp add: continuous_within) + apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) + using has_field_derivative_at_within has_field_derivative_iff f' + apply (fastforce simp add:)+ + done + next + case False + then have dxz: "dist x z > 0" by auto + have cf: "continuous (at x within s) f" + using conf continuous_on_eq_continuous_within that by blast + have "continuous (at x within s) (\w. (f w - f z) / (w - z))" + by (rule cf continuous_intros | simp add: False)+ + then show ?thesis + apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) + apply (force simp: dist_commute) + done + qed + have fink': "finite (insert z k)" using \finite k\ by blast + have *: "((\w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \" + apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) + using c apply (force simp: continuous_on_eq_continuous_within) + apply (rename_tac w) + apply (rule_tac d="dist w z" and f = "\w. (f w - f z)/(w - z)" in field_differentiable_transform_within) + apply (simp_all add: dist_pos_lt dist_commute) + apply (metis less_irrefl) + apply (rule derivative_intros fcd | simp)+ + done + show ?thesis + apply (rule has_contour_integral_eq) + using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] + apply (auto simp: ac_simps divide_simps) + done +qed + +theorem Cauchy_integral_formula_convex_simple: + "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; + pathfinish \ = pathstart \\ + \ ((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" + apply (rule Cauchy_integral_formula_weak [where k = "{}"]) + using holomorphic_on_imp_continuous_on + by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) + +subsection\Homotopy forms of Cauchy's theorem\ + +lemma Cauchy_theorem_homotopic: + assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" +proof - + have pathsf: "linked_paths atends g h" + using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) + obtain k :: "real \ real \ complex" + where contk: "continuous_on ({0..1} \ {0..1}) k" + and ks: "k ` ({0..1} \ {0..1}) \ s" + and k [simp]: "\x. k (0, x) = g x" "\x. k (1, x) = h x" + and ksf: "\t\{0..1}. linked_paths atends g (\x. k (t, x))" + using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) + have ucontk: "uniformly_continuous_on ({0..1} \ {0..1}) k" + by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) + { fix t::real assume t: "t \ {0..1}" + have pak: "path (k \ (\u. (t, u)))" + unfolding path_def + apply (rule continuous_intros continuous_on_subset [OF contk])+ + using t by force + have pik: "path_image (k \ Pair t) \ s" + using ks t by (auto simp: path_image_def) + obtain e where "e>0" and e: + "\g h. \valid_path g; valid_path h; + \u\{0..1}. cmod (g u - (k \ Pair t) u) < e \ cmod (h u - (k \ Pair t) u) < e; + linked_paths atends g h\ + \ contour_integral h f = contour_integral g f" + using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis + obtain d where "d>0" and d: + "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" + by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) + { fix t1 t2 + assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" + have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" + using \e > 0\ + apply (rule_tac y = k1 in norm_triangle_half_l) + apply (auto simp: norm_minus_commute intro: order_less_trans) + done + have "\d>0. \g1 g2. valid_path g1 \ valid_path g2 \ + (\u\{0..1}. cmod (g1 u - k (t1, u)) < d \ cmod (g2 u - k (t2, u)) < d) \ + linked_paths atends g1 g2 \ + contour_integral g2 f = contour_integral g1 f" + apply (rule_tac x="e/4" in exI) + using t t1 t2 ltd \e > 0\ + apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) + done + } + then have "\e. 0 < e \ + (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < e \ \t2 - t\ < e + \ (\d. 0 < d \ + (\g1 g2. valid_path g1 \ valid_path g2 \ + (\u \ {0..1}. + norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ + linked_paths atends g1 g2 + \ contour_integral g2 f = contour_integral g1 f)))" + by (rule_tac x=d in exI) (simp add: \d > 0\) + } + then obtain ee where ee: + "\t. t \ {0..1} \ ee t > 0 \ + (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < ee t \ \t2 - t\ < ee t + \ (\d. 0 < d \ + (\g1 g2. valid_path g1 \ valid_path g2 \ + (\u \ {0..1}. + norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ + linked_paths atends g1 g2 + \ contour_integral g2 f = contour_integral g1 f)))" + by metis + note ee_rule = ee [THEN conjunct2, rule_format] + define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" + obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" + proof (rule compactE [OF compact_interval]) + show "{0..1} \ \C" + using ee [THEN conjunct1] by (auto simp: C_def dist_norm) + qed (use C_def in auto) + define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" + have kk01: "kk \ {0..1}" by (auto simp: kk_def) + define e where "e = Min (ee ` kk)" + have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" + using C' by (auto simp: kk_def C_def) + have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" + by (simp add: kk_def ee) + moreover have "finite kk" + using \finite C'\ kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) + moreover have "kk \ {}" using \{0..1} \ \C'\ C'_eq by force + ultimately have "e > 0" + using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) + then obtain N::nat where "N > 0" and N: "1/N < e/3" + by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) + have e_le_ee: "\i. i \ kk \ e \ ee i" + using \finite kk\ by (simp add: e_def Min_le_iff [of "ee ` kk"]) + have plus: "\t \ kk. x \ ball t (ee t / 3)" if "x \ {0..1}" for x + using C' subsetD [OF C'01 that] unfolding C'_eq by blast + have [OF order_refl]: + "\d. 0 < d \ (\j. valid_path j \ (\u \ {0..1}. norm(j u - k (n/N, u)) < d) \ linked_paths atends g j + \ contour_integral j f = contour_integral g f)" + if "n \ N" for n + using that + proof (induct n) + case 0 show ?case using ee_rule [of 0 0 0] + apply clarsimp + apply (rule_tac x=d in exI, safe) + by (metis diff_self vpg norm_zero) + next + case (Suc n) + then have N01: "n/N \ {0..1}" "(Suc n)/N \ {0..1}" by auto + then obtain t where t: "t \ kk" "n/N \ ball t (ee t / 3)" + using plus [of "n/N"] by blast + then have nN_less: "\n/N - t\ < ee t" + by (simp add: dist_norm del: less_divide_eq_numeral1) + have n'N_less: "\real (Suc n) / real N - t\ < ee t" + using t N \N > 0\ e_le_ee [of t] + by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) + have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast + obtain d1 where "d1 > 0" and d1: + "\g1 g2. \valid_path g1; valid_path g2; + \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; + linked_paths atends g1 g2\ + \ contour_integral g2 f = contour_integral g1 f" + using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce + have "n \ N" using Suc.prems by auto + with Suc.hyps + obtain d2 where "d2 > 0" + and d2: "\j. \valid_path j; \u\{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\ + \ contour_integral j f = contour_integral g f" + by auto + have "continuous_on {0..1} (k \ (\u. (n/N, u)))" + apply (rule continuous_intros continuous_on_subset [OF contk])+ + using N01 by auto + then have pkn: "path (\u. k (n/N, u))" + by (simp add: path_def) + have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) + obtain p where "polynomial_function p" + and psf: "pathstart p = pathstart (\u. k (n/N, u))" + "pathfinish p = pathfinish (\u. k (n/N, u))" + and pk_le: "\t. t\{0..1} \ cmod (p t - k (n/N, t)) < min d1 d2" + using path_approx_polynomial_function [OF pkn min12] by blast + then have vpp: "valid_path p" using valid_path_polynomial_function by blast + have lpa: "linked_paths atends g p" + by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) + show ?case + proof (intro exI; safe) + fix j + assume "valid_path j" "linked_paths atends g j" + and "\u\{0..1}. cmod (j u - k (real (Suc n) / real N, u)) < min d1 d2" + then have "contour_integral j f = contour_integral p f" + using pk_le N01(1) ksf by (force intro!: vpp d1 simp add: linked_paths_def psf) + also have "... = contour_integral g f" + using pk_le by (force intro!: vpp d2 lpa) + finally show "contour_integral j f = contour_integral g f" . + qed (simp add: \0 < d1\ \0 < d2\) + qed + then obtain d where "0 < d" + "\j. valid_path j \ (\u \ {0..1}. norm(j u - k (1,u)) < d) \ linked_paths atends g j + \ contour_integral j f = contour_integral g f" + using \N>0\ by auto + then have "linked_paths atends g h \ contour_integral h f = contour_integral g f" + using \N>0\ vph by fastforce + then show ?thesis + by (simp add: pathsf) +qed + +proposition Cauchy_theorem_homotopic_paths: + assumes hom: "homotopic_paths s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" + using Cauchy_theorem_homotopic [of True s g h] assms by simp + +proposition Cauchy_theorem_homotopic_loops: + assumes hom: "homotopic_loops s g h" + and "open s" and f: "f holomorphic_on s" + and vpg: "valid_path g" and vph: "valid_path h" + shows "contour_integral g f = contour_integral h f" + using Cauchy_theorem_homotopic [of False s g h] assms by simp + +lemma has_contour_integral_newpath: + "\(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\ + \ (f has_contour_integral y) g" + using has_contour_integral_integral contour_integral_unique by auto + +lemma Cauchy_theorem_null_homotopic: + "\f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\ \ (f has_contour_integral 0) g" + apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) + using contour_integrable_holomorphic_simple + apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) + by (simp add: Cauchy_theorem_homotopic_loops) + +subsection\<^marker>\tag unimportant\ \More winding number properties\ + +text\including the fact that it's +-1 inside a simple closed curve.\ + +lemma winding_number_homotopic_paths: + assumes "homotopic_paths (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto + moreover have pag: "z \ path_image g" and pah: "z \ path_image h" + using homotopic_paths_imp_subset [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ + \ homotopic_paths (-{z}) g p" + and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ + \ homotopic_paths (-{z}) h q" + using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \ path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\t\{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" + using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast + obtain q where q: + "valid_path q" "z \ path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\t\{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" + using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast + have "homotopic_paths (- {z}) g p" + by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) + moreover have "homotopic_paths (- {z}) h q" + by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) + ultimately have "homotopic_paths (- {z}) p q" + by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) + then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" + by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_homotopic_loops: + assumes "homotopic_loops (-{z}) g h" + shows "winding_number g z = winding_number h z" +proof - + have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto + moreover have pag: "z \ path_image g" and pah: "z \ path_image h" + using homotopic_loops_imp_subset [OF assms] by auto + moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" + using homotopic_loops_imp_loop [OF assms] by auto + ultimately obtain d e where "d > 0" "e > 0" + and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ + \ homotopic_loops (-{z}) g p" + and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ + \ homotopic_loops (-{z}) h q" + using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force + obtain p where p: + "valid_path p" "z \ path_image p" + "pathstart p = pathstart g" "pathfinish p = pathfinish g" + and gp_less:"\t\{0..1}. cmod (g t - p t) < d" + and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" + using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast + obtain q where q: + "valid_path q" "z \ path_image q" + "pathstart q = pathstart h" "pathfinish q = pathfinish h" + and hq_less: "\t\{0..1}. cmod (h t - q t) < e" + and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" + using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast + have gp: "homotopic_loops (- {z}) g p" + by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) + have hq: "homotopic_loops (- {z}) h q" + by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) + have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" + proof (rule Cauchy_theorem_homotopic_loops) + show "homotopic_loops (- {z}) p q" + by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) + qed (auto intro!: holomorphic_intros simp: p q) + then show ?thesis + by (simp add: pap paq) +qed + +lemma winding_number_paths_linear_eq: + "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; + \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ + \ winding_number h z = winding_number g z" + by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) + +lemma winding_number_loops_linear_eq: + "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; + \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ + \ winding_number h z = winding_number g z" + by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) + +lemma winding_number_nearby_paths_eq: + "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; + \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ + \ winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) + +lemma winding_number_nearby_loops_eq: + "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; + \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ + \ winding_number h z = winding_number g z" + by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) + + +lemma winding_number_subpath_combine: + "\path g; z \ path_image g; + u \ {0..1}; v \ {0..1}; w \ {0..1}\ + \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = + winding_number (subpath u w g) z" +apply (rule trans [OF winding_number_join [THEN sym] + winding_number_homotopic_paths [OF homotopic_join_subpaths]]) + using path_image_subpath_subset by auto + +subsection\Partial circle path\ + +definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" + where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" + +lemma pathstart_part_circlepath [simp]: + "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" +by (metis part_circlepath_def pathstart_def pathstart_linepath) + +lemma pathfinish_part_circlepath [simp]: + "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" +by (metis part_circlepath_def pathfinish_def pathfinish_linepath) + +lemma reversepath_part_circlepath[simp]: + "reversepath (part_circlepath z r s t) = part_circlepath z r t s" + unfolding part_circlepath_def reversepath_def linepath_def + by (auto simp:algebra_simps) + +lemma has_vector_derivative_part_circlepath [derivative_intros]: + "((part_circlepath z r s t) has_vector_derivative + (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) + (at x within X)" + apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) + apply (rule has_vector_derivative_real_field) + apply (rule derivative_eq_intros | simp)+ + done + +lemma differentiable_part_circlepath: + "part_circlepath c r a b differentiable at x within A" + using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast + +lemma vector_derivative_part_circlepath: + "vector_derivative (part_circlepath z r s t) (at x) = + \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" + using has_vector_derivative_part_circlepath vector_derivative_at by blast + +lemma vector_derivative_part_circlepath01: + "\0 \ x; x \ 1\ + \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = + \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" + using has_vector_derivative_part_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" + apply (simp add: valid_path_def) + apply (rule C1_differentiable_imp_piecewise) + apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath + intro!: continuous_intros) + done + +lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" + by (simp add: valid_path_imp_path) + +proposition path_image_part_circlepath: + assumes "s \ t" + shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" +proof - + { fix z::real + assume "0 \ z" "z \ 1" + with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" + apply (rule_tac x="(1 - z) * s + z * t" in exI) + apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) + apply (rule conjI) + using mult_right_mono apply blast + using affine_ineq by (metis "mult.commute") + } + moreover + { fix z + assume "s \ z" "z \ t" + then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" + apply (rule_tac x="(z - s)/(t - s)" in image_eqI) + apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) + apply (auto simp: field_split_simps) + done + } + ultimately show ?thesis + by (fastforce simp add: path_image_def part_circlepath_def) +qed + +lemma path_image_part_circlepath': + "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" +proof - + have "path_image (part_circlepath z r s t) = + (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" + by (simp add: image_image path_image_def part_circlepath_def) + also have "linepath s t ` {0..1} = closed_segment s t" + by (rule linepath_image_01) + finally show ?thesis by (simp add: cis_conv_exp) +qed + +lemma path_image_part_circlepath_subset: + "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" +by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) + +lemma in_path_image_part_circlepath: + assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" + shows "norm(w - z) = r" +proof - + have "w \ {c. dist z c = r}" + by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) + thus ?thesis + by (simp add: dist_norm norm_minus_commute) +qed + +lemma path_image_part_circlepath_subset': + assumes "r \ 0" + shows "path_image (part_circlepath z r s t) \ sphere z r" +proof (cases "s \ t") + case True + thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp +next + case False + thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms + by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all +qed + +lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" + by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) + +lemma contour_integral_bound_part_circlepath: + assumes "f contour_integrable_on part_circlepath c r a b" + assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" + shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" +proof - + let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * + exp (\ * linepath a b x))" + have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" + proof (rule integral_norm_bound_integral, goal_cases) + case 1 + with assms(1) show ?case + by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) + next + case (3 x) + with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult + by (intro mult_mono) (auto simp: path_image_def) + qed auto + also have "?I = contour_integral (part_circlepath c r a b) f" + by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) + finally show ?thesis by simp +qed + +lemma has_contour_integral_part_circlepath_iff: + assumes "a < b" + shows "(f has_contour_integral I) (part_circlepath c r a b) \ + ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" +proof - + have "(f has_contour_integral I) (part_circlepath c r a b) \ + ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) + (at x within {0..1})) has_integral I) {0..1}" + unfolding has_contour_integral_def .. + also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * + cis (linepath a b x)) has_integral I) {0..1}" + by (intro has_integral_cong, subst vector_derivative_part_circlepath01) + (simp_all add: cis_conv_exp) + also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * + r * \ * exp (\ * linepath (of_real a) (of_real b) x) * + vector_derivative (linepath (of_real a) (of_real b)) + (at x within {0..1})) has_integral I) {0..1}" + by (intro has_integral_cong, subst vector_derivative_linepath_within) + (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) + also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) + (linepath (of_real a) (of_real b))" + by (simp add: has_contour_integral_def) + also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms + by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) + finally show ?thesis . +qed + +lemma contour_integrable_part_circlepath_iff: + assumes "a < b" + shows "f contour_integrable_on (part_circlepath c r a b) \ + (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (auto simp: contour_integrable_on_def integrable_on_def + has_contour_integral_part_circlepath_iff) + +lemma contour_integral_part_circlepath_eq: + assumes "a < b" + shows "contour_integral (part_circlepath c r a b) f = + integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" +proof (cases "f contour_integrable_on part_circlepath c r a b") + case True + hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (simp add: contour_integrable_part_circlepath_iff) + with True show ?thesis + using has_contour_integral_part_circlepath_iff[OF assms] + contour_integral_unique has_integral_integrable_integral by blast +next + case False + hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" + using assms by (simp add: contour_integrable_part_circlepath_iff) + with False show ?thesis + by (simp add: not_integrable_contour_integral not_integrable_integral) +qed + +lemma contour_integral_part_circlepath_reverse: + "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" + by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all + +lemma contour_integral_part_circlepath_reverse': + "b < a \ contour_integral (part_circlepath c r a b) f = + -contour_integral (part_circlepath c r b a) f" + by (rule contour_integral_part_circlepath_reverse) + +lemma finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" +proof (cases "w = 0") + case True then show ?thesis by auto +next + case False + have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" + apply (simp add: norm_mult finite_int_iff_bounded_le) + apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) + apply (auto simp: field_split_simps le_floor_iff) + done + have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" + by blast + show ?thesis + apply (subst exp_Ln [OF False, symmetric]) + apply (simp add: exp_eq) + using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) + done +qed + +lemma finite_bounded_log2: + fixes a::complex + assumes "a \ 0" + shows "finite {z. norm z \ b \ exp(a*z) = w}" +proof - + have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" + by (rule finite_imageI [OF finite_bounded_log]) + show ?thesis + by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) +qed + +lemma has_contour_integral_bound_part_circlepath_strong: + assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" + and "finite k" and le: "0 \ B" "0 < r" "s \ t" + and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" + shows "cmod i \ B * r * (t - s)" +proof - + consider "s = t" | "s < t" using \s \ t\ by linarith + then show ?thesis + proof cases + case 1 with fi [unfolded has_contour_integral] + have "i = 0" by (simp add: vector_derivative_part_circlepath) + with assms show ?thesis by simp + next + case 2 + have [simp]: "\r\ = r" using \r > 0\ by linarith + have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" + by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) + have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y + proof - + define w where "w = (y - z)/of_real r / exp(\ * of_real s)" + have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" + apply (rule finite_vimageI [OF finite_bounded_log2]) + using \s < t\ apply (auto simp: inj_of_real) + done + show ?thesis + apply (simp add: part_circlepath_def linepath_def vimage_def) + apply (rule finite_subset [OF _ fin]) + using le + apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) + done + qed + then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" + by (rule finite_finite_vimage_IntI [OF \finite k\]) + have **: "((\x. if (part_circlepath z r s t x) \ k then 0 + else f(part_circlepath z r s t x) * + vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" + by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) + have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" + by (auto intro!: B [unfolded path_image_def image_def, simplified]) + show ?thesis + apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) + using assms apply force + apply (simp add: norm_mult vector_derivative_part_circlepath) + using le * "2" \r > 0\ by auto + qed +qed + +lemma has_contour_integral_bound_part_circlepath: + "\(f has_contour_integral i) (part_circlepath z r s t); + 0 \ B; 0 < r; s \ t; + \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ + \ norm i \ B*r*(t - s)" + by (auto intro: has_contour_integral_bound_part_circlepath_strong) + +lemma contour_integrable_continuous_part_circlepath: + "continuous_on (path_image (part_circlepath z r s t)) f + \ f contour_integrable_on (part_circlepath z r s t)" + apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) + apply (rule integrable_continuous_real) + apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) + done + +proposition winding_number_part_circlepath_pos_less: + assumes "s < t" and no: "norm(w - z) < r" + shows "0 < Re (winding_number(part_circlepath z r s t) w)" +proof - + have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) + note valid_path_part_circlepath + moreover have " w \ path_image (part_circlepath z r s t)" + using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) + moreover have "0 < r * (t - s) * (r - cmod (w - z))" + using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) + ultimately show ?thesis + apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) + apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) + apply (rule mult_left_mono)+ + using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] + apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) + using assms \0 < r\ by auto +qed + +lemma simple_path_part_circlepath: + "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" +proof (cases "r = 0 \ s = t") + case True + then show ?thesis + unfolding part_circlepath_def simple_path_def + by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ +next + case False then have "r \ 0" "s \ t" by auto + have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" + by (simp add: algebra_simps) + have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 + \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" + by auto + have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ + (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" + by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] + intro: exI [where x = "-n" for n]) + have 1: "\s - t\ \ 2 * pi" + if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" + proof (rule ccontr) + assume "\ \s - t\ \ 2 * pi" + then have *: "\n. t - s \ of_int n * \s - t\" + using False that [of "2*pi / \t - s\"] + by (simp add: abs_minus_commute divide_simps) + show False + using * [of 1] * [of "-1"] by auto + qed + have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n + proof - + have "t-s = 2 * (real_of_int n * pi)/x" + using that by (simp add: field_simps) + then show ?thesis by (metis abs_minus_commute) + qed + have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" + by force + show ?thesis using False + apply (simp add: simple_path_def) + apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) + apply (subst abs_away) + apply (auto simp: 1) + apply (rule ccontr) + apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) + done +qed + +lemma arc_part_circlepath: + assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" + shows "arc (part_circlepath z r s t)" +proof - + have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" + and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n + proof (rule ccontr) + assume "x \ y" + have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" + by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_div_cancel_left eq) + then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" + by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) + with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" + by (force simp: field_simps) + have "\real_of_int n\ < \y - x\" + using assms \x \ y\ by (simp add: st abs_mult field_simps) + then show False + using assms x y st by (auto dest: of_int_lessD) + qed + show ?thesis + using assms + apply (simp add: arc_def) + apply (simp add: part_circlepath_def inj_on_def exp_eq) + apply (blast intro: *) + done +qed + +subsection\Special case of one complete circle\ + +definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" + where "circlepath z r \ part_circlepath z r 0 (2*pi)" + +lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" + by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) + +lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" + by (simp add: circlepath_def) + +lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" + by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) + +lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" +proof - + have "z + of_real r * exp (2 * pi * \ * (x + 1/2)) = + z + of_real r * exp (2 * pi * \ * x + pi * \)" + by (simp add: divide_simps) (simp add: algebra_simps) + also have "\ = z - r * exp (2 * pi * \ * x)" + by (simp add: exp_add) + finally show ?thesis + by (simp add: circlepath path_image_def sphere_def dist_norm) +qed + +lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" + using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] + by (simp add: add.commute) + +lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" + using circlepath_add1 [of z r "x-1/2"] + by (simp add: add.commute) + +lemma path_image_circlepath_minus_subset: + "path_image (circlepath z (-r)) \ path_image (circlepath z r)" + apply (simp add: path_image_def image_def circlepath_minus, clarify) + apply (case_tac "xa \ 1/2", force) + apply (force simp: circlepath_add_half)+ + done + +lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" + using path_image_circlepath_minus_subset by fastforce + +lemma has_vector_derivative_circlepath [derivative_intros]: + "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) + (at x within X)" + apply (simp add: circlepath_def scaleR_conv_of_real) + apply (rule derivative_eq_intros) + apply (simp add: algebra_simps) + done + +lemma vector_derivative_circlepath: + "vector_derivative (circlepath z r) (at x) = + 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" +using has_vector_derivative_circlepath vector_derivative_at by blast + +lemma vector_derivative_circlepath01: + "\0 \ x; x \ 1\ + \ vector_derivative (circlepath z r) (at x within {0..1}) = + 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" + using has_vector_derivative_circlepath + by (auto simp: vector_derivative_at_within_ivl) + +lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" + by (simp add: circlepath_def) + +lemma path_circlepath [simp]: "path (circlepath z r)" + by (simp add: valid_path_imp_path) + +lemma path_image_circlepath_nonneg: + assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" +proof - + have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x + proof (cases "x = z") + case True then show ?thesis by force + next + case False + define w where "w = x - z" + then have "w \ 0" by (simp add: False) + have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" + using cis_conv_exp complex_eq_iff by auto + show ?thesis + apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) + apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) + apply (rule_tac x="t / (2*pi)" in image_eqI) + apply (simp add: field_simps \w \ 0\) + using False ** + apply (auto simp: w_def) + done + qed + show ?thesis + unfolding circlepath path_image_def sphere_def dist_norm + by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) +qed + +lemma path_image_circlepath [simp]: + "path_image (circlepath z r) = sphere z \r\" + using path_image_circlepath_minus + by (force simp: path_image_circlepath_nonneg abs_if) + +lemma has_contour_integral_bound_circlepath_strong: + "\(f has_contour_integral i) (circlepath z r); + finite k; 0 \ B; 0 < r; + \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ + \ norm i \ B*(2*pi*r)" + unfolding circlepath_def + by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) + +lemma has_contour_integral_bound_circlepath: + "\(f has_contour_integral i) (circlepath z r); + 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ + \ norm i \ B*(2*pi*r)" + by (auto intro: has_contour_integral_bound_circlepath_strong) + +lemma contour_integrable_continuous_circlepath: + "continuous_on (path_image (circlepath z r)) f + \ f contour_integrable_on (circlepath z r)" + by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) + +lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" + by (simp add: circlepath_def simple_path_part_circlepath) + +lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" + by (simp add: sphere_def dist_norm norm_minus_commute) + +lemma contour_integral_circlepath: + assumes "r > 0" + shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" +proof (rule contour_integral_unique) + show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" + unfolding has_contour_integral_def using assms + apply (subst has_integral_cong) + apply (simp add: vector_derivative_circlepath01) + using has_integral_const_real [of _ 0 1] apply (force simp: circlepath) + done +qed + +lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" + apply (rule winding_number_unique_loop) + apply (simp_all add: sphere_def valid_path_imp_path) + apply (rule_tac x="circlepath z r" in exI) + apply (simp add: sphere_def contour_integral_circlepath) + done + +proposition winding_number_circlepath: + assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" +proof (cases "w = z") + case True then show ?thesis + using assms winding_number_circlepath_centre by auto +next + case False + have [simp]: "r > 0" + using assms le_less_trans norm_ge_zero by blast + define r' where "r' = norm(w - z)" + have "r' < r" + by (simp add: assms r'_def) + have disjo: "cball z r' \ sphere z r = {}" + using \r' < r\ by (force simp: cball_def sphere_def) + have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" + proof (rule winding_number_around_inside [where s = "cball z r'"]) + show "winding_number (circlepath z r) z \ 0" + by (simp add: winding_number_circlepath_centre) + show "cball z r' \ path_image (circlepath z r) = {}" + by (simp add: disjo less_eq_real_def) + qed (auto simp: r'_def dist_norm norm_minus_commute) + also have "\ = 1" + by (simp add: winding_number_circlepath_centre) + finally show ?thesis . +qed + + +text\ Hence the Cauchy formula for points inside a circle.\ + +theorem Cauchy_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r" + shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) + (circlepath z r)" +proof - + have "r > 0" + using assms le_less_trans norm_ge_zero by blast + have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) + (circlepath z r)" + proof (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) + show "\x. x \ interior (cball z r) - {} \ + f field_differentiable at x" + using holf holomorphic_on_imp_differentiable_at by auto + have "w \ sphere z r" + by simp (metis dist_commute dist_norm not_le order_refl wz) + then show "path_image (circlepath z r) \ cball z r - {w}" + using \r > 0\ by (auto simp add: cball_def sphere_def) + qed (use wz in \simp_all add: dist_norm norm_minus_commute contf\) + then show ?thesis + by (simp add: winding_number_circlepath assms) +qed + +corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: + assumes "f holomorphic_on cball z r" "norm(w - z) < r" + shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) + (circlepath z r)" +using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) + + +lemma no_bounded_connected_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" + and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" + shows "winding_number g z = 0" +apply (rule winding_number_zero_in_outside) +apply (simp_all add: assms) +by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) + +lemma no_bounded_path_component_imp_winding_number_zero: + assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" + and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" + shows "winding_number g z = 0" +apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) +by (simp add: bounded_subset nb path_component_subset_connected_component) + + +subsection\ Uniform convergence of path integral\ + +text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ + +proposition contour_integral_uniform_limit: + assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" + and ul_f: "uniform_limit (path_image \) f l F" + and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" + and \: "valid_path \" + and [simp]: "\ trivial_limit F" + shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" +proof - + have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) + { fix e::real + assume "0 < e" + then have "0 < e / (\B\ + 1)" by simp + then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" + using ul_f [unfolded uniform_limit_iff dist_norm] by auto + with ev_fint + obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" + and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" + using eventually_happens [OF eventually_conj] + by (fastforce simp: contour_integrable_on path_image_def) + have Ble: "B * e / (\B\ + 1) \ e" + using \0 \ B\ \0 < e\ by (simp add: field_split_simps) + have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" + proof (intro exI conjI ballI) + show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" + if "x \ {0..1}" for x + apply (rule order_trans [OF _ Ble]) + using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ + apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) + apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le]) + done + qed (rule inta) + } + then show lintg: "l contour_integrable_on \" + unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) + { fix e::real + define B' where "B' = B + 1" + have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) + assume "0 < e" + then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" + using ul_f [unfolded uniform_limit_iff dist_norm, rule_format, of "e / B' / 2"] B' + by (simp add: field_simps) + have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp + have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" + if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t + proof - + have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" + using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto + also have "\ < e" + by (simp add: B' \0 < e\ mult_imp_div_pos_less) + finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . + then show ?thesis + by (simp add: left_diff_distrib [symmetric] norm_mult) + qed + have le_e: "\x. \\xa\{0..1}. 2 * cmod (f x (\ xa) - l (\ xa)) < e / B'; f x contour_integrable_on \\ + \ cmod (integral {0..1} + (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" + apply (rule le_less_trans [OF integral_norm_bound_integral ie]) + apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) + apply (blast intro: *)+ + done + have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" + apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) + apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) + apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric] le_e) + done + } + then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" + by (rule tendstoI) +qed + +corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: + assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" + and "uniform_limit (sphere z r) f l F" + and "\ trivial_limit F" "0 < r" + shows "l contour_integrable_on (circlepath z r)" + "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" + using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) + + +subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ + +lemma Cauchy_next_derivative: + assumes "continuous_on (path_image \) f'" + and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" + and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" + and k: "k \ 0" + and "open s" + and \: "valid_path \" + and w: "w \ s - path_image \" + shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" + and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) + (at w)" (is "?thes2") +proof - + have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast + then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w + using open_contains_ball by blast + have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" + by (metis norm_of_nat of_nat_Suc) + have cint: "\x. \x \ w; cmod (x - w) < d\ + \ (\z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \" + apply (rule contour_integrable_div [OF contour_integrable_diff]) + using int w d + by (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ + have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) + contour_integrable_on \" + unfolding eventually_at + apply (rule_tac x=d in exI) + apply (simp add: \d > 0\ dist_norm field_simps cint) + done + have bim_g: "bounded (image f' (path_image \))" + by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) + then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" + by (force simp: bounded_pos path_image_def) + have twom: "\\<^sub>F n in at w. + \x\path_image \. + cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" + if "0 < e" for e + proof - + have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" + if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" + and uw_less: "cmod (u - w) < e * (d/2) ^ (k+2) / (1 + real k)" + for u x + proof - + define ff where [abs_def]: + "ff n w = + (if n = 0 then inverse(x - w)^k + else if n = 1 then k / (x - w)^(Suc k) + else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w + have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" + by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) + have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d/2))" + if "z \ ball w (d/2)" "i \ 1" for i z + proof - + have "z \ path_image \" + using \x \ path_image \\ d that ball_divide_subset_numeral by blast + then have xz[simp]: "x \ z" using \x \ path_image \\ by blast + then have neq: "x * x + z * z \ x * (z * 2)" + by (blast intro: dest!: sum_sqs_eq) + with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto + then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" + by (simp add: algebra_simps) + show ?thesis using \i \ 1\ + apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) + apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ + done + qed + { fix a::real and b::real assume ab: "a > 0" "b > 0" + then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" + by (subst mult_le_cancel_left_pos) + (use \k \ 0\ in \auto simp: divide_simps\) + with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" + by (simp add: field_simps) + } note canc = this + have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d/2) ^ (k + 2)" + if "v \ ball w (d/2)" for v + proof - + have lessd: "\z. cmod (\ z - v) < d/2 \ cmod (w - \ z) < d" + by (metis that norm_minus_commute norm_triangle_half_r dist_norm mem_ball) + have "d/2 \ cmod (x - v)" using d x that + using lessd d x + by (auto simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps) + then have "d \ cmod (x - v) * 2" + by (simp add: field_split_simps) + then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" + using \0 < d\ order_less_imp_le power_mono by blast + have "x \ v" using that + using \x \ path_image \\ ball_divide_subset_numeral d by fastforce + then show ?thesis + using \d > 0\ apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) + using dpow_le apply (simp add: field_split_simps) + done + qed + have ub: "u \ ball w (d/2)" + using uwd by (simp add: dist_commute dist_norm) + have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d/2) ^ k))" + using complex_Taylor [OF _ ff1 ff2 _ ub, of w, simplified] + by (simp add: ff_def \0 < d\) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" + by (simp add: field_simps) + then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) + / (cmod (u - w) * real k) + \ (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" + using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) + also have "\ < e" + using uw_less \0 < d\ by (simp add: mult_ac divide_simps) + finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) + / cmod ((u - w) * real k) < e" + by (simp add: norm_mult) + have "x \ u" + using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) + show ?thesis + apply (rule le_less_trans [OF _ e]) + using \k \ 0\ \x \ u\ \u \ w\ + apply (simp add: field_simps norm_divide [symmetric]) + done + qed + show ?thesis + unfolding eventually_at + apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) + apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) + done + qed + have 2: "uniform_limit (path_image \) (\n x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) (\x. f' x / (x - w) ^ Suc k) (at w)" + unfolding uniform_limit_iff dist_norm + proof clarify + fix e::real + assume "0 < e" + have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" + if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" + and x: "0 \ x" "x \ 1" + for u x + proof (cases "(f' (\ x)) = 0") + case True then show ?thesis by (simp add: \0 < e\) + next + case False + have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = + cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k))" + by (simp add: field_simps) + also have "\ = cmod (f' (\ x)) * + cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - + inverse (\ x - w) * inverse (\ x - w) ^ k)" + by (simp add: norm_mult) + also have "\ < cmod (f' (\ x)) * (e/C)" + using False mult_strict_left_mono [OF ec] by force + also have "\ \ e" using C + by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) + finally show ?thesis . + qed + show "\\<^sub>F n in at w. + \x\path_image \. + cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" + using twom [OF divide_pos_pos [OF \0 < e\ \C > 0\]] unfolding path_image_def + by (force intro: * elim: eventually_mono) + qed + show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" + by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto + have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) + \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" + by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto + have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = + (f u - f w) / (u - w) / k" + if "dist u w < d" for u + proof - + have u: "u \ s - path_image \" + by (metis subsetD d dist_commute mem_ball that) + show ?thesis + apply (rule contour_integral_unique) + apply (simp add: diff_divide_distrib algebra_simps) + apply (intro has_contour_integral_diff has_contour_integral_div) + using u w apply (simp_all add: field_simps int) + done + qed + show ?thes2 + apply (simp add: has_field_derivative_iff del: power_Suc) + apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) + apply (simp add: \k \ 0\ **) + done +qed + +lemma Cauchy_next_derivative_circlepath: + assumes contf: "continuous_on (path_image (circlepath z r)) f" + and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" + and k: "k \ 0" + and w: "w \ ball z r" + shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" + (is "?thes2") +proof - + have "r > 0" using w + using ball_eq_empty by fastforce + have wim: "w \ ball z r - path_image (circlepath z r)" + using w by (auto simp: dist_norm) + show ?thes1 ?thes2 + by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; + auto simp: vector_derivative_circlepath norm_mult)+ +qed + + +text\ In particular, the first derivative formula.\ + +lemma Cauchy_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" + (is "?thes2") +proof - + have [simp]: "r \ 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def) + have int: "\w. dist z w < r \ + ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" + by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) + show ?thes1 + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) + apply (blast intro: int) + done + have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" + apply (simp add: power2_eq_square) + apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) + apply (blast intro: int) + done + then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" + by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) + show ?thes2 + by simp (rule fder) +qed + +subsection\Existence of all higher derivatives\ + +proposition derivative_is_holomorphic: + assumes "open S" + and fder: "\z. z \ S \ (f has_field_derivative f' z) (at z)" + shows "f' holomorphic_on S" +proof - + have *: "\h. (f' has_field_derivative h) (at z)" if "z \ S" for z + proof - + obtain r where "r > 0" and r: "cball z r \ S" + using open_contains_cball \z \ S\ \open S\ by blast + then have holf_cball: "f holomorphic_on cball z r" + apply (simp add: holomorphic_on_def) + using field_differentiable_at_within field_differentiable_def fder by blast + then have "continuous_on (path_image (circlepath z r)) f" + using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) + then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" + by (auto intro: continuous_intros)+ + have contf_cball: "continuous_on (cball z r) f" using holf_cball + by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) + have holf_ball: "f holomorphic_on ball z r" using holf_cball + using ball_subset_cball holomorphic_on_subset by blast + { fix w assume w: "w \ ball z r" + have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) + (at w)" + by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) + have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" + using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) + have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral + contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) + (circlepath z r)" + by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) + then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral + contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) + (circlepath z r)" + by (simp add: algebra_simps) + then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" + by (simp add: f'_eq) + } note * = this + show ?thesis + apply (rule exI) + apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) + apply (simp_all add: \0 < r\ * dist_norm) + done + qed + show ?thesis + by (simp add: holomorphic_on_open [OF \open S\] *) +qed + +lemma holomorphic_deriv [holomorphic_intros]: + "\f holomorphic_on S; open S\ \ (deriv f) holomorphic_on S" +by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) + +lemma analytic_deriv [analytic_intros]: "f analytic_on S \ (deriv f) analytic_on S" + using analytic_on_holomorphic holomorphic_deriv by auto + +lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv ^^ n) f holomorphic_on S" + by (induction n) (auto simp: holomorphic_deriv) + +lemma analytic_higher_deriv [analytic_intros]: "f analytic_on S \ (deriv ^^ n) f analytic_on S" + unfolding analytic_on_def using holomorphic_higher_deriv by blast + +lemma has_field_derivative_higher_deriv: + "\f holomorphic_on S; open S; x \ S\ + \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" +by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply + funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) + +lemma valid_path_compose_holomorphic: + assumes "valid_path g" and holo:"f holomorphic_on S" and "open S" "path_image g \ S" + shows "valid_path (f \ g)" +proof (rule valid_path_compose[OF \valid_path g\]) + fix x assume "x \ path_image g" + then show "f field_differentiable at x" + using analytic_on_imp_differentiable_at analytic_on_open assms holo by blast +next + have "deriv f holomorphic_on S" + using holomorphic_deriv holo \open S\ by auto + then show "continuous_on (path_image g) (deriv f)" + using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto +qed + + +subsection\Morera's theorem\ + +lemma Morera_local_triangle_ball: + assumes "\z. z \ S + \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ + (\b c. closed_segment b c \ ball a e + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on S" +proof - + { fix z assume "z \ S" + with assms obtain e a where + "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" + and 0: "\b c. closed_segment b c \ ball a e + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by blast + have az: "dist a z < e" using mem_ball z by blast + have sb_ball: "ball z (e - dist a z) \ ball a e" + by (simp add: dist_commute ball_subset_ball_iff) + have "\e>0. f holomorphic_on ball z e" + proof (intro exI conjI) + have sub_ball: "\y. dist a y < e \ closed_segment a y \ ball a e" + by (meson \0 < e\ centre_in_ball convex_ball convex_contains_segment mem_ball) + show "f holomorphic_on ball z (e - dist a z)" + apply (rule holomorphic_on_subset [OF _ sb_ball]) + apply (rule derivative_is_holomorphic[OF open_ball]) + apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) + apply (simp_all add: 0 \0 < e\ sub_ball) + done + qed (simp add: az) + } + then show ?thesis + by (simp add: analytic_on_def) +qed + +lemma Morera_local_triangle: + assumes "\z. z \ S + \ \t. open t \ z \ t \ continuous_on t f \ + (\a b c. convex hull {a,b,c} \ t + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + shows "f analytic_on S" +proof - + { fix z assume "z \ S" + with assms obtain t where + "open t" and z: "z \ t" and contf: "continuous_on t f" + and 0: "\a b c. convex hull {a,b,c} \ t + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" + by force + then obtain e where "e>0" and e: "ball z e \ t" + using open_contains_ball by blast + have [simp]: "continuous_on (ball z e) f" using contf + using continuous_on_subset e by blast + have eq0: "\b c. closed_segment b c \ ball z e \ + contour_integral (linepath z b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c z) f = 0" + by (meson 0 z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) + have "\e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ + (\b c. closed_segment b c \ ball a e \ + contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" + using \e > 0\ eq0 by force + } + then show ?thesis + by (simp add: Morera_local_triangle_ball) +qed + +proposition Morera_triangle: + "\continuous_on S f; open S; + \a b c. convex hull {a,b,c} \ S + \ contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0\ + \ f analytic_on S" + using Morera_local_triangle by blast + +subsection\Combining theorems for higher derivatives including Leibniz rule\ + +lemma higher_deriv_linear [simp]: + "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" + by (induction n) auto + +lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" + by (induction n) auto + +lemma higher_deriv_ident [simp]: + "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" + apply (induction n, simp) + apply (metis higher_deriv_linear lambda_one) + done + +lemma higher_deriv_id [simp]: + "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" + by (simp add: id_def) + +lemma has_complex_derivative_funpow_1: + "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" + apply (induction n, auto) + apply (simp add: id_def) + by (metis DERIV_chain comp_funpow comp_id funpow_swap1 mult.right_neutral) + +lemma higher_deriv_uminus: + assumes "f holomorphic_on S" "open S" and z: "z \ S" + shows "(deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have "((deriv ^^ n) (\w. - f w) has_field_derivative - deriv ((deriv ^^ n) f) z) (at z)" + apply (rule has_field_derivative_transform_within_open [of "\w. -((deriv ^^ n) f w)"]) + apply (rule derivative_eq_intros | rule * refl assms)+ + apply (auto simp add: Suc) + done + then show ?case + by (simp add: DERIV_imp_deriv) +qed + +lemma higher_deriv_add: + fixes z::complex + assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" + shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have "((deriv ^^ n) (\w. f w + g w) has_field_derivative + deriv ((deriv ^^ n) f) z + deriv ((deriv ^^ n) g) z) (at z)" + apply (rule has_field_derivative_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) + apply (rule derivative_eq_intros | rule * refl assms)+ + apply (auto simp add: Suc) + done + then show ?case + by (simp add: DERIV_imp_deriv) +qed + +lemma higher_deriv_diff: + fixes z::complex + assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" + shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" + apply (simp only: Groups.group_add_class.diff_conv_add_uminus higher_deriv_add) + apply (subst higher_deriv_add) + using assms holomorphic_on_minus apply (auto simp: higher_deriv_uminus) + done + +lemma bb: "Suc n choose k = (n choose k) + (if k = 0 then 0 else (n choose (k - 1)))" + by (cases k) simp_all + +lemma higher_deriv_mult: + fixes z::complex + assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" + shows "(deriv ^^ n) (\w. f w * g w) z = + (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have *: "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) z) (at z)" + "\n. ((deriv ^^ n) g has_field_derivative deriv ((deriv ^^ n) g) z) (at z)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + have sumeq: "(\i = 0..n. + of_nat (n choose i) * (deriv ((deriv ^^ i) f) z * (deriv ^^ (n - i)) g z + deriv ((deriv ^^ (n - i)) g) z * (deriv ^^ i) f z)) = + g z * deriv ((deriv ^^ n) f) z + (\i = 0..n. (deriv ^^ i) f z * (of_nat (Suc n choose i) * (deriv ^^ (Suc n - i)) g z))" + apply (simp add: bb algebra_simps sum.distrib) + apply (subst (4) sum_Suc_reindex) + apply (auto simp: algebra_simps Suc_diff_le intro: sum.cong) + done + have "((deriv ^^ n) (\w. f w * g w) has_field_derivative + (\i = 0..Suc n. (Suc n choose i) * (deriv ^^ i) f z * (deriv ^^ (Suc n - i)) g z)) + (at z)" + apply (rule has_field_derivative_transform_within_open + [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) + apply (simp add: algebra_simps) + apply (rule DERIV_cong [OF DERIV_sum]) + apply (rule DERIV_cmult) + apply (auto intro: DERIV_mult * sumeq \open S\ Suc.prems Suc.IH [symmetric]) + done + then show ?case + unfolding funpow.simps o_apply + by (simp add: DERIV_imp_deriv) +qed + +lemma higher_deriv_transform_within_open: + fixes z::complex + assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" + and fg: "\w. w \ S \ f w = g w" + shows "(deriv ^^ i) f z = (deriv ^^ i) g z" +using z +by (induction i arbitrary: z) + (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) + +lemma higher_deriv_compose_linear: + fixes z::complex + assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \ S" + and fg: "\w. w \ S \ u * w \ T" + shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" +using z +proof (induction n arbitrary: z) + case 0 then show ?case by simp +next + case (Suc n z) + have holo0: "f holomorphic_on (*) u ` S" + by (meson fg f holomorphic_on_subset image_subset_iff) + have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S" + by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T) + have holo3: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S" + by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros) + have holo1: "(\w. f (u * w)) holomorphic_on S" + apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) + apply (rule holo0 holomorphic_intros)+ + done + have "deriv ((deriv ^^ n) (\w. f (u * w))) z = deriv (\z. u^n * (deriv ^^ n) f (u*z)) z" + apply (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems]) + apply (rule holomorphic_higher_deriv [OF holo1 S]) + apply (simp add: Suc.IH) + done + also have "\ = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" + apply (rule deriv_cmult) + apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) + apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and T=T, unfolded o_def]) + apply (simp) + apply (simp add: analytic_on_open f holomorphic_higher_deriv T) + apply (blast intro: fg) + done + also have "\ = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" + apply (subst deriv_chain [where g = "(deriv ^^ n) f" and f = "(*) u", unfolded o_def]) + apply (rule derivative_intros) + using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv T apply blast + apply (simp) + done + finally show ?case + by simp +qed + +lemma higher_deriv_add_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_add show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_diff_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_diff show ?thesis + by (auto simp: analytic_at_two) +qed + +lemma higher_deriv_uminus_at: + "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" + using higher_deriv_uminus + by (auto simp: analytic_at) + +lemma higher_deriv_mult_at: + assumes "f analytic_on {z}" "g analytic_on {z}" + shows "(deriv ^^ n) (\w. f w * g w) z = + (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" +proof - + have "f analytic_on {z} \ g analytic_on {z}" + using assms by blast + with higher_deriv_mult show ?thesis + by (auto simp: analytic_at_two) +qed + + +text\ Nonexistence of isolated singularities and a stronger integral formula.\ + +proposition no_isolated_singularity: + fixes z::complex + assumes f: "continuous_on S f" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" + shows "f holomorphic_on S" +proof - + { fix z + assume "z \ S" and cdf: "\x. x \ S - K \ f field_differentiable at x" + have "f field_differentiable at z" + proof (cases "z \ K") + case False then show ?thesis by (blast intro: cdf \z \ S\) + next + case True + with finite_set_avoid [OF K, of z] + obtain d where "d>0" and d: "\x. \x\K; x \ z\ \ d \ dist z x" + by blast + obtain e where "e>0" and e: "ball z e \ S" + using S \z \ S\ by (force simp: open_contains_ball) + have fde: "continuous_on (ball z (min d e)) f" + by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) + have cont: "{a,b,c} \ ball z (min d e) \ continuous_on (convex hull {a, b, c}) f" for a b c + by (simp add: hull_minimal continuous_on_subset [OF fde]) + have fd: "\{a,b,c} \ ball z (min d e); x \ interior (convex hull {a, b, c}) - K\ + \ f field_differentiable at x" for a b c x + by (metis cdf Diff_iff Int_iff ball_min_Int subsetD convex_ball e interior_mono interior_subset subset_hull) + obtain g where "\w. w \ ball z (min d e) \ (g has_field_derivative f w) (at w within ball z (min d e))" + apply (rule contour_integral_convex_primitive + [OF convex_ball fde Cauchy_theorem_triangle_cofinite [OF _ K]]) + using cont fd by auto + then have "f holomorphic_on ball z (min d e)" + by (metis open_ball at_within_open derivative_is_holomorphic) + then show ?thesis + unfolding holomorphic_on_def + by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) + qed + } + with holf S K show ?thesis + by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) +qed + +lemma no_isolated_singularity': + fixes z::complex + assumes f: "\z. z \ K \ (f \ f z) (at z within S)" + and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" + shows "f holomorphic_on S" +proof (rule no_isolated_singularity[OF _ assms(2-)]) + show "continuous_on S f" unfolding continuous_on_def + proof + fix z assume z: "z \ S" + show "(f \ f z) (at z within S)" + proof (cases "z \ K") + case False + from holf have "continuous_on (S - K) f" + by (rule holomorphic_on_imp_continuous_on) + with z False have "(f \ f z) (at z within (S - K))" + by (simp add: continuous_on_def) + also from z K S False have "at z within (S - K) = at z within S" + by (subst (1 2) at_within_open) (auto intro: finite_imp_closed) + finally show "(f \ f z) (at z within S)" . + qed (insert assms z, simp_all) + qed +qed + +proposition Cauchy_integral_formula_convex: + assumes S: "convex S" and K: "finite K" and contf: "continuous_on S f" + and fcd: "(\x. x \ interior S - K \ f field_differentiable at x)" + and z: "z \ interior S" and vpg: "valid_path \" + and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + have *: "\x. x \ interior S \ f field_differentiable at x" + unfolding holomorphic_on_open [symmetric] field_differentiable_def + using no_isolated_singularity [where S = "interior S"] + by (meson K contf continuous_at_imp_continuous_on continuous_on_interior fcd + field_differentiable_at_within field_differentiable_def holomorphic_onI + holomorphic_on_imp_differentiable_at open_interior) + show ?thesis + by (rule Cauchy_integral_formula_weak [OF S finite.emptyI contf]) (use * assms in auto) +qed + +text\ Formula for higher derivatives.\ + +lemma Cauchy_has_contour_integral_higher_derivative_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) + (circlepath z r)" +using w +proof (induction k arbitrary: w) + case 0 then show ?case + using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) +next + case (Suc k) + have [simp]: "r > 0" using w + using ball_eq_empty by fastforce + have f: "continuous_on (path_image (circlepath z r)) f" + by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def less_imp_le) + obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" + using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] + by (auto simp: contour_integrable_on_def) + then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" + by (rule contour_integral_unique) + have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" + using Suc.prems assms has_field_derivative_higher_deriv by auto + then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" + by (force simp: field_differentiable_def) + have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = + of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" + by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) + also have "\ = of_nat (Suc k) * X" + by (simp only: con) + finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . + then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" + by (metis deriv_cmult dnf_diff) + then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" + by (simp add: field_simps) + then show ?case + using of_nat_eq_0_iff X by fastforce +qed + +lemma Cauchy_higher_derivative_integral_circlepath: + assumes contf: "continuous_on (cball z r) f" + and holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" + (is "?thes1") + and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" + (is "?thes2") +proof - + have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) + (circlepath z r)" + using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] + by simp + show ?thes1 using * + using contour_integrable_on_def by blast + show ?thes2 + unfolding contour_integral_unique [OF *] by (simp add: field_split_simps) +qed + +corollary Cauchy_contour_integral_circlepath: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" + shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" +by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) + +lemma Cauchy_contour_integral_circlepath_2: + assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" + shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" + using Cauchy_contour_integral_circlepath [OF assms, of 1] + by (simp add: power2_eq_square) + + +subsection\A holomorphic function is analytic, i.e. has local power series\ + +theorem holomorphic_power_series: + assumes holf: "f holomorphic_on ball z r" + and w: "w \ ball z r" + shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" +proof - + \ \Replacing \<^term>\r\ and the original (weak) premises with stronger ones\ + obtain r where "r > 0" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" + proof + have "cball z ((r + dist w z) / 2) \ ball z r" + using w by (simp add: dist_commute field_sum_of_halves subset_eq) + then show "f holomorphic_on cball z ((r + dist w z) / 2)" + by (rule holomorphic_on_subset [OF holf]) + have "r > 0" + using w by clarsimp (metis dist_norm le_less_trans norm_ge_zero) + then show "0 < (r + dist w z) / 2" + by simp (use zero_le_dist [of w z] in linarith) + qed (use w in \auto simp: dist_commute\) + then have holf: "f holomorphic_on ball z r" + using ball_subset_cball holomorphic_on_subset by blast + have contf: "continuous_on (cball z r) f" + by (simp add: holfc holomorphic_on_imp_continuous_on) + have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" + by (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) (simp add: \0 < r\) + obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" + by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) + obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" + and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" + proof + show "\u. cmod (u - z) = r \ r - dist z w \ cmod (u - w)" + by (metis add_diff_eq diff_add_cancel dist_norm norm_diff_ineq) + qed (use w in \auto simp: dist_norm norm_minus_commute\) + have ul: "uniform_limit (sphere z r) (\n x. (\kx. f x / (x - w)) sequentially" + unfolding uniform_limit_iff dist_norm + proof clarify + fix e::real + assume "0 < e" + have rr: "0 \ (r - k) / r" "(r - k) / r < 1" using k by auto + obtain n where n: "((r - k) / r) ^ n < e / B * k" + using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force + have "norm ((\k N" and r: "r = dist z u" for N u + proof - + have N: "((r - k) / r) ^ N < e / B * k" + apply (rule le_less_trans [OF power_decreasing n]) + using \n \ N\ k by auto + have u [simp]: "(u \ z) \ (u \ w)" + using \0 < r\ r w by auto + have wzu_not1: "(w - z) / (u - z) \ 1" + by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) + have "norm ((\kk = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" + using \0 < B\ + apply (auto simp: geometric_sum [OF wzu_not1]) + apply (simp add: field_simps norm_mult [symmetric]) + done + also have "\ = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" + using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) + also have "\ = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" + by (simp add: algebra_simps) + also have "\ = norm (w - z) ^ N * norm (f u) / r ^ N" + by (simp add: norm_mult norm_power norm_minus_commute) + also have "\ \ (((r - k)/r)^N) * B" + using \0 < r\ w k + apply (simp add: divide_simps) + apply (rule mult_mono [OF power_mono]) + apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) + done + also have "\ < e * k" + using \0 < B\ N by (simp add: divide_simps) + also have "\ \ e * norm (u - w)" + using r kle \0 < e\ by (simp add: dist_commute dist_norm) + finally show ?thesis + by (simp add: field_split_simps norm_divide del: power_Suc) + qed + with \0 < r\ show "\\<^sub>F n in sequentially. \x\sphere z r. + norm ((\k\<^sub>F x in sequentially. + contour_integral (circlepath z r) (\u. \kku. f u / (u - z) ^ Suc k) * (w - z) ^ k)" + apply (rule eventuallyI) + apply (subst contour_integral_sum, simp) + using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) + apply (simp only: contour_integral_lmul cint algebra_simps) + done + have cic: "\u. (\y. \k0 < r\ by (force intro!: Cauchy_higher_derivative_integral_circlepath [OF contf holf]) + have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) + sums contour_integral (circlepath z r) (\u. f u/(u - w))" + unfolding sums_def + apply (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul] cic) + using \0 < r\ apply auto + done + then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) + sums (2 * of_real pi * \ * f w)" + using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) + then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) + sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" + by (rule sums_divide) + then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) + sums f w" + by (simp add: field_simps) + then show ?thesis + by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) +qed + + +subsection\The Liouville theorem and the Fundamental Theorem of Algebra\ + +text\ These weak Liouville versions don't even need the derivative formula.\ + +lemma Liouville_weak_0: + assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" + shows "f z = 0" +proof (rule ccontr) + assume fz: "f z \ 0" + with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] + obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" + by (auto simp: dist_norm) + define R where "R = 1 + \B\ + norm z" + have "R > 0" unfolding R_def + proof - + have "0 \ cmod z + \B\" + by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) + then show "0 < 1 + \B\ + cmod z" + by linarith + qed + have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" + apply (rule Cauchy_integral_circlepath) + using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ + done + have "cmod (x - z) = R \ cmod (f x) * 2 < cmod (f z)" for x + unfolding R_def + by (rule B) (use norm_triangle_ineq4 [of x z] in auto) + with \R > 0\ fz show False + using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] + by (auto simp: less_imp_le norm_mult norm_divide field_split_simps) +qed + +proposition Liouville_weak: + assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" + shows "f z = l" + using Liouville_weak_0 [of "\z. f z - l"] + by (simp add: assms holomorphic_on_diff LIM_zero) + +proposition Liouville_weak_inverse: + assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" + obtains z where "f z = 0" +proof - + { assume f: "\z. f z \ 0" + have 1: "(\x. 1 / f x) holomorphic_on UNIV" + by (simp add: holomorphic_on_divide assms f) + have 2: "((\x. 1 / f x) \ 0) at_infinity" + apply (rule tendstoI [OF eventually_mono]) + apply (rule_tac B="2/e" in unbounded) + apply (simp add: dist_norm norm_divide field_split_simps) + done + have False + using Liouville_weak_0 [OF 1 2] f by simp + } + then show ?thesis + using that by blast +qed + +text\ In particular we get the Fundamental Theorem of Algebra.\ + +theorem fundamental_theorem_of_algebra: + fixes a :: "nat \ complex" + assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" + obtains z where "(\i\n. a i * z^i) = 0" +using assms +proof (elim disjE bexE) + assume "a 0 = 0" then show ?thesis + by (auto simp: that [of 0]) +next + fix i + assume i: "i \ {1..n}" and nz: "a i \ 0" + have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" + by (rule holomorphic_intros)+ + show thesis + proof (rule Liouville_weak_inverse [OF 1]) + show "\\<^sub>F x in at_infinity. B \ cmod (\i\n. a i * x ^ i)" for B + using i nz by (intro polyfun_extremal exI[of _ i]) auto + qed (use that in auto) +qed + +subsection\Weierstrass convergence theorem\ + +lemma holomorphic_uniform_limit: + assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" + and ulim: "uniform_limit (cball z r) f g F" + and F: "\ trivial_limit F" + obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" +proof (cases r "0::real" rule: linorder_cases) + case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) +next + case equal then show ?thesis + by (force simp: holomorphic_on_def intro: that) +next + case greater + have contg: "continuous_on (cball z r) g" + using cont uniform_limit_theorem [OF eventually_mono ulim F] by blast + have "path_image (circlepath z r) \ cball z r" + using \0 < r\ by auto + then have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" + by (intro continuous_intros continuous_on_subset [OF contg]) + have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" + if w: "w \ ball z r" for w + proof - + define d where "d = (r - norm(w - z))" + have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) + have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" + unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) + have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" + 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\ \path \\ compact_path_image + by blast + obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ U" + apply (rule that [of "d0/2"]) + using \0 < d0\ + apply (auto simp: dist_norm dest: d0) + done + have "\x x'. \x \ path_image \; dist x x' * 2 < dd\ \ \y k. x' = y + k \ y \ path_image \ \ dist 0 k * 2 \ dd" + apply (rule_tac x=x in exI) + apply (rule_tac x="x'-x" in exI) + apply (force simp: dist_norm) + done + then have 1: "path_image \ \ interior {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" + apply (clarsimp simp add: mem_interior) + using \0 < dd\ + apply (rule_tac x="dd/2" in exI, auto) + done + obtain T where "compact T" and subt: "path_image \ \ interior T" and T: "T \ U" + apply (rule that [OF _ 1]) + apply (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) + apply (rule order_trans [OF _ dd]) + using \0 < dd\ by fastforce + obtain L where "L>0" + and L: "\f B. \f holomorphic_on interior T; \z. z\interior T \ cmod (f z) \ B\ \ + cmod (contour_integral \ f) \ L * B" + using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] + by blast + have "bounded(f ` T)" + by (meson \compact T\ compact_continuous_image compact_imp_bounded conf continuous_on_subset T) + then obtain D where "D>0" and D: "\x. x \ T \ norm (f x) \ D" + by (auto simp: bounded_pos) + obtain C where "C>0" and C: "\x. x \ T \ norm x \ C" + using \compact T\ bounded_pos compact_imp_bounded by force + have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y + proof - + have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp + with le have ybig: "norm y > C" by force + with C have "y \ T" by force + then have ynot: "y \ path_image \" + using subt interior_subset by blast + have [simp]: "winding_number \ y = 0" + apply (rule winding_number_zero_outside [of _ "cball 0 C"]) + using ybig interior_subset subt + apply (force simp: loop \path \\ dist_norm intro!: C)+ + done + have [simp]: "h y = contour_integral \ (\w. f w/(w - y))" + by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) + have holint: "(\w. f w / (w - y)) holomorphic_on interior T" + apply (rule holomorphic_on_divide) + using holf holomorphic_on_subset interior_subset T apply blast + apply (rule holomorphic_intros)+ + using \y \ T\ interior_subset by auto + have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior T" for z + proof - + have "D * L / e + cmod z \ cmod y" + using le C [of z] z using interior_subset by force + then have DL2: "D * L / e \ cmod (z - y)" + using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) + have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" + by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) + also have "\ \ D * (e / L / D)" + apply (rule mult_mono) + using that D interior_subset apply blast + using \L>0\ \e>0\ \D>0\ DL2 + apply (auto simp: norm_divide field_split_simps) + done + finally show ?thesis . + qed + have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" + by (simp add: dist_norm) + also have "\ \ L * (D * (e / L / D))" + by (rule L [OF holint leD]) + also have "\ = e" + using \L>0\ \0 < D\ by auto + finally show ?thesis . + qed + then have "(h \ 0) at_infinity" + by (meson Lim_at_infinityI) + moreover have "h holomorphic_on UNIV" + proof - + have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" + if "x \ U" "z \ U" "x \ z" for x z + using that conf + apply (simp add: split_def continuous_on_eq_continuous_at \open U\) + apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ + done + have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" + by (rule continuous_intros)+ + have open_uu_Id: "open (U \ U - Id)" + apply (rule open_Diff) + apply (simp add: open_Times \open U\) + using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] + apply (auto simp: Id_fstsnd_eq algebra_simps) + done + have con_derf: "continuous (at z) (deriv f)" if "z \ U" for z + apply (rule continuous_on_interior [of U]) + apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open U\) + by (simp add: interior_open that \open U\) + have tendsto_f': "((\(x,y). if y = x then deriv f (x) + else (f (y) - f (x)) / (y - x)) \ deriv f x) + (at (x, x) within U \ U)" if "x \ U" for x + proof (rule Lim_withinI) + fix e::real assume "0 < e" + obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" + using \0 < e\ continuous_within_E [OF con_derf [OF \x \ U\]] + by (metis UNIV_I dist_norm) + obtain k2 where "k2>0" and k2: "ball x k2 \ U" + by (blast intro: openE [OF \open U\] \x \ U\) + have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" + if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" + for x' z' + proof - + have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w + apply (drule segment_furthest_le [where y=x]) + by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) + have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w + by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) + have f_has_der: "\x. x \ U \ (f has_field_derivative deriv f x) (at x within U)" + by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \open U\) + have "closed_segment x' z' \ U" + by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) + then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" + using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp + then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" + by (rule has_contour_integral_div) + have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" + apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) + using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] + \e > 0\ \z' \ x'\ + apply (auto simp: norm_divide divide_simps derf_le) + done + also have "\ \ e" using \0 < e\ by simp + finally show ?thesis . + qed + show "\d>0. \xa\U \ U. + 0 < dist xa (x, x) \ dist xa (x, x) < d \ + dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" + apply (rule_tac x="min k1 k2" in exI) + using \k1>0\ \k2>0\ \e>0\ + apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) + done + qed + have con_pa_f: "continuous_on (path_image \) f" + by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt T) + have le_B: "\T. T \ {0..1} \ cmod (vector_derivative \ (at T)) \ B" + apply (rule B) + using \' using path_image_def vector_derivative_at by fastforce + have f_has_cint: "\w. w \ v - path_image \ \ ((\u. f u / (u - w) ^ 1) has_contour_integral h w) \" + by (simp add: V) + have cond_uu: "continuous_on (U \ U) (\(x,y). d x y)" + apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') + apply (simp add: tendsto_within_open_NO_MATCH open_Times \open U\, clarify) + apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) + using con_ff + apply (auto simp: continuous_within) + done + have hol_dw: "(\z. d z w) holomorphic_on U" if "w \ U" for w + proof - + have "continuous_on U ((\(x,y). d x y) \ (\z. (w,z)))" + by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ + then have *: "continuous_on U (\z. if w = z then deriv f z else (f w - f z) / (w - z))" + by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) + have **: "\x. \x \ U; x \ w\ \ (\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" + apply (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) + apply (rule \open U\ derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+ + done + show ?thesis + unfolding d_def + apply (rule no_isolated_singularity [OF * _ \open U\, where K = "{w}"]) + apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \open U\ **) + done + qed + { fix a b + assume abu: "closed_segment a b \ U" + then have "\w. w \ U \ (\z. d z w) contour_integrable_on (linepath a b)" + by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) + then have cont_cint_d: "continuous_on U (\w. contour_integral (linepath a b) (\z. d z w))" + apply (rule contour_integral_continuous_on_linepath_2D [OF \open U\ _ _ abu]) + apply (auto intro: continuous_on_swap_args cond_uu) + done + have cont_cint_d\: "continuous_on {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" + proof (rule continuous_on_compose) + show "continuous_on {0..1} \" + using \path \\ path_def by blast + show "continuous_on (\ ` {0..1}) (\w. contour_integral (linepath a b) (\z. d z w))" + using pasz unfolding path_image_def + by (auto intro!: continuous_on_subset [OF cont_cint_d]) + qed + have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" + apply (simp add: contour_integrable_on) + apply (rule integrable_continuous_real) + apply (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) + using pf\' + by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) + have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\z. contour_integral \ (d z))" + using abu by (force simp: h_def intro: contour_integral_eq) + also have "\ = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" + apply (rule contour_integral_swap) + apply (rule continuous_on_subset [OF cond_uu]) + using abu pasz \valid_path \\ + apply (auto intro!: continuous_intros) + by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) + finally have cint_h_eq: + "contour_integral (linepath a b) h = + contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . + note cint_cint cint_h_eq + } note cint_h = this + have conthu: "continuous_on U h" + proof (simp add: continuous_on_sequentially, clarify) + fix a x + assume x: "x \ U" and au: "\n. a n \ U" and ax: "a \ x" + then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" + by (meson U contour_integrable_on_def eventuallyI) + obtain dd where "dd>0" and dd: "cball x dd \ U" using open_contains_cball \open U\ x by force + have A2: "uniform_limit (path_image \) (\n. d (a n)) (d x) sequentially" + unfolding uniform_limit_iff dist_norm + proof clarify + fix ee::real + assume "0 < ee" + show "\\<^sub>F n in sequentially. \\\path_image \. cmod (d (a n) \ - d x \) < ee" + proof - + let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" + have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" + apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) + using dd pasz \valid_path \\ + apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) + done + then obtain kk where "kk>0" + and kk: "\x x'. \x \ ?ddpa; x' \ ?ddpa; dist x' x < kk\ \ + dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" + by (rule uniformly_continuous_onE [where e = ee]) (use \0 < ee\ in auto) + have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" + for w z + using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm) + show ?thesis + using ax unfolding lim_sequentially eventually_sequentially + apply (drule_tac x="min dd kk" in spec) + using \dd > 0\ \kk > 0\ + apply (fastforce simp: kk dist_norm) + done + qed + qed + have "(\n. contour_integral \ (d (a n))) \ contour_integral \ (d x)" + by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \valid_path \\) + then have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" + by (simp add: h_def x) + then show "(h \ a) \ h x" + by (simp add: h_def x au o_def) + qed + show ?thesis + proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) + fix z0 + consider "z0 \ v" | "z0 \ U" using uv_Un by blast + then show "h field_differentiable at z0" + proof cases + assume "z0 \ v" then show ?thesis + using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ + by (auto simp: field_differentiable_def v_def) + next + assume "z0 \ U" then + obtain e where "e>0" and e: "ball z0 e \ U" by (blast intro: openE [OF \open U\]) + have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" + if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c + proof - + have *: "\x1 x2 z. z \ U \ closed_segment x1 x2 \ U \ (\w. d w z) contour_integrable_on linepath x1 x2" + using hol_dw holomorphic_on_imp_continuous_on \open U\ + by (auto intro!: contour_integrable_holomorphic_simple) + have abc: "closed_segment a b \ U" "closed_segment b c \ U" "closed_segment c a \ U" + using that e segments_subset_convex_hull by fastforce+ + have eq0: "\w. w \ U \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" + apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) + apply (rule holomorphic_on_subset [OF hol_dw]) + using e abc_subset by auto + have "contour_integral \ + (\x. contour_integral (linepath a b) (\z. d z x) + + (contour_integral (linepath b c) (\z. d z x) + + contour_integral (linepath c a) (\z. d z x))) = 0" + apply (rule contour_integral_eq_0) + using abc pasz U + apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ + done + then show ?thesis + by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) + qed + show ?thesis + using e \e > 0\ + by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic + Morera_triangle continuous_on_subset [OF conthu] *) + qed + qed + qed + ultimately have [simp]: "h z = 0" for z + by (meson Liouville_weak) + have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" + by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) + then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" + by (metis mult.commute has_contour_integral_lmul) + then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" + by (simp add: field_split_simps) + moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" + using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) + show ?thesis + using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) +qed + +theorem Cauchy_integral_formula_global: + assumes S: "open S" and holf: "f holomorphic_on S" + and z: "z \ S" and vpg: "valid_path \" + and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" + and zero: "\w. w \ S \ winding_number \ w = 0" + shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" +proof - + have "path \" using vpg by (blast intro: valid_path_imp_path) + have hols: "(\w. f w / (w - z)) holomorphic_on S - {z}" "(\w. 1 / (w - z)) holomorphic_on S - {z}" + by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ + then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" + by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete S vpg pasz) + obtain d where "d>0" + and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; + pathstart h = pathstart g \ pathfinish h = pathfinish g\ + \ path_image h \ S - {z} \ (\f. f holomorphic_on S - {z} \ contour_integral h f = contour_integral g f)" + using contour_integral_nearby_ends [OF _ \path \\ pasz] S by (simp add: open_Diff) metis + obtain p where polyp: "polynomial_function p" + and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" + using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast + then have ploop: "pathfinish p = pathstart p" using loop by auto + have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast + have [simp]: "z \ path_image \" using pasz by blast + have paps: "path_image p \ S - {z}" and cint_eq: "(\f. f holomorphic_on S - {z} \ contour_integral p f = contour_integral \ f)" + using pf ps led d [OF vpg vpp] \d > 0\ by auto + have wn_eq: "winding_number p z = winding_number \ z" + using vpp paps + by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) + have "winding_number p w = winding_number \ w" if "w \ S" for w + proof - + have hol: "(\v. 1 / (v - w)) holomorphic_on S - {z}" + using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) + have "w \ path_image p" "w \ path_image \" using paps pasz that by auto + then show ?thesis + using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) + qed + then have wn0: "\w. w \ S \ winding_number p w = 0" + by (simp add: zero) + show ?thesis + using Cauchy_integral_formula_global_weak [OF S holf z polyp paps ploop wn0] hols + by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) +qed + +theorem Cauchy_theorem_global: + assumes S: "open S" and holf: "f holomorphic_on S" + and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" + and pas: "path_image \ \ S" + and zero: "\w. w \ S \ winding_number \ w = 0" + shows "(f has_contour_integral 0) \" +proof - + obtain z where "z \ S" and znot: "z \ path_image \" + proof - + have "compact (path_image \)" + using compact_valid_path_image vpg by blast + then have "path_image \ \ S" + by (metis (no_types) compact_open path_image_nonempty S) + with pas show ?thesis by (blast intro: that) + qed + then have pasz: "path_image \ \ S - {z}" using pas by blast + have hol: "(\w. (w - z) * f w) holomorphic_on S" + by (rule holomorphic_intros holf)+ + show ?thesis + using Cauchy_integral_formula_global [OF S hol \z \ S\ vpg pasz loop zero] + by (auto simp: znot elim!: has_contour_integral_eq) +qed + +corollary Cauchy_theorem_global_outside: + assumes "open S" "f holomorphic_on S" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ S" + "\w. w \ S \ w \ outside(path_image \)" + shows "(f has_contour_integral 0) \" +by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) + +lemma simply_connected_imp_winding_number_zero: + assumes "simply_connected S" "path g" + "path_image g \ S" "pathfinish g = pathstart g" "z \ S" + shows "winding_number g z = 0" +proof - + have hom: "homotopic_loops S g (linepath (pathstart g) (pathstart g))" + by (meson assms homotopic_paths_imp_homotopic_loops pathfinish_linepath simply_connected_eq_contractible_path) + then have "homotopic_paths (- {z}) g (linepath (pathstart g) (pathstart g))" + by (meson \z \ S\ homotopic_loops_imp_homotopic_paths_null homotopic_paths_subset subset_Compl_singleton) + then have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" + by (rule winding_number_homotopic_paths) + also have "\ = 0" + using assms by (force intro: winding_number_trivial) + finally show ?thesis . +qed + +lemma Cauchy_theorem_simply_connected: + assumes "open S" "simply_connected S" "f holomorphic_on S" "valid_path g" + "path_image g \ S" "pathfinish g = pathstart g" + shows "(f has_contour_integral 0) g" +using assms +apply (simp add: simply_connected_eq_contractible_path) +apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"] + homotopic_paths_imp_homotopic_loops) +using valid_path_imp_path by blast + +proposition\<^marker>\tag unimportant\ holomorphic_logarithm_exists: + assumes A: "convex A" "open A" + and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" + and z0: "z0 \ A" + obtains g where "g holomorphic_on A" and "\x. x \ A \ exp (g x) = f x" +proof - + note f' = holomorphic_derivI [OF f(1) A(2)] + obtain g where g: "\x. x \ A \ (g has_field_derivative deriv f x / f x) (at x)" + proof (rule holomorphic_convex_primitive' [OF A]) + show "(\x. deriv f x / f x) holomorphic_on A" + by (intro holomorphic_intros f A) + qed (auto simp: A at_within_open[of _ A]) + define h where "h = (\x. -g z0 + ln (f z0) + g x)" + from g and A have g_holo: "g holomorphic_on A" + by (auto simp: holomorphic_on_def at_within_open[of _ A] field_differentiable_def) + hence h_holo: "h holomorphic_on A" + by (auto simp: h_def intro!: holomorphic_intros) + have "\c. \x\A. f x / exp (h x) - 1 = c" + proof (rule has_field_derivative_zero_constant, goal_cases) + case (2 x) + note [simp] = at_within_open[OF _ \open A\] + from 2 and z0 and f show ?case + by (auto simp: h_def exp_diff field_simps intro!: derivative_eq_intros g f') + qed fact+ + then obtain c where c: "\x. x \ A \ f x / exp (h x) - 1 = c" + by blast + from c[OF z0] and z0 and f have "c = 0" + by (simp add: h_def) + with c have "\x. x \ A \ exp (h x) = f x" by simp + from that[OF h_holo this] show ?thesis . +qed + +subsection \Complex functions and power series\ + +text \ + The following defines the power series expansion of a complex function at a given point + (assuming that it is analytic at that point). +\ +definition\<^marker>\tag important\ fps_expansion :: "(complex \ complex) \ complex \ complex fps" where + "fps_expansion f z0 = Abs_fps (\n. (deriv ^^ n) f z0 / fact n)" + +lemma + fixes r :: ereal + assumes "f holomorphic_on eball z0 r" + shows conv_radius_fps_expansion: "fps_conv_radius (fps_expansion f z0) \ r" + and eval_fps_expansion: "\z. z \ eball z0 r \ eval_fps (fps_expansion f z0) (z - z0) = f z" + and eval_fps_expansion': "\z. norm z < r \ eval_fps (fps_expansion f z0) z = f (z0 + z)" +proof - + have "(\n. fps_nth (fps_expansion f z0) n * (z - z0) ^ n) sums f z" + if "z \ ball z0 r'" "ereal r' < r" for z r' + proof - + from that(2) have "ereal r' \ r" by simp + from assms(1) and this have "f holomorphic_on ball z0 r'" + by (rule holomorphic_on_subset[OF _ ball_eball_mono]) + from holomorphic_power_series [OF this that(1)] + show ?thesis by (simp add: fps_expansion_def) + qed + hence *: "(\n. fps_nth (fps_expansion f z0) n * (z - z0) ^ n) sums f z" + if "z \ eball z0 r" for z + using that by (subst (asm) eball_conv_UNION_balls) blast + show "fps_conv_radius (fps_expansion f z0) \ r" unfolding fps_conv_radius_def + proof (rule conv_radius_geI_ex) + fix r' :: real assume r': "r' > 0" "ereal r' < r" + thus "\z. norm z = r' \ summable (\n. fps_nth (fps_expansion f z0) n * z ^ n)" + using *[of "z0 + of_real r'"] + by (intro exI[of _ "of_real r'"]) (auto simp: summable_def dist_norm) + qed + show "eval_fps (fps_expansion f z0) (z - z0) = f z" if "z \ eball z0 r" for z + using *[OF that] by (simp add: eval_fps_def sums_iff) + show "eval_fps (fps_expansion f z0) z = f (z0 + z)" if "ereal (norm z) < r" for z + using *[of "z0 + z"] and that by (simp add: eval_fps_def sums_iff dist_norm) +qed + + +text \ + We can now show several more facts about power series expansions (at least in the complex case) + with relative ease that would have been trickier without complex analysis. +\ +lemma + fixes f :: "complex fps" and r :: ereal + assumes "\z. ereal (norm z) < r \ eval_fps f z \ 0" + shows fps_conv_radius_inverse: "fps_conv_radius (inverse f) \ min r (fps_conv_radius f)" + and eval_fps_inverse: "\z. ereal (norm z) < fps_conv_radius f \ ereal (norm z) < r \ + eval_fps (inverse f) z = inverse (eval_fps f z)" +proof - + define R where "R = min (fps_conv_radius f) r" + have *: "fps_conv_radius (inverse f) \ min r (fps_conv_radius f) \ + (\z\eball 0 (min (fps_conv_radius f) r). eval_fps (inverse f) z = inverse (eval_fps f z))" + proof (cases "min r (fps_conv_radius f) > 0") + case True + define f' where "f' = fps_expansion (\z. inverse (eval_fps f z)) 0" + have holo: "(\z. inverse (eval_fps f z)) holomorphic_on eball 0 (min r (fps_conv_radius f))" + using assms by (intro holomorphic_intros) auto + from holo have radius: "fps_conv_radius f' \ min r (fps_conv_radius f)" + unfolding f'_def by (rule conv_radius_fps_expansion) + have eval_f': "eval_fps f' z = inverse (eval_fps f z)" + if "norm z < fps_conv_radius f" "norm z < r" for z + using that unfolding f'_def by (subst eval_fps_expansion'[OF holo]) auto + + have "f * f' = 1" + proof (rule eval_fps_eqD) + from radius and True have "0 < min (fps_conv_radius f) (fps_conv_radius f')" + by (auto simp: min_def split: if_splits) + also have "\ \ fps_conv_radius (f * f')" by (rule fps_conv_radius_mult) + finally show "\ > 0" . + next + from True have "R > 0" by (auto simp: R_def) + hence "eventually (\z. z \ eball 0 R) (nhds 0)" + by (intro eventually_nhds_in_open) (auto simp: zero_ereal_def) + thus "eventually (\z. eval_fps (f * f') z = eval_fps 1 z) (nhds 0)" + proof eventually_elim + case (elim z) + hence "eval_fps (f * f') z = eval_fps f z * eval_fps f' z" + using radius by (intro eval_fps_mult) + (auto simp: R_def min_def split: if_splits intro: less_trans) + also have "eval_fps f' z = inverse (eval_fps f z)" + using elim by (intro eval_f') (auto simp: R_def) + also from elim have "eval_fps f z \ 0" + by (intro assms) (auto simp: R_def) + hence "eval_fps f z * inverse (eval_fps f z) = eval_fps 1 z" + by simp + finally show "eval_fps (f * f') z = eval_fps 1 z" . + qed + qed simp_all + hence "f' = inverse f" + by (intro fps_inverse_unique [symmetric]) (simp_all add: mult_ac) + with eval_f' and radius show ?thesis by simp + next + case False + hence *: "eball 0 R = {}" + by (intro eball_empty) (auto simp: R_def min_def split: if_splits) + show ?thesis + proof safe + from False have "min r (fps_conv_radius f) \ 0" + by (simp add: min_def) + also have "0 \ fps_conv_radius (inverse f)" + by (simp add: fps_conv_radius_def conv_radius_nonneg) + finally show "min r (fps_conv_radius f) \ \" . + qed (unfold * [unfolded R_def], auto) + qed + + from * show "fps_conv_radius (inverse f) \ min r (fps_conv_radius f)" by blast + from * show "eval_fps (inverse f) z = inverse (eval_fps f z)" + if "ereal (norm z) < fps_conv_radius f" "ereal (norm z) < r" for z + using that by auto +qed + +lemma + fixes f g :: "complex fps" and r :: ereal + defines "R \ Min {r, fps_conv_radius f, fps_conv_radius g}" + assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0" "r > 0" + assumes nz: "\z. z \ eball 0 r \ eval_fps g z \ 0" + shows fps_conv_radius_divide': "fps_conv_radius (f / g) \ R" + and eval_fps_divide': + "ereal (norm z) < R \ eval_fps (f / g) z = eval_fps f z / eval_fps g z" +proof - + from nz[of 0] and \r > 0\ have nz': "fps_nth g 0 \ 0" + by (auto simp: eval_fps_at_0 zero_ereal_def) + have "R \ min r (fps_conv_radius g)" + by (auto simp: R_def intro: min.coboundedI2) + also have "min r (fps_conv_radius g) \ fps_conv_radius (inverse g)" + by (intro fps_conv_radius_inverse assms) (auto simp: zero_ereal_def) + finally have radius: "fps_conv_radius (inverse g) \ R" . + have "R \ min (fps_conv_radius f) (fps_conv_radius (inverse g))" + by (intro radius min.boundedI) (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) + also have "\ \ fps_conv_radius (f * inverse g)" + by (rule fps_conv_radius_mult) + also have "f * inverse g = f / g" + by (intro fps_divide_unit [symmetric] nz') + finally show "fps_conv_radius (f / g) \ R" . + + assume z: "ereal (norm z) < R" + have "eval_fps (f * inverse g) z = eval_fps f z * eval_fps (inverse g) z" + using radius by (intro eval_fps_mult less_le_trans[OF z]) + (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) + also have "eval_fps (inverse g) z = inverse (eval_fps g z)" using \r > 0\ + by (intro eval_fps_inverse[where r = r] less_le_trans[OF z] nz) + (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) + also have "f * inverse g = f / g" by fact + finally show "eval_fps (f / g) z = eval_fps f z / eval_fps g z" by (simp add: field_split_simps) +qed + +lemma + fixes f g :: "complex fps" and r :: ereal + defines "R \ Min {r, fps_conv_radius f, fps_conv_radius g}" + assumes "subdegree g \ subdegree f" + assumes "fps_conv_radius f > 0" "fps_conv_radius g > 0" "r > 0" + assumes "\z. z \ eball 0 r \ z \ 0 \ eval_fps g z \ 0" + shows fps_conv_radius_divide: "fps_conv_radius (f / g) \ R" + and eval_fps_divide: + "ereal (norm z) < R \ c = fps_nth f (subdegree g) / fps_nth g (subdegree g) \ + eval_fps (f / g) z = (if z = 0 then c else eval_fps f z / eval_fps g z)" +proof - + define f' g' where "f' = fps_shift (subdegree g) f" and "g' = fps_shift (subdegree g) g" + have f_eq: "f = f' * fps_X ^ subdegree g" and g_eq: "g = g' * fps_X ^ subdegree g" + unfolding f'_def g'_def by (rule subdegree_decompose' le_refl | fact)+ + have subdegree: "subdegree f' = subdegree f - subdegree g" "subdegree g' = 0" + using assms(2) by (simp_all add: f'_def g'_def) + have [simp]: "fps_conv_radius f' = fps_conv_radius f" "fps_conv_radius g' = fps_conv_radius g" + by (simp_all add: f'_def g'_def) + have [simp]: "fps_nth f' 0 = fps_nth f (subdegree g)" + "fps_nth g' 0 = fps_nth g (subdegree g)" by (simp_all add: f'_def g'_def) + have g_nz: "g \ 0" + proof - + define z :: complex where "z = (if r = \ then 1 else of_real (real_of_ereal r / 2))" + from \r > 0\ have "z \ eball 0 r" + by (cases r) (auto simp: z_def eball_def) + moreover have "z \ 0" using \r > 0\ + by (cases r) (auto simp: z_def) + ultimately have "eval_fps g z \ 0" by (rule assms(6)) + thus "g \ 0" by auto + qed + have fg: "f / g = f' * inverse g'" + by (subst f_eq, subst (2) g_eq) (insert g_nz, simp add: fps_divide_unit) + + have g'_nz: "eval_fps g' z \ 0" if z: "norm z < min r (fps_conv_radius g)" for z + proof (cases "z = 0") + case False + with assms and z have "eval_fps g z \ 0" by auto + also from z have "eval_fps g z = eval_fps g' z * z ^ subdegree g" + by (subst g_eq) (auto simp: eval_fps_mult) + finally show ?thesis by auto + qed (insert \g \ 0\, auto simp: g'_def eval_fps_at_0) + + have "R \ min (min r (fps_conv_radius g)) (fps_conv_radius g')" + by (auto simp: R_def min.coboundedI1 min.coboundedI2) + also have "\ \ fps_conv_radius (inverse g')" + using g'_nz by (rule fps_conv_radius_inverse) + finally have conv_radius_inv: "R \ fps_conv_radius (inverse g')" . + hence "R \ fps_conv_radius (f' * inverse g')" + by (intro order.trans[OF _ fps_conv_radius_mult]) + (auto simp: R_def intro: min.coboundedI1 min.coboundedI2) + thus "fps_conv_radius (f / g) \ R" by (simp add: fg) + + fix z c :: complex assume z: "ereal (norm z) < R" + assume c: "c = fps_nth f (subdegree g) / fps_nth g (subdegree g)" + show "eval_fps (f / g) z = (if z = 0 then c else eval_fps f z / eval_fps g z)" + proof (cases "z = 0") + case False + from z and conv_radius_inv have "ereal (norm z) < fps_conv_radius (inverse g')" + by simp + with z have "eval_fps (f / g) z = eval_fps f' z * eval_fps (inverse g') z" + unfolding fg by (subst eval_fps_mult) (auto simp: R_def) + also have "eval_fps (inverse g') z = inverse (eval_fps g' z)" + using z by (intro eval_fps_inverse[of "min r (fps_conv_radius g')"] g'_nz) (auto simp: R_def) + also have "eval_fps f' z * \ = eval_fps f z / eval_fps g z" + using z False assms(2) by (simp add: f'_def g'_def eval_fps_shift R_def) + finally show ?thesis using False by simp + qed (simp_all add: eval_fps_at_0 fg field_simps c) +qed + +lemma has_fps_expansion_fps_expansion [intro]: + assumes "open A" "0 \ A" "f holomorphic_on A" + shows "f has_fps_expansion fps_expansion f 0" +proof - + from assms(1,2) obtain r where r: "r > 0 " "ball 0 r \ A" + by (auto simp: open_contains_ball) + have holo: "f holomorphic_on eball 0 (ereal r)" + using r(2) and assms(3) by auto + from r(1) have "0 < ereal r" by simp + also have "r \ fps_conv_radius (fps_expansion f 0)" + using holo by (intro conv_radius_fps_expansion) auto + finally have "\ > 0" . + moreover have "eventually (\z. z \ ball 0 r) (nhds 0)" + using r(1) by (intro eventually_nhds_in_open) auto + hence "eventually (\z. eval_fps (fps_expansion f 0) z = f z) (nhds 0)" + by eventually_elim (subst eval_fps_expansion'[OF holo], auto) + ultimately show ?thesis using r(1) by (auto simp: has_fps_expansion_def) +qed + +lemma fps_conv_radius_tan: + fixes c :: complex + assumes "c \ 0" + shows "fps_conv_radius (fps_tan c) \ pi / (2 * norm c)" +proof - + have "fps_conv_radius (fps_tan c) \ + Min {pi / (2 * norm c), fps_conv_radius (fps_sin c), fps_conv_radius (fps_cos c)}" + unfolding fps_tan_def + proof (rule fps_conv_radius_divide) + fix z :: complex assume "z \ eball 0 (pi / (2 * norm c))" + with cos_eq_zero_imp_norm_ge[of "c*z"] assms + show "eval_fps (fps_cos c) z \ 0" by (auto simp: norm_mult field_simps) + qed (insert assms, auto) + thus ?thesis by (simp add: min_def) +qed + +lemma eval_fps_tan: + fixes c :: complex + assumes "norm z < pi / (2 * norm c)" + shows "eval_fps (fps_tan c) z = tan (c * z)" +proof (cases "c = 0") + case False + show ?thesis unfolding fps_tan_def + proof (subst eval_fps_divide'[where r = "pi / (2 * norm c)"]) + fix z :: complex assume "z \ eball 0 (pi / (2 * norm c))" + with cos_eq_zero_imp_norm_ge[of "c*z"] assms + show "eval_fps (fps_cos c) z \ 0" using False by (auto simp: norm_mult field_simps) + qed (insert False assms, auto simp: field_simps tan_def) +qed simp_all + +end diff --git a/src/HOL/Complex_Analysis/Complex_Analysis.thy b/src/HOL/Complex_Analysis/Complex_Analysis.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/Complex_Analysis.thy @@ -0,0 +1,6 @@ +theory Complex_Analysis + imports + Winding_Numbers +begin + +end diff --git a/src/HOL/Complex_Analysis/Conformal_Mappings.thy b/src/HOL/Complex_Analysis/Conformal_Mappings.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy @@ -0,0 +1,5116 @@ +section \Conformal Mappings and Consequences of Cauchy's Integral Theorem\ + +text\By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\ + +text\Also Cauchy's residue theorem by Wenda Li (2016)\ + +theory Conformal_Mappings +imports Cauchy_Integral_Theorem + +begin + +(* FIXME mv to Cauchy_Integral_Theorem.thy *) +subsection\Cauchy's inequality and more versions of Liouville\ + +lemma Cauchy_higher_deriv_bound: + assumes holf: "f holomorphic_on (ball z r)" + and contf: "continuous_on (cball z r) f" + and fin : "\w. w \ ball z r \ f w \ ball y B0" + and "0 < r" and "0 < n" + shows "norm ((deriv ^^ n) f z) \ (fact n) * B0 / r^n" +proof - + have "0 < B0" using \0 < r\ fin [of z] + by (metis ball_eq_empty ex_in_conv fin not_less) + have le_B0: "\w. cmod (w - z) \ r \ cmod (f w - y) \ B0" + apply (rule continuous_on_closure_norm_le [of "ball z r" "\w. f w - y"]) + apply (auto simp: \0 < r\ dist_norm norm_minus_commute) + apply (rule continuous_intros contf)+ + using fin apply (simp add: dist_commute dist_norm less_eq_real_def) + done + have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w) z - (deriv ^^ n) (\w. y) z" + using \0 < n\ by simp + also have "... = (deriv ^^ n) (\w. f w - y) z" + by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \0 < r\) + finally have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w - y) z" . + have contf': "continuous_on (cball z r) (\u. f u - y)" + by (rule contf continuous_intros)+ + have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" + by (simp add: holf holomorphic_on_diff) + define a where "a = (2 * pi)/(fact n)" + have "0 < a" by (simp add: a_def) + have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" + using \0 < r\ by (simp add: a_def field_split_simps) + have der_dif: "(deriv ^^ n) (\w. f w - y) z = (deriv ^^ n) f z" + using \0 < r\ \0 < n\ + by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) + have "norm ((2 * of_real pi * \)/(fact n) * (deriv ^^ n) (\w. f w - y) z) + \ (B0/r^(Suc n)) * (2 * pi * r)" + apply (rule has_contour_integral_bound_circlepath [of "(\u. (f u - y)/(u - z)^(Suc n))" _ z]) + using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] + using \0 < B0\ \0 < r\ + apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) + done + then show ?thesis + using \0 < r\ + by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) +qed + +lemma Cauchy_inequality: + assumes holf: "f holomorphic_on (ball \ r)" + and contf: "continuous_on (cball \ r) f" + and "0 < r" + and nof: "\x. norm(\-x) = r \ norm(f x) \ B" + shows "norm ((deriv ^^ n) f \) \ (fact n) * B / r^n" +proof - + obtain x where "norm (\-x) = r" + by (metis abs_of_nonneg add_diff_cancel_left' \0 < r\ diff_add_cancel + dual_order.strict_implies_order norm_of_real) + then have "0 \ B" + by (metis nof norm_not_less_zero not_le order_trans) + have "((\u. f u / (u - \) ^ Suc n) has_contour_integral (2 * pi) * \ / fact n * (deriv ^^ n) f \) + (circlepath \ r)" + apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) + using \0 < r\ by simp + then have "norm ((2 * pi * \)/(fact n) * (deriv ^^ n) f \) \ (B / r^(Suc n)) * (2 * pi * r)" + apply (rule has_contour_integral_bound_circlepath) + using \0 \ B\ \0 < r\ + apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) + done + then show ?thesis using \0 < r\ + by (simp add: norm_divide norm_mult field_simps) +qed + +lemma Liouville_polynomial: + assumes holf: "f holomorphic_on UNIV" + and nof: "\z. A \ norm z \ norm(f z) \ B * norm z ^ n" + shows "f \ = (\k\n. (deriv^^k) f 0 / fact k * \ ^ k)" +proof (cases rule: le_less_linear [THEN disjE]) + assume "B \ 0" + then have "\z. A \ norm z \ norm(f z) = 0" + by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) + then have f0: "(f \ 0) at_infinity" + using Lim_at_infinity by force + then have [simp]: "f = (\w. 0)" + using Liouville_weak [OF holf, of 0] + by (simp add: eventually_at_infinity f0) meson + show ?thesis by simp +next + assume "0 < B" + have "((\k. (deriv ^^ k) f 0 / (fact k) * (\ - 0)^k) sums f \)" + apply (rule holomorphic_power_series [where r = "norm \ + 1"]) + using holf holomorphic_on_subset apply auto + done + then have sumsf: "((\k. (deriv ^^ k) f 0 / (fact k) * \^k) sums f \)" by simp + have "(deriv ^^ k) f 0 / fact k * \ ^ k = 0" if "k>n" for k + proof (cases "(deriv ^^ k) f 0 = 0") + case True then show ?thesis by simp + next + case False + define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" + have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" + using \0 < B\ by simp + then have wge1: "1 \ norm w" + by (metis norm_of_real w_def) + then have "w \ 0" by auto + have kB: "0 < fact k * B" + using \0 < B\ by simp + then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" + by simp + then have wgeA: "A \ cmod w" + by (simp only: w_def norm_of_real) + have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" + using \0 < B\ by simp + then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" + by (metis norm_of_real w_def) + then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" + using False by (simp add: field_split_simps mult.commute split: if_split_asm) + also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" + apply (rule Cauchy_inequality) + using holf holomorphic_on_subset apply force + using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast + using \w \ 0\ apply simp + by (metis nof wgeA dist_0_norm dist_norm) + also have "... = fact k * (B * 1 / cmod w ^ (k-n))" + apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) + using \k>n\ \w \ 0\ \0 < B\ apply (simp add: field_split_simps semiring_normalization_rules) + done + also have "... = fact k * B / cmod w ^ (k-n)" + by simp + finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . + then have "1 / cmod w < 1 / cmod w ^ (k - n)" + by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) + then have "cmod w ^ (k - n) < cmod w" + by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) + with self_le_power [OF wge1] have False + by (meson diff_is_0_eq not_gr0 not_le that) + then show ?thesis by blast + qed + then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \ ^ (k + Suc n) = 0" for k + using not_less_eq by blast + then have "(\i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \ ^ (i + Suc n)) sums 0" + by (rule sums_0) + with sums_split_initial_segment [OF sumsf, where n = "Suc n"] + show ?thesis + using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce +qed + +text\Every bounded entire function is a constant function.\ +theorem Liouville_theorem: + assumes holf: "f holomorphic_on UNIV" + and bf: "bounded (range f)" + obtains c where "\z. f z = c" +proof - + obtain B where "\z. cmod (f z) \ B" + by (meson bf bounded_pos rangeI) + then show ?thesis + using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast +qed + +text\A holomorphic function f has only isolated zeros unless f is 0.\ + +lemma powser_0_nonzero: + fixes a :: "nat \ 'a::{real_normed_field,banach}" + assumes r: "0 < r" + and sm: "\x. norm (x - \) < r \ (\n. a n * (x - \) ^ n) sums (f x)" + and [simp]: "f \ = 0" + and m0: "a m \ 0" and "m>0" + obtains s where "0 < s" and "\z. z \ cball \ s - {\} \ f z \ 0" +proof - + have "r \ conv_radius a" + using sm sums_summable by (auto simp: le_conv_radius_iff [where \=\]) + obtain m where am: "a m \ 0" and az [simp]: "(\n. n a n = 0)" + apply (rule_tac m = "LEAST n. a n \ 0" in that) + using m0 + apply (rule LeastI2) + apply (fastforce intro: dest!: not_less_Least)+ + done + define b where "b i = a (i+m) / a m" for i + define g where "g x = suminf (\i. b i * (x - \) ^ i)" for x + have [simp]: "b 0 = 1" + by (simp add: am b_def) + { fix x::'a + assume "norm (x - \) < r" + then have "(\n. (a m * (x - \)^m) * (b n * (x - \)^n)) sums (f x)" + using am az sm sums_zero_iff_shift [of m "(\n. a n * (x - \) ^ n)" "f x"] + by (simp add: b_def monoid_mult_class.power_add algebra_simps) + then have "x \ \ \ (\n. b n * (x - \)^n) sums (f x / (a m * (x - \)^m))" + using am by (simp add: sums_mult_D) + } note bsums = this + then have "norm (x - \) < r \ summable (\n. b n * (x - \)^n)" for x + using sums_summable by (cases "x=\") auto + then have "r \ conv_radius b" + by (simp add: le_conv_radius_iff [where \=\]) + then have "r/2 < conv_radius b" + using not_le order_trans r by fastforce + then have "continuous_on (cball \ (r/2)) g" + using powser_continuous_suminf [of "r/2" b \] by (simp add: g_def) + then obtain s where "s>0" "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ dist (g x) (g \) < 1/2" + apply (rule continuous_onE [where x=\ and e = "1/2"]) + using r apply (auto simp: norm_minus_commute dist_norm) + done + moreover have "g \ = 1" + by (simp add: g_def) + ultimately have gnz: "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ (g x) \ 0" + by fastforce + have "f x \ 0" if "x \ \" "norm (x - \) \ s" "norm (x - \) \ r/2" for x + using bsums [of x] that gnz [of x] + apply (auto simp: g_def) + using r sums_iff by fastforce + then show ?thesis + apply (rule_tac s="min s (r/2)" in that) + using \0 < r\ \0 < s\ by (auto simp: dist_commute dist_norm) +qed + +subsection \Analytic continuation\ + +proposition isolated_zeros: + assumes holf: "f holomorphic_on S" + and "open S" "connected S" "\ \ S" "f \ = 0" "\ \ S" "f \ \ 0" + obtains r where "0 < r" and "ball \ r \ S" and + "\z. z \ ball \ r - {\} \ f z \ 0" +proof - + obtain r where "0 < r" and r: "ball \ r \ S" + using \open S\ \\ \ S\ open_contains_ball_eq by blast + have powf: "((\n. (deriv ^^ n) f \ / (fact n) * (z - \)^n) sums f z)" if "z \ ball \ r" for z + apply (rule holomorphic_power_series [OF _ that]) + apply (rule holomorphic_on_subset [OF holf r]) + done + obtain m where m: "(deriv ^^ m) f \ / (fact m) \ 0" + using holomorphic_fun_eq_0_on_connected [OF holf \open S\ \connected S\ _ \\ \ S\ \\ \ S\] \f \ \ 0\ + by auto + then have "m \ 0" using assms(5) funpow_0 by fastforce + obtain s where "0 < s" and s: "\z. z \ cball \ s - {\} \ f z \ 0" + apply (rule powser_0_nonzero [OF \0 < r\ powf \f \ = 0\ m]) + using \m \ 0\ by (auto simp: dist_commute dist_norm) + have "0 < min r s" by (simp add: \0 < r\ \0 < s\) + then show ?thesis + apply (rule that) + using r s by auto +qed + +proposition analytic_continuation: + assumes holf: "f holomorphic_on S" + and "open S" and "connected S" + and "U \ S" and "\ \ S" + and "\ islimpt U" + and fU0 [simp]: "\z. z \ U \ f z = 0" + and "w \ S" + shows "f w = 0" +proof - + obtain e where "0 < e" and e: "cball \ e \ S" + using \open S\ \\ \ S\ open_contains_cball_eq by blast + define T where "T = cball \ e \ U" + have contf: "continuous_on (closure T) f" + by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on + holomorphic_on_subset inf.cobounded1) + have fT0 [simp]: "\x. x \ T \ f x = 0" + by (simp add: T_def) + have "\r. \\e>0. \x'\U. x' \ \ \ dist x' \ < e; 0 < r\ \ \x'\cball \ e \ U. x' \ \ \ dist x' \ < r" + by (metis \0 < e\ IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj) + then have "\ islimpt T" using \\ islimpt U\ + by (auto simp: T_def islimpt_approachable) + then have "\ \ closure T" + by (simp add: closure_def) + then have "f \ = 0" + by (auto simp: continuous_constant_on_closure [OF contf]) + show ?thesis + apply (rule ccontr) + apply (rule isolated_zeros [OF holf \open S\ \connected S\ \\ \ S\ \f \ = 0\ \w \ S\], assumption) + by (metis open_ball \\ islimpt T\ centre_in_ball fT0 insertE insert_Diff islimptE) +qed + +corollary analytic_continuation_open: + assumes "open s" and "open s'" and "s \ {}" and "connected s'" + and "s \ s'" + assumes "f holomorphic_on s'" and "g holomorphic_on s'" + and "\z. z \ s \ f z = g z" + assumes "z \ s'" + shows "f z = g z" +proof - + from \s \ {}\ obtain \ where "\ \ s" by auto + with \open s\ have \: "\ islimpt s" + by (intro interior_limit_point) (auto simp: interior_open) + have "f z - g z = 0" + by (rule analytic_continuation[of "\z. f z - g z" s' s \]) + (insert assms \\ \ s\ \, auto intro: holomorphic_intros) + thus ?thesis by simp +qed + +subsection\Open mapping theorem\ + +lemma holomorphic_contract_to_zero: + assumes contf: "continuous_on (cball \ r) f" + and holf: "f holomorphic_on ball \ r" + and "0 < r" + and norm_less: "\z. norm(\ - z) = r \ norm(f \) < norm(f z)" + obtains z where "z \ ball \ r" "f z = 0" +proof - + { assume fnz: "\w. w \ ball \ r \ f w \ 0" + then have "0 < norm (f \)" + by (simp add: \0 < r\) + have fnz': "\w. w \ cball \ r \ f w \ 0" + by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) + have "frontier(cball \ r) \ {}" + using \0 < r\ by simp + define g where [abs_def]: "g z = inverse (f z)" for z + have contg: "continuous_on (cball \ r) g" + unfolding g_def using contf continuous_on_inverse fnz' by blast + have holg: "g holomorphic_on ball \ r" + unfolding g_def using fnz holf holomorphic_on_inverse by blast + have "frontier (cball \ r) \ cball \ r" + by (simp add: subset_iff) + then have contf': "continuous_on (frontier (cball \ r)) f" + and contg': "continuous_on (frontier (cball \ r)) g" + by (blast intro: contf contg continuous_on_subset)+ + have froc: "frontier(cball \ r) \ {}" + using \0 < r\ by simp + moreover have "continuous_on (frontier (cball \ r)) (norm o f)" + using contf' continuous_on_compose continuous_on_norm_id by blast + ultimately obtain w where w: "w \ frontier(cball \ r)" + and now: "\x. x \ frontier(cball \ r) \ norm (f w) \ norm (f x)" + apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]]) + apply simp + done + then have fw: "0 < norm (f w)" + by (simp add: fnz') + have "continuous_on (frontier (cball \ r)) (norm o g)" + using contg' continuous_on_compose continuous_on_norm_id by blast + then obtain v where v: "v \ frontier(cball \ r)" + and nov: "\x. x \ frontier(cball \ r) \ norm (g v) \ norm (g x)" + apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]]) + apply simp + done + then have fv: "0 < norm (f v)" + by (simp add: fnz') + have "norm ((deriv ^^ 0) g \) \ fact 0 * norm (g v) / r ^ 0" + by (rule Cauchy_inequality [OF holg contg \0 < r\]) (simp add: dist_norm nov) + then have "cmod (g \) \ norm (g v)" + by simp + with w have wr: "norm (\ - w) = r" and nfw: "norm (f w) \ norm (f \)" + apply (simp_all add: dist_norm) + by (metis \0 < cmod (f \)\ g_def less_imp_inverse_less norm_inverse not_le now order_trans v) + with fw have False + using norm_less by force + } + with that show ?thesis by blast +qed + +theorem open_mapping_thm: + assumes holf: "f holomorphic_on S" + and S: "open S" and "connected S" + and "open U" and "U \ S" + and fne: "\ f constant_on S" + shows "open (f ` U)" +proof - + have *: "open (f ` U)" + if "U \ {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\x. \y \ U. f y \ x" + for U + proof (clarsimp simp: open_contains_ball) + fix \ assume \: "\ \ U" + show "\e>0. ball (f \) e \ f ` U" + proof - + have hol: "(\z. f z - f \) holomorphic_on U" + by (rule holomorphic_intros that)+ + obtain s where "0 < s" and sbU: "ball \ s \ U" + and sne: "\z. z \ ball \ s - {\} \ (\z. f z - f \) z \ 0" + using isolated_zeros [OF hol U \] by (metis fneU right_minus_eq) + obtain r where "0 < r" and r: "cball \ r \ ball \ s" + apply (rule_tac r="s/2" in that) + using \0 < s\ by auto + have "cball \ r \ U" + using sbU r by blast + then have frsbU: "frontier (cball \ r) \ U" + using Diff_subset frontier_def order_trans by fastforce + then have cof: "compact (frontier(cball \ r))" + by blast + have frne: "frontier (cball \ r) \ {}" + using \0 < r\ by auto + have contfr: "continuous_on (frontier (cball \ r)) (\z. norm (f z - f \))" + by (metis continuous_on_norm continuous_on_subset frsbU hol holomorphic_on_imp_continuous_on) + obtain w where "norm (\ - w) = r" + and w: "(\z. norm (\ - z) = r \ norm (f w - f \) \ norm(f z - f \))" + apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) + apply (simp add: dist_norm) + done + moreover define \ where "\ \ norm (f w - f \) / 3" + ultimately have "0 < \" + using \0 < r\ dist_complex_def r sne by auto + have "ball (f \) \ \ f ` U" + proof + fix \ + assume \: "\ \ ball (f \) \" + have *: "cmod (\ - f \) < cmod (\ - f z)" if "cmod (\ - z) = r" for z + proof - + have lt: "cmod (f w - f \) / 3 < cmod (\ - f z)" + using w [OF that] \ + using dist_triangle2 [of "f \" "\" "f z"] dist_triangle2 [of "f \" "f z" \] + by (simp add: \_def dist_norm norm_minus_commute) + show ?thesis + by (metis \_def dist_commute dist_norm less_trans lt mem_ball \) + qed + have "continuous_on (cball \ r) (\z. \ - f z)" + apply (rule continuous_intros)+ + using \cball \ r \ U\ \f holomorphic_on U\ + apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) + done + moreover have "(\z. \ - f z) holomorphic_on ball \ r" + apply (rule holomorphic_intros)+ + apply (metis \cball \ r \ U\ \f holomorphic_on U\ holomorphic_on_subset interior_cball interior_subset) + done + ultimately obtain z where "z \ ball \ r" "\ - f z = 0" + apply (rule holomorphic_contract_to_zero) + apply (blast intro!: \0 < r\ *)+ + done + then show "\ \ f ` U" + using \cball \ r \ U\ by fastforce + qed + then show ?thesis using \0 < \\ by blast + qed + qed + have "open (f ` X)" if "X \ components U" for X + proof - + have holfU: "f holomorphic_on U" + using \U \ S\ holf holomorphic_on_subset by blast + have "X \ {}" + using that by (simp add: in_components_nonempty) + moreover have "open X" + using that \open U\ open_components by auto + moreover have "connected X" + using that in_components_maximal by blast + moreover have "f holomorphic_on X" + by (meson that holfU holomorphic_on_subset in_components_maximal) + moreover have "\y\X. f y \ x" for x + proof (rule ccontr) + assume not: "\ (\y\X. f y \ x)" + have "X \ S" + using \U \ S\ in_components_subset that by blast + obtain w where w: "w \ X" using \X \ {}\ by blast + have wis: "w islimpt X" + using w \open X\ interior_eq by auto + have hol: "(\z. f z - x) holomorphic_on S" + by (simp add: holf holomorphic_on_diff) + with fne [unfolded constant_on_def] + analytic_continuation[OF hol S \connected S\ \X \ S\ _ wis] not \X \ S\ w + show False by auto + qed + ultimately show ?thesis + by (rule *) + qed + then have "open (f ` \(components U))" + by (metis (no_types, lifting) imageE image_Union open_Union) + then show ?thesis + by force +qed + +text\No need for \<^term>\S\ to be connected. But the nonconstant condition is stronger.\ +corollary\<^marker>\tag unimportant\ open_mapping_thm2: + assumes holf: "f holomorphic_on S" + and S: "open S" + and "open U" "U \ S" + and fnc: "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" + shows "open (f ` U)" +proof - + have "S = \(components S)" by simp + with \U \ S\ have "U = (\C \ components S. C \ U)" by auto + then have "f ` U = (\C \ components S. f ` (C \ U))" + using image_UN by fastforce + moreover + { fix C assume "C \ components S" + with S \C \ components S\ open_components in_components_connected + have C: "open C" "connected C" by auto + have "C \ S" + by (metis \C \ components S\ in_components_maximal) + have nf: "\ f constant_on C" + apply (rule fnc) + using C \C \ S\ \C \ components S\ in_components_nonempty by auto + have "f holomorphic_on C" + by (metis holf holomorphic_on_subset \C \ S\) + then have "open (f ` (C \ U))" + apply (rule open_mapping_thm [OF _ C _ _ nf]) + apply (simp add: C \open U\ open_Int, blast) + done + } ultimately show ?thesis + by force +qed + +corollary\<^marker>\tag unimportant\ open_mapping_thm3: + assumes holf: "f holomorphic_on S" + and "open S" and injf: "inj_on f S" + shows "open (f ` S)" +apply (rule open_mapping_thm2 [OF holf]) +using assms +apply (simp_all add:) +using injective_not_constant subset_inj_on by blast + +subsection\Maximum modulus principle\ + +text\If \<^term>\f\ is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is + properly within the domain of \<^term>\f\.\ + +proposition maximum_modulus_principle: + assumes holf: "f holomorphic_on S" + and S: "open S" and "connected S" + and "open U" and "U \ S" and "\ \ U" + and no: "\z. z \ U \ norm(f z) \ norm(f \)" + shows "f constant_on S" +proof (rule ccontr) + assume "\ f constant_on S" + then have "open (f ` U)" + using open_mapping_thm assms by blast + moreover have "\ open (f ` U)" + proof - + have "\t. cmod (f \ - t) < e \ t \ f ` U" if "0 < e" for e + apply (rule_tac x="if 0 < Re(f \) then f \ + (e/2) else f \ - (e/2)" in exI) + using that + apply (simp add: dist_norm) + apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym) + done + then show ?thesis + unfolding open_contains_ball by (metis \\ \ U\ contra_subsetD dist_norm imageI mem_ball) + qed + ultimately show False + by blast +qed + +proposition maximum_modulus_frontier: + assumes holf: "f holomorphic_on (interior S)" + and contf: "continuous_on (closure S) f" + and bos: "bounded S" + and leB: "\z. z \ frontier S \ norm(f z) \ B" + and "\ \ S" + shows "norm(f \) \ B" +proof - + have "compact (closure S)" using bos + by (simp add: bounded_closure compact_eq_bounded_closed) + moreover have "continuous_on (closure S) (cmod \ f)" + using contf continuous_on_compose continuous_on_norm_id by blast + ultimately obtain z where zin: "z \ closure S" and z: "\y. y \ closure S \ (cmod \ f) y \ (cmod \ f) z" + using continuous_attains_sup [of "closure S" "norm o f"] \\ \ S\ by auto + then consider "z \ frontier S" | "z \ interior S" using frontier_def by auto + then have "norm(f z) \ B" + proof cases + case 1 then show ?thesis using leB by blast + next + case 2 + have zin: "z \ connected_component_set (interior S) z" + by (simp add: 2) + have "f constant_on (connected_component_set (interior S) z)" + apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin]) + apply (metis connected_component_subset holf holomorphic_on_subset) + apply (simp_all add: open_connected_component) + by (metis closure_subset comp_eq_dest_lhs interior_subset subsetCE z connected_component_in) + then obtain c where c: "\w. w \ connected_component_set (interior S) z \ f w = c" + by (auto simp: constant_on_def) + have "f ` closure(connected_component_set (interior S) z) \ {c}" + apply (rule image_closure_subset) + apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) + using c + apply auto + done + then have cc: "\w. w \ closure(connected_component_set (interior S) z) \ f w = c" by blast + have "frontier(connected_component_set (interior S) z) \ {}" + apply (simp add: frontier_eq_empty) + by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV) + then obtain w where w: "w \ frontier(connected_component_set (interior S) z)" + by auto + then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def) + also have "... \ B" + apply (rule leB) + using w +using frontier_interior_subset frontier_of_connected_component_subset by blast + finally show ?thesis . + qed + then show ?thesis + using z \\ \ S\ closure_subset by fastforce +qed + +corollary\<^marker>\tag unimportant\ maximum_real_frontier: + assumes holf: "f holomorphic_on (interior S)" + and contf: "continuous_on (closure S) f" + and bos: "bounded S" + and leB: "\z. z \ frontier S \ Re(f z) \ B" + and "\ \ S" + shows "Re(f \) \ B" +using maximum_modulus_frontier [of "exp o f" S "exp B"] + Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms +by auto + +subsection\<^marker>\tag unimportant\ \Factoring out a zero according to its order\ + +lemma holomorphic_factor_order_of_zero: + assumes holf: "f holomorphic_on S" + and os: "open S" + and "\ \ S" "0 < n" + and dnz: "(deriv ^^ n) f \ \ 0" + and dfz: "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" + obtains g r where "0 < r" + "g holomorphic_on ball \ r" + "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" + "\w. w \ ball \ r \ g w \ 0" +proof - + obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) + then have holfb: "f holomorphic_on ball \ r" + using holf holomorphic_on_subset by blast + define g where "g w = suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" for w + have sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" + and feq: "f w - f \ = (w - \)^n * g w" + if w: "w \ ball \ r" for w + proof - + define powf where "powf = (\i. (deriv ^^ i) f \/(fact i) * (w - \)^i)" + have sing: "{.. = 0 then {} else {0})" + unfolding powf_def using \0 < n\ dfz by (auto simp: dfz; metis funpow_0 not_gr0) + have "powf sums f w" + unfolding powf_def by (rule holomorphic_power_series [OF holfb w]) + moreover have "(\i" + apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric]) + apply simp + apply (simp only: dfz sing) + apply (simp add: powf_def) + done + ultimately have fsums: "(\i. powf (i+n)) sums (f w - f \)" + using w sums_iff_shift' by metis + then have *: "summable (\i. (w - \) ^ n * ((deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n)))" + unfolding powf_def using sums_summable + by (auto simp: power_add mult_ac) + have "summable (\i. (deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n))" + proof (cases "w=\") + case False then show ?thesis + using summable_mult [OF *, of "1 / (w - \) ^ n"] by simp + next + case True then show ?thesis + by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"] + split: if_split_asm) + qed + then show sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" + by (simp add: summable_sums_iff g_def) + show "f w - f \ = (w - \)^n * g w" + apply (rule sums_unique2) + apply (rule fsums [unfolded powf_def]) + using sums_mult [OF sumsg, of "(w - \) ^ n"] + by (auto simp: power_add mult_ac) + qed + then have holg: "g holomorphic_on ball \ r" + by (meson sumsg power_series_holomorphic) + then have contg: "continuous_on (ball \ r) g" + by (blast intro: holomorphic_on_imp_continuous_on) + have "g \ \ 0" + using dnz unfolding g_def + by (subst suminf_finite [of "{0}"]) auto + obtain d where "0 < d" and d: "\w. w \ ball \ d \ g w \ 0" + apply (rule exE [OF continuous_on_avoid [OF contg _ \g \ \ 0\]]) + using \0 < r\ + apply force + by (metis \0 < r\ less_trans mem_ball not_less_iff_gr_or_eq) + show ?thesis + apply (rule that [where g=g and r ="min r d"]) + using \0 < r\ \0 < d\ holg + apply (auto simp: feq holomorphic_on_subset subset_ball d) + done +qed + + +lemma holomorphic_factor_order_of_zero_strong: + assumes holf: "f holomorphic_on S" "open S" "\ \ S" "0 < n" + and "(deriv ^^ n) f \ \ 0" + and "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" + obtains g r where "0 < r" + "g holomorphic_on ball \ r" + "\w. w \ ball \ r \ f w - f \ = ((w - \) * g w) ^ n" + "\w. w \ ball \ r \ g w \ 0" +proof - + obtain g r where "0 < r" + and holg: "g holomorphic_on ball \ r" + and feq: "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" + and gne: "\w. w \ ball \ r \ g w \ 0" + by (auto intro: holomorphic_factor_order_of_zero [OF assms]) + have con: "continuous_on (ball \ r) (\z. deriv g z / g z)" + by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on) + have cd: "\x. dist \ x < r \ (\z. deriv g z / g z) field_differentiable at x" + apply (rule derivative_intros)+ + using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) + apply (metis open_ball at_within_open holg holomorphic_on_def mem_ball) + using gne mem_ball by blast + obtain h where h: "\x. x \ ball \ r \ (h has_field_derivative deriv g x / g x) (at x)" + apply (rule exE [OF holomorphic_convex_primitive [of "ball \ r" "{}" "\z. deriv g z / g z"]]) + apply (auto simp: con cd) + apply (metis open_ball at_within_open mem_ball) + done + then have "continuous_on (ball \ r) h" + by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open) + then have con: "continuous_on (ball \ r) (\x. exp (h x) / g x)" + by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne) + have 0: "dist \ x < r \ ((\x. exp (h x) / g x) has_field_derivative 0) (at x)" for x + apply (rule h derivative_eq_intros | simp)+ + apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2]) + using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h) + done + obtain c where c: "\x. x \ ball \ r \ exp (h x) / g x = c" + by (rule DERIV_zero_connected_constant [of "ball \ r" "{}" "\x. exp(h x) / g x"]) (auto simp: con 0) + have hol: "(\z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \ r" + apply (rule holomorphic_on_compose [unfolded o_def, where g = exp]) + apply (rule holomorphic_intros)+ + using h holomorphic_on_open apply blast + apply (rule holomorphic_intros)+ + using \0 < n\ apply simp + apply (rule holomorphic_intros)+ + done + show ?thesis + apply (rule that [where g="\z. exp((Ln(inverse c) + h z)/n)" and r =r]) + using \0 < r\ \0 < n\ + apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric]) + apply (rule hol) + apply (simp add: Transcendental.exp_add gne) + done +qed + + +lemma + fixes k :: "'a::wellorder" + assumes a_def: "a == LEAST x. P x" and P: "P k" + shows def_LeastI: "P a" and def_Least_le: "a \ k" +unfolding a_def +by (rule LeastI Least_le; rule P)+ + +lemma holomorphic_factor_zero_nonconstant: + assumes holf: "f holomorphic_on S" and S: "open S" "connected S" + and "\ \ S" "f \ = 0" + and nonconst: "\ f constant_on S" + obtains g r n + where "0 < n" "0 < r" "ball \ r \ S" + "g holomorphic_on ball \ r" + "\w. w \ ball \ r \ f w = (w - \)^n * g w" + "\w. w \ ball \ r \ g w \ 0" +proof (cases "\n>0. (deriv ^^ n) f \ = 0") + case True then show ?thesis + using holomorphic_fun_eq_const_on_connected [OF holf S _ \\ \ S\] nonconst by (simp add: constant_on_def) +next + case False + then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \ \ 0" by blast + obtain r0 where "r0 > 0" "ball \ r0 \ S" using S openE \\ \ S\ by auto + define n where "n \ LEAST n. (deriv ^^ n) f \ \ 0" + have n_ne: "(deriv ^^ n) f \ \ 0" + by (rule def_LeastI [OF n_def]) (rule n0) + then have "0 < n" using \f \ = 0\ + using funpow_0 by fastforce + have n_min: "\k. k < n \ (deriv ^^ k) f \ = 0" + using def_Least_le [OF n_def] not_le by blast + then obtain g r1 + where "0 < r1" "g holomorphic_on ball \ r1" + "\w. w \ ball \ r1 \ f w = (w - \) ^ n * g w" + "\w. w \ ball \ r1 \ g w \ 0" + by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne] simp: \f \ = 0\) + then show ?thesis + apply (rule_tac g=g and r="min r0 r1" and n=n in that) + using \0 < n\ \0 < r0\ \0 < r1\ \ball \ r0 \ S\ + apply (auto simp: subset_ball intro: holomorphic_on_subset) + done +qed + + +lemma holomorphic_lower_bound_difference: + assumes holf: "f holomorphic_on S" and S: "open S" "connected S" + and "\ \ S" and "\ \ S" + and fne: "f \ \ f \" + obtains k n r + where "0 < k" "0 < r" + "ball \ r \ S" + "\w. w \ ball \ r \ k * norm(w - \)^n \ norm(f w - f \)" +proof - + define n where "n = (LEAST n. 0 < n \ (deriv ^^ n) f \ \ 0)" + obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \ \ 0" + using fne holomorphic_fun_eq_const_on_connected [OF holf S] \\ \ S\ \\ \ S\ by blast + then have "0 < n" and n_ne: "(deriv ^^ n) f \ \ 0" + unfolding n_def by (metis (mono_tags, lifting) LeastI)+ + have n_min: "\k. \0 < k; k < n\ \ (deriv ^^ k) f \ = 0" + unfolding n_def by (blast dest: not_less_Least) + then obtain g r + where "0 < r" and holg: "g holomorphic_on ball \ r" + and fne: "\w. w \ ball \ r \ f w - f \ = (w - \) ^ n * g w" + and gnz: "\w. w \ ball \ r \ g w \ 0" + by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne]) + obtain e where "e>0" and e: "ball \ e \ S" using assms by (blast elim!: openE) + then have holfb: "f holomorphic_on ball \ e" + using holf holomorphic_on_subset by blast + define d where "d = (min e r) / 2" + have "0 < d" using \0 < r\ \0 < e\ by (simp add: d_def) + have "d < r" + using \0 < r\ by (auto simp: d_def) + then have cbb: "cball \ d \ ball \ r" + by (auto simp: cball_subset_ball_iff) + then have "g holomorphic_on cball \ d" + by (rule holomorphic_on_subset [OF holg]) + then have "closed (g ` cball \ d)" + by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on) + moreover have "g ` cball \ d \ {}" + using \0 < d\ by auto + ultimately obtain x where x: "x \ g ` cball \ d" and "\y. y \ g ` cball \ d \ dist 0 x \ dist 0 y" + by (rule distance_attains_inf) blast + then have leg: "\w. w \ cball \ d \ norm x \ norm (g w)" + by auto + have "ball \ d \ cball \ d" by auto + also have "... \ ball \ e" using \0 < d\ d_def by auto + also have "... \ S" by (rule e) + finally have dS: "ball \ d \ S" . + moreover have "x \ 0" using gnz x \d < r\ by auto + ultimately show ?thesis + apply (rule_tac k="norm x" and n=n and r=d in that) + using \d < r\ leg + apply (auto simp: \0 < d\ fne norm_mult norm_power algebra_simps mult_right_mono) + done +qed + +lemma + assumes holf: "f holomorphic_on (S - {\})" and \: "\ \ interior S" + shows holomorphic_on_extend_lim: + "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ + ((\z. (z - \) * f z) \ 0) (at \)" + (is "?P = ?Q") + and holomorphic_on_extend_bounded: + "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ + (\B. eventually (\z. norm(f z) \ B) (at \))" + (is "?P = ?R") +proof - + obtain \ where "0 < \" and \: "ball \ \ \ S" + using \ mem_interior by blast + have "?R" if holg: "g holomorphic_on S" and gf: "\z. z \ S - {\} \ g z = f z" for g + proof - + have *: "\\<^sub>F z in at \. dist (g z) (g \) < 1 \ cmod (f z) \ cmod (g \) + 1" + apply (simp add: eventually_at) + apply (rule_tac x="\" in exI) + using \ \0 < \\ + apply (clarsimp simp:) + apply (drule_tac c=x in subsetD) + apply (simp add: dist_commute) + by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD) + have "continuous_on (interior S) g" + by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset) + then have "\x. x \ interior S \ (g \ g x) (at x)" + using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast + then have "(g \ g \) (at \)" + by (simp add: \) + then show ?thesis + apply (rule_tac x="norm(g \) + 1" in exI) + apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) + done + qed + moreover have "?Q" if "\\<^sub>F z in at \. cmod (f z) \ B" for B + by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) + moreover have "?P" if "(\z. (z - \) * f z) \\\ 0" + proof - + define h where [abs_def]: "h z = (z - \)^2 * f z" for z + have h0: "(h has_field_derivative 0) (at \)" + apply (simp add: h_def has_field_derivative_iff) + apply (rule Lim_transform_within [OF that, of 1]) + apply (auto simp: field_split_simps power2_eq_square) + done + have holh: "h holomorphic_on S" + proof (simp add: holomorphic_on_def, clarify) + fix z assume "z \ S" + show "h field_differentiable at z within S" + proof (cases "z = \") + case True then show ?thesis + using field_differentiable_at_within field_differentiable_def h0 by blast + next + case False + then have "f field_differentiable at z within S" + using holomorphic_onD [OF holf, of z] \z \ S\ + unfolding field_differentiable_def has_field_derivative_iff + by (force intro: exI [where x="dist \ z"] elim: Lim_transform_within_set [unfolded eventually_at]) + then show ?thesis + by (simp add: h_def power2_eq_square derivative_intros) + qed + qed + define g where [abs_def]: "g z = (if z = \ then deriv h \ else (h z - h \) / (z - \))" for z + have holg: "g holomorphic_on S" + unfolding g_def by (rule pole_lemma [OF holh \]) + show ?thesis + apply (rule_tac x="\z. if z = \ then deriv g \ else (g z - g \)/(z - \)" in exI) + apply (rule conjI) + apply (rule pole_lemma [OF holg \]) + apply (auto simp: g_def power2_eq_square divide_simps) + using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square) + done + qed + ultimately show "?P = ?Q" and "?P = ?R" + by meson+ +qed + +lemma pole_at_infinity: + assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \ l) at_infinity" + obtains a n where "\z. f z = (\i\n. a i * z^i)" +proof (cases "l = 0") + case False + with tendsto_inverse [OF lim] show ?thesis + apply (rule_tac a="(\n. inverse l)" and n=0 in that) + apply (simp add: Liouville_weak [OF holf, of "inverse l"]) + done +next + case True + then have [simp]: "l = 0" . + show ?thesis + proof (cases "\r. 0 < r \ (\z \ ball 0 r - {0}. f(inverse z) \ 0)") + case True + then obtain r where "0 < r" and r: "\z. z \ ball 0 r - {0} \ f(inverse z) \ 0" + by auto + have 1: "inverse \ f \ inverse holomorphic_on ball 0 r - {0}" + by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+ + have 2: "0 \ interior (ball 0 r)" + using \0 < r\ by simp + have "\B. 0 eventually (\z. cmod ((inverse \ f \ inverse) z) \ B) (at 0)" + apply (rule exI [where x=1]) + apply simp + using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] + apply (rule eventually_mono) + apply (simp add: dist_norm) + done + with holomorphic_on_extend_bounded [OF 1 2] + obtain g where holg: "g holomorphic_on ball 0 r" + and geq: "\z. z \ ball 0 r - {0} \ g z = (inverse \ f \ inverse) z" + by meson + have ifi0: "(inverse \ f \ inverse) \0\ 0" + using \l = 0\ lim lim_at_infinity_0 by blast + have g2g0: "g \0\ g 0" + using \0 < r\ centre_in_ball continuous_at continuous_on_eq_continuous_at holg + by (blast intro: holomorphic_on_imp_continuous_on) + have g2g1: "g \0\ 0" + apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) + using \0 < r\ by (auto simp: geq) + have [simp]: "g 0 = 0" + by (rule tendsto_unique [OF _ g2g0 g2g1]) simp + have "ball 0 r - {0::complex} \ {}" + using \0 < r\ + apply (clarsimp simp: ball_def dist_norm) + apply (drule_tac c="of_real r/2" in subsetD, auto) + done + then obtain w::complex where "w \ 0" and w: "norm w < r" by force + then have "g w \ 0" by (simp add: geq r) + obtain B n e where "0 < B" "0 < e" "e \ r" + and leg: "\w. norm w < e \ B * cmod w ^ n \ cmod (g w)" + apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w]) + using \0 < r\ w \g w \ 0\ by (auto simp: ball_subset_ball_iff) + have "cmod (f z) \ cmod z ^ n / B" if "2/e \ cmod z" for z + proof - + have ize: "inverse z \ ball 0 e - {0}" using that \0 < e\ + by (auto simp: norm_divide field_split_simps algebra_simps) + then have [simp]: "z \ 0" and izr: "inverse z \ ball 0 r - {0}" using \e \ r\ + by auto + then have [simp]: "f z \ 0" + using r [of "inverse z"] by simp + have [simp]: "f z = inverse (g (inverse z))" + using izr geq [of "inverse z"] by simp + show ?thesis using ize leg [of "inverse z"] \0 < B\ \0 < e\ + by (simp add: field_split_simps norm_divide algebra_simps) + qed + then show ?thesis + apply (rule_tac a = "\k. (deriv ^^ k) f 0 / (fact k)" and n=n in that) + apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp) + done + next + case False + then have fi0: "\r. r > 0 \ \z\ball 0 r - {0}. f (inverse z) = 0" + by simp + have fz0: "f z = 0" if "0 < r" and lt1: "\x. x \ 0 \ cmod x < r \ inverse (cmod (f (inverse x))) < 1" + for z r + proof - + have f0: "(f \ 0) at_infinity" + proof - + have DIM_complex[intro]: "2 \ DIM(complex)" \ \should not be necessary!\ + by simp + have "f (inverse x) \ 0 \ x \ 0 \ cmod x < r \ 1 < cmod (f (inverse x))" for x + using lt1[of x] by (auto simp: field_simps) + then have **: "cmod (f (inverse x)) \ 1 \ x \ 0 \ cmod x < r \ f (inverse x) = 0" for x + by force + then have *: "(f \ inverse) ` (ball 0 r - {0}) \ {0} \ - ball 0 1" + by force + have "continuous_on (inverse ` (ball 0 r - {0})) f" + using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast + then have "connected ((f \ inverse) ` (ball 0 r - {0}))" + apply (intro connected_continuous_image continuous_intros) + apply (force intro: connected_punctured_ball)+ + done + then have "{0} \ (f \ inverse) ` (ball 0 r - {0}) = {} \ - ball 0 1 \ (f \ inverse) ` (ball 0 r - {0}) = {}" + by (rule connected_closedD) (use * in auto) + then have "w \ 0 \ cmod w < r \ f (inverse w) = 0" for w + using fi0 **[of w] \0 < r\ + apply (auto simp add: inf.commute [of "- ball 0 1"] Diff_eq [symmetric] image_subset_iff dest: less_imp_le) + apply fastforce + apply (drule bspec [of _ _ w]) + apply (auto dest: less_imp_le) + done + then show ?thesis + apply (simp add: lim_at_infinity_0) + apply (rule tendsto_eventually) + apply (simp add: eventually_at) + apply (rule_tac x=r in exI) + apply (simp add: \0 < r\ dist_norm) + done + qed + obtain w where "w \ ball 0 r - {0}" and "f (inverse w) = 0" + using False \0 < r\ by blast + then show ?thesis + by (auto simp: f0 Liouville_weak [OF holf, of 0]) + qed + show ?thesis + apply (rule that [of "\n. 0" 0]) + using lim [unfolded lim_at_infinity_0] + apply (simp add: Lim_at dist_norm norm_inverse) + apply (drule_tac x=1 in spec) + using fz0 apply auto + done + qed +qed + +subsection\<^marker>\tag unimportant\ \Entire proper functions are precisely the non-trivial polynomials\ + +lemma proper_map_polyfun: + fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" + assumes "closed S" and "compact K" and c: "c i \ 0" "1 \ i" "i \ n" + shows "compact (S \ {z. (\i\n. c i * z^i) \ K})" +proof - + obtain B where "B > 0" and B: "\x. x \ K \ norm x \ B" + by (metis compact_imp_bounded \compact K\ bounded_pos) + have *: "norm x \ b" + if "\x. b \ norm x \ B + 1 \ norm (\i\n. c i * x ^ i)" + "(\i\n. c i * x ^ i) \ K" for b x + proof - + have "norm (\i\n. c i * x ^ i) \ B" + using B that by blast + moreover have "\ B + 1 \ B" + by simp + ultimately show "norm x \ b" + using that by (metis (no_types) less_eq_real_def not_less order_trans) + qed + have "bounded {z. (\i\n. c i * z ^ i) \ K}" + using Limits.polyfun_extremal [where c=c and B="B+1", OF c] + by (auto simp: bounded_pos eventually_at_infinity_pos *) + moreover have "closed ((\z. (\i\n. c i * z ^ i)) -` K)" + apply (intro allI continuous_closed_vimage continuous_intros) + using \compact K\ compact_eq_bounded_closed by blast + ultimately show ?thesis + using closed_Int_compact [OF \closed S\] compact_eq_bounded_closed + by (auto simp add: vimage_def) +qed + +lemma proper_map_polyfun_univ: + fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" + assumes "compact K" "c i \ 0" "1 \ i" "i \ n" + shows "compact ({z. (\i\n. c i * z^i) \ K})" + using proper_map_polyfun [of UNIV K c i n] assms by simp + +lemma proper_map_polyfun_eq: + assumes "f holomorphic_on UNIV" + shows "(\k. compact k \ compact {z. f z \ k}) \ + (\c n. 0 < n \ (c n \ 0) \ f = (\z. \i\n. c i * z^i))" + (is "?lhs = ?rhs") +proof + assume compf [rule_format]: ?lhs + have 2: "\k. 0 < k \ a k \ 0 \ f = (\z. \i \ k. a i * z ^ i)" + if "\z. f z = (\i\n. a i * z ^ i)" for a n + proof (cases "\i\n. 0 a i = 0") + case True + then have [simp]: "\z. f z = a 0" + by (simp add: that sum.atMost_shift) + have False using compf [of "{a 0}"] by simp + then show ?thesis .. + next + case False + then obtain k where k: "0 < k" "k\n" "a k \ 0" by force + define m where "m = (GREATEST k. k\n \ a k \ 0)" + have m: "m\n \ a m \ 0" + unfolding m_def + apply (rule GreatestI_nat [where b = n]) + using k apply auto + done + have [simp]: "a i = 0" if "m < i" "i \ n" for i + using Greatest_le_nat [where b = "n" and P = "\k. k\n \ a k \ 0"] + using m_def not_le that by auto + have "k \ m" + unfolding m_def + apply (rule Greatest_le_nat [where b = "n"]) + using k apply auto + done + with k m show ?thesis + by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right) + qed + have "((inverse \ f) \ 0) at_infinity" + proof (rule Lim_at_infinityI) + fix e::real assume "0 < e" + with compf [of "cball 0 (inverse e)"] + show "\B. \x. B \ cmod x \ dist ((inverse \ f) x) 0 \ e" + apply simp + apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) + apply (rule_tac x="b+1" in exI) + apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one) + done + qed + then show ?rhs + apply (rule pole_at_infinity [OF assms]) + using 2 apply blast + done +next + assume ?rhs + then obtain c n where "0 < n" "c n \ 0" "f = (\z. \i\n. c i * z ^ i)" by blast + then have "compact {z. f z \ k}" if "compact k" for k + by (auto intro: proper_map_polyfun_univ [OF that]) + then show ?lhs by blast +qed + +subsection \Relating invertibility and nonvanishing of derivative\ + +lemma has_complex_derivative_locally_injective: + assumes holf: "f holomorphic_on S" + and S: "\ \ S" "open S" + and dnz: "deriv f \ \ 0" + obtains r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" +proof - + have *: "\d>0. \x. dist \ x < d \ onorm (\v. deriv f x * v - deriv f \ * v) < e" if "e > 0" for e + proof - + have contdf: "continuous_on S (deriv f)" + by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open S\) + obtain \ where "\>0" and \: "\x. \x \ S; dist x \ \ \\ \ cmod (deriv f x - deriv f \) \ e/2" + using continuous_onE [OF contdf \\ \ S\, of "e/2"] \0 < e\ + by (metis dist_complex_def half_gt_zero less_imp_le) + obtain \ where "\>0" "ball \ \ \ S" + by (metis openE [OF \open S\ \\ \ S\]) + with \\>0\ have "\\>0. \x. dist \ x < \ \ onorm (\v. deriv f x * v - deriv f \ * v) \ e/2" + apply (rule_tac x="min \ \" in exI) + apply (intro conjI allI impI Operator_Norm.onorm_le) + apply simp + apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult) + apply (rule mult_right_mono [OF \]) + apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \) + done + with \e>0\ show ?thesis by force + qed + have "inj ((*) (deriv f \))" + using dnz by simp + then obtain g' where g': "linear g'" "g' \ (*) (deriv f \) = id" + using linear_injective_left_inverse [of "(*) (deriv f \)"] + by (auto simp: linear_times) + show ?thesis + apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\z h. deriv f z * h" and g' = g']) + using g' * + apply (simp_all add: linear_conv_bounded_linear that) + using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf + holomorphic_on_imp_differentiable_at \open S\ apply blast + done +qed + +lemma has_complex_derivative_locally_invertible: + assumes holf: "f holomorphic_on S" + and S: "\ \ S" "open S" + and dnz: "deriv f \ \ 0" + obtains r where "r > 0" "ball \ r \ S" "open (f ` (ball \ r))" "inj_on f (ball \ r)" +proof - + obtain r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" + by (blast intro: that has_complex_derivative_locally_injective [OF assms]) + then have \: "\ \ ball \ r" by simp + then have nc: "\ f constant_on ball \ r" + using \inj_on f (ball \ r)\ injective_not_constant by fastforce + have holf': "f holomorphic_on ball \ r" + using \ball \ r \ S\ holf holomorphic_on_subset by blast + have "open (f ` ball \ r)" + apply (rule open_mapping_thm [OF holf']) + using nc apply auto + done + then show ?thesis + using \0 < r\ \ball \ r \ S\ \inj_on f (ball \ r)\ that by blast +qed + +lemma holomorphic_injective_imp_regular: + assumes holf: "f holomorphic_on S" + and "open S" and injf: "inj_on f S" + and "\ \ S" + shows "deriv f \ \ 0" +proof - + obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) + have holf': "f holomorphic_on ball \ r" + using \ball \ r \ S\ holf holomorphic_on_subset by blast + show ?thesis + proof (cases "\n>0. (deriv ^^ n) f \ = 0") + case True + have fcon: "f w = f \" if "w \ ball \ r" for w + apply (rule holomorphic_fun_eq_const_on_connected [OF holf']) + using True \0 < r\ that by auto + have False + using fcon [of "\ + r/2"] \0 < r\ r injf unfolding inj_on_def + by (metis \\ \ S\ contra_subsetD dist_commute fcon mem_ball perfect_choose_dist) + then show ?thesis .. + next + case False + then obtain n0 where n0: "n0 > 0 \ (deriv ^^ n0) f \ \ 0" by blast + define n where [abs_def]: "n = (LEAST n. n > 0 \ (deriv ^^ n) f \ \ 0)" + have n_ne: "n > 0" "(deriv ^^ n) f \ \ 0" + using def_LeastI [OF n_def n0] by auto + have n_min: "\k. 0 < k \ k < n \ (deriv ^^ k) f \ = 0" + using def_Least_le [OF n_def] not_le by auto + obtain g \ where "0 < \" + and holg: "g holomorphic_on ball \ \" + and fd: "\w. w \ ball \ \ \ f w - f \ = ((w - \) * g w) ^ n" + and gnz: "\w. w \ ball \ \ \ g w \ 0" + apply (rule holomorphic_factor_order_of_zero_strong [OF holf \open S\ \\ \ S\ n_ne]) + apply (blast intro: n_min)+ + done + show ?thesis + proof (cases "n=1") + case True + with n_ne show ?thesis by auto + next + case False + have holgw: "(\w. (w - \) * g w) holomorphic_on ball \ (min r \)" + apply (rule holomorphic_intros)+ + using holg by (simp add: holomorphic_on_subset subset_ball) + have gd: "\w. dist \ w < \ \ (g has_field_derivative deriv g w) (at w)" + using holg + by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH) + have *: "\w. w \ ball \ (min r \) + \ ((\w. (w - \) * g w) has_field_derivative ((w - \) * deriv g w + g w)) + (at w)" + by (rule gd derivative_eq_intros | simp)+ + have [simp]: "deriv (\w. (w - \) * g w) \ \ 0" + using * [of \] \0 < \\ \0 < r\ by (simp add: DERIV_imp_deriv gnz) + obtain T where "\ \ T" "open T" and Tsb: "T \ ball \ (min r \)" and oimT: "open ((\w. (w - \) * g w) ` T)" + apply (rule has_complex_derivative_locally_invertible [OF holgw, of \]) + using \0 < r\ \0 < \\ + apply (simp_all add:) + by (meson open_ball centre_in_ball) + define U where "U = (\w. (w - \) * g w) ` T" + have "open U" by (metis oimT U_def) + have "0 \ U" + apply (auto simp: U_def) + apply (rule image_eqI [where x = \]) + apply (auto simp: \\ \ T\) + done + then obtain \ where "\>0" and \: "cball 0 \ \ U" + using \open U\ open_contains_cball by blast + then have "\ * exp(2 * of_real pi * \ * (0/n)) \ cball 0 \" + "\ * exp(2 * of_real pi * \ * (1/n)) \ cball 0 \" + by (auto simp: norm_mult) + with \ have "\ * exp(2 * of_real pi * \ * (0/n)) \ U" + "\ * exp(2 * of_real pi * \ * (1/n)) \ U" by blast+ + then obtain y0 y1 where "y0 \ T" and y0: "(y0 - \) * g y0 = \ * exp(2 * of_real pi * \ * (0/n))" + and "y1 \ T" and y1: "(y1 - \) * g y1 = \ * exp(2 * of_real pi * \ * (1/n))" + by (auto simp: U_def) + then have "y0 \ ball \ \" "y1 \ ball \ \" using Tsb by auto + moreover have "y0 \ y1" + using y0 y1 \\ > 0\ complex_root_unity_eq_1 [of n 1] \n > 0\ False by auto + moreover have "T \ S" + by (meson Tsb min.cobounded1 order_trans r subset_ball) + ultimately have False + using inj_onD [OF injf, of y0 y1] \y0 \ T\ \y1 \ T\ + using fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne + apply (simp add: y0 y1 power_mult_distrib) + apply (force simp: algebra_simps) + done + then show ?thesis .. + qed + qed +qed + +text\Hence a nice clean inverse function theorem\ + +lemma has_field_derivative_inverse_strong: + fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" + shows "\DERIV f x :> f'; f' \ 0; open S; x \ S; continuous_on S f; + \z. z \ S \ g (f z) = z\ + \ DERIV g (f x) :> inverse (f')" + unfolding has_field_derivative_def + by (rule has_derivative_inverse_strong [of S x f g]) auto + +lemma has_field_derivative_inverse_strong_x: + fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" + shows "\DERIV f (g y) :> f'; f' \ 0; open S; continuous_on S f; g y \ S; f(g y) = y; + \z. z \ S \ g (f z) = z\ + \ DERIV g y :> inverse (f')" + unfolding has_field_derivative_def + by (rule has_derivative_inverse_strong_x [of S g y f]) auto + +proposition holomorphic_has_inverse: + assumes holf: "f holomorphic_on S" + and "open S" and injf: "inj_on f S" + obtains g where "g holomorphic_on (f ` S)" + "\z. z \ S \ deriv f z * deriv g (f z) = 1" + "\z. z \ S \ g(f z) = z" +proof - + have ofs: "open (f ` S)" + by (rule open_mapping_thm3 [OF assms]) + have contf: "continuous_on S f" + by (simp add: holf holomorphic_on_imp_continuous_on) + have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \ S" for z + proof - + have 1: "(f has_field_derivative deriv f z) (at z)" + using DERIV_deriv_iff_field_differentiable \z \ S\ \open S\ holf holomorphic_on_imp_differentiable_at + by blast + have 2: "deriv f z \ 0" + using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast + show ?thesis + apply (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) + apply (simp add: holf holomorphic_on_imp_continuous_on) + by (simp add: injf the_inv_into_f_f) + qed + show ?thesis + proof + show "the_inv_into S f holomorphic_on f ` S" + by (simp add: holomorphic_on_open ofs) (blast intro: *) + next + fix z assume "z \ S" + have "deriv f z \ 0" + using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast + then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" + using * [OF \z \ S\] by (simp add: DERIV_imp_deriv) + next + fix z assume "z \ S" + show "the_inv_into S f (f z) = z" + by (simp add: \z \ S\ injf the_inv_into_f_f) + qed +qed + +subsection\The Schwarz Lemma\ + +lemma Schwarz1: + assumes holf: "f holomorphic_on S" + and contf: "continuous_on (closure S) f" + and S: "open S" "connected S" + and boS: "bounded S" + and "S \ {}" + obtains w where "w \ frontier S" + "\z. z \ closure S \ norm (f z) \ norm (f w)" +proof - + have connf: "continuous_on (closure S) (norm o f)" + using contf continuous_on_compose continuous_on_norm_id by blast + have coc: "compact (closure S)" + by (simp add: \bounded S\ bounded_closure compact_eq_bounded_closed) + then obtain x where x: "x \ closure S" and xmax: "\z. z \ closure S \ norm(f z) \ norm(f x)" + apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]]) + using \S \ {}\ apply auto + done + then show ?thesis + proof (cases "x \ frontier S") + case True + then show ?thesis using that xmax by blast + next + case False + then have "x \ S" + using \open S\ frontier_def interior_eq x by auto + then have "f constant_on S" + apply (rule maximum_modulus_principle [OF holf S \open S\ order_refl]) + using closure_subset apply (blast intro: xmax) + done + then have "f constant_on (closure S)" + by (rule constant_on_closureI [OF _ contf]) + then obtain c where c: "\x. x \ closure S \ f x = c" + by (meson constant_on_def) + obtain w where "w \ frontier S" + by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) + then show ?thesis + by (simp add: c frontier_def that) + qed +qed + +lemma Schwarz2: + "\f holomorphic_on ball 0 r; + 0 < s; ball w s \ ball 0 r; + \z. norm (w-z) < s \ norm(f z) \ norm(f w)\ + \ f constant_on ball 0 r" +by (rule maximum_modulus_principle [where U = "ball w s" and \ = w]) (simp_all add: dist_norm) + +lemma Schwarz3: + assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" + obtains h where "h holomorphic_on (ball 0 r)" and "\z. norm z < r \ f z = z * (h z)" and "deriv f 0 = h 0" +proof - + define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z + have d0: "deriv f 0 = h 0" + by (simp add: h_def) + moreover have "h holomorphic_on (ball 0 r)" + by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) + moreover have "norm z < r \ f z = z * h z" for z + by (simp add: h_def) + ultimately show ?thesis + using that by blast +qed + +proposition Schwarz_Lemma: + assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" + and no: "\z. norm z < 1 \ norm (f z) < 1" + and \: "norm \ < 1" + shows "norm (f \) \ norm \" and "norm(deriv f 0) \ 1" + and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) + \ norm(deriv f 0) = 1) + \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" + (is "?P \ ?Q") +proof - + obtain h where holh: "h holomorphic_on (ball 0 1)" + and fz_eq: "\z. norm z < 1 \ f z = z * (h z)" and df0: "deriv f 0 = h 0" + by (rule Schwarz3 [OF holf]) auto + have noh_le: "norm (h z) \ 1" if z: "norm z < 1" for z + proof - + have "norm (h z) < a" if a: "1 < a" for a + proof - + have "max (inverse a) (norm z) < 1" + using z a by (simp_all add: inverse_less_1_iff) + then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" + using Rats_dense_in_real by blast + then have nzr: "norm z < r" and ira: "inverse r < a" + using z a less_imp_inverse_less by force+ + then have "0 < r" + by (meson norm_not_less_zero not_le order.strict_trans2) + have holh': "h holomorphic_on ball 0 r" + by (meson holh \r < 1\ holomorphic_on_subset less_eq_real_def subset_ball) + have conth': "continuous_on (cball 0 r) h" + by (meson \r < 1\ dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) + obtain w where w: "norm w = r" and lenw: "\z. norm z < r \ norm(h z) \ norm(h w)" + apply (rule Schwarz1 [OF holh']) using conth' \0 < r\ by auto + have "h w = f w / w" using fz_eq \r < 1\ nzr w by auto + then have "cmod (h z) < inverse r" + by (metis \0 < r\ \r < 1\ divide_strict_right_mono inverse_eq_divide + le_less_trans lenw no norm_divide nzr w) + then show ?thesis using ira by linarith + qed + then show "norm (h z) \ 1" + using not_le by blast + qed + show "cmod (f \) \ cmod \" + proof (cases "\ = 0") + case True then show ?thesis by auto + next + case False + then show ?thesis + by (simp add: noh_le fz_eq \ mult_left_le norm_mult) + qed + show no_df0: "norm(deriv f 0) \ 1" + by (simp add: \\z. cmod z < 1 \ cmod (h z) \ 1\ df0) + show "?Q" if "?P" + using that + proof + assume "\z. cmod z < 1 \ z \ 0 \ cmod (f z) = cmod z" + then obtain \ where \: "cmod \ < 1" "\ \ 0" "cmod (f \) = cmod \" by blast + then have [simp]: "norm (h \) = 1" + by (simp add: fz_eq norm_mult) + have "ball \ (1 - cmod \) \ ball 0 1" + by (simp add: ball_subset_ball_iff) + moreover have "\z. cmod (\ - z) < 1 - cmod \ \ cmod (h z) \ cmod (h \)" + apply (simp add: algebra_simps) + by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) + ultimately obtain c where c: "\z. norm z < 1 \ h z = c" + using Schwarz2 [OF holh, of "1 - norm \" \, unfolded constant_on_def] \ by auto + then have "norm c = 1" + using \ by force + with c show ?thesis + using fz_eq by auto + next + assume [simp]: "cmod (deriv f 0) = 1" + then obtain c where c: "\z. norm z < 1 \ h z = c" + using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le + by auto + moreover have "norm c = 1" using df0 c by auto + ultimately show ?thesis + using fz_eq by auto + qed +qed + +corollary Schwarz_Lemma': + assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" + and no: "\z. norm z < 1 \ norm (f z) < 1" + shows "((\\. norm \ < 1 \ norm (f \) \ norm \) + \ norm(deriv f 0) \ 1) + \ (((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) + \ norm(deriv f 0) = 1) + \ (\\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1))" + using Schwarz_Lemma [OF assms] + by (metis (no_types) norm_eq_zero zero_less_one) + +subsection\The Schwarz reflection principle\ + +lemma hol_pal_lem0: + assumes "d \ a \ k" "k \ d \ b" + obtains c where + "c \ closed_segment a b" "d \ c = k" + "\z. z \ closed_segment a c \ d \ z \ k" + "\z. z \ closed_segment c b \ k \ d \ z" +proof - + obtain c where cin: "c \ closed_segment a b" and keq: "k = d \ c" + using connected_ivt_hyperplane [of "closed_segment a b" a b d k] + by (auto simp: assms) + have "closed_segment a c \ {z. d \ z \ k}" "closed_segment c b \ {z. k \ d \ z}" + unfolding segment_convex_hull using assms keq + by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) + then show ?thesis using cin that by fastforce +qed + +lemma hol_pal_lem1: + assumes "convex S" "open S" + and abc: "a \ S" "b \ S" "c \ S" + "d \ 0" and lek: "d \ a \ k" "d \ b \ k" "d \ c \ k" + and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof - + have "interior (convex hull {a, b, c}) \ interior(S \ {x. d \ x \ k})" + apply (rule interior_mono) + apply (rule hull_minimal) + apply (simp add: abc lek) + apply (rule convex_Int [OF \convex S\ convex_halfspace_le]) + done + also have "... \ {z \ S. d \ z < k}" + by (force simp: interior_open [OF \open S\] \d \ 0\) + finally have *: "interior (convex hull {a, b, c}) \ {z \ S. d \ z < k}" . + have "continuous_on (convex hull {a,b,c}) f" + using \convex S\ contf abc continuous_on_subset subset_hull + by fastforce + moreover have "f holomorphic_on interior (convex hull {a,b,c})" + by (rule holomorphic_on_subset [OF holf1 *]) + ultimately show ?thesis + using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 + by blast +qed + +lemma hol_pal_lem2: + assumes S: "convex S" "open S" + and abc: "a \ S" "b \ S" "c \ S" + and "d \ 0" and lek: "d \ a \ k" "d \ b \ k" + and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" + and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "d \ c \ k") + case True show ?thesis + by (rule hol_pal_lem1 [OF S abc \d \ 0\ lek True holf1 contf]) +next + case False + then have "d \ c > k" by force + obtain a' where a': "a' \ closed_segment b c" and "d \ a' = k" + and ba': "\z. z \ closed_segment b a' \ d \ z \ k" + and a'c: "\z. z \ closed_segment a' c \ k \ d \ z" + apply (rule hol_pal_lem0 [of d b k c, OF \d \ b \ k\]) + using False by auto + obtain b' where b': "b' \ closed_segment a c" and "d \ b' = k" + and ab': "\z. z \ closed_segment a b' \ d \ z \ k" + and b'c: "\z. z \ closed_segment b' c \ k \ d \ z" + apply (rule hol_pal_lem0 [of d a k c, OF \d \ a \ k\]) + using False by auto + have a'b': "a' \ S \ b' \ S" + using a' abc b' convex_contains_segment \convex S\ by auto + have "continuous_on (closed_segment c a) f" + by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) + then have 1: "contour_integral (linepath c a) f = + contour_integral (linepath c b') f + contour_integral (linepath b' a) f" + apply (rule contour_integral_split_linepath) + using b' by (simp add: closed_segment_commute) + have "continuous_on (closed_segment b c) f" + by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) + then have 2: "contour_integral (linepath b c) f = + contour_integral (linepath b a') f + contour_integral (linepath a' c) f" + by (rule contour_integral_split_linepath [OF _ a']) + have 3: "contour_integral (reversepath (linepath b' a')) f = + - contour_integral (linepath b' a') f" + by (rule contour_integral_reversepath [OF valid_path_linepath]) + have fcd_le: "f field_differentiable at x" + if "x \ interior S \ x \ interior {x. d \ x \ k}" for x + proof - + have "f holomorphic_on S \ {c. d \ c < k}" + by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) + then have "\C D. x \ interior C \ interior D \ f holomorphic_on interior C \ interior D" + using that + by (metis Collect_mem_eq Int_Collect \d \ 0\ interior_halfspace_le interior_open \open S\) + then show "f field_differentiable at x" + by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) + qed + have ab_le: "\x. x \ closed_segment a b \ d \ x \ k" + proof - + fix x :: complex + assume "x \ closed_segment a b" + then have "\C. x \ C \ b \ C \ a \ C \ \ convex C" + by (meson contra_subsetD convex_contains_segment) + then show "d \ x \ k" + by (metis lek convex_halfspace_le mem_Collect_eq) + qed + have "continuous_on (S \ {x. d \ x \ k}) f" using contf + by (simp add: continuous_on_subset) + then have "(f has_contour_integral 0) + (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" + apply (rule Cauchy_theorem_convex [where K = "{}"]) + apply (simp_all add: path_image_join convex_Int convex_halfspace_le \convex S\ fcd_le ab_le + closed_segment_subset abc a'b' ba') + by (metis \d \ a' = k\ \d \ b' = k\ convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl) + then have 4: "contour_integral (linepath a b) f + + contour_integral (linepath b a') f + + contour_integral (linepath a' b') f + + contour_integral (linepath b' a) f = 0" + by (rule has_chain_integral_chain_integral4) + have fcd_ge: "f field_differentiable at x" + if "x \ interior S \ x \ interior {x. k \ d \ x}" for x + proof - + have f2: "f holomorphic_on S \ {c. k < d \ c}" + by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) + have f3: "interior S = S" + by (simp add: interior_open \open S\) + then have "x \ S \ interior {c. k \ d \ c}" + using that by simp + then show "f field_differentiable at x" + using f3 f2 unfolding holomorphic_on_def + by (metis (no_types) \d \ 0\ at_within_interior interior_Int interior_halfspace_ge interior_interior) + qed + have "continuous_on (S \ {x. k \ d \ x}) f" using contf + by (simp add: continuous_on_subset) + then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" + apply (rule Cauchy_theorem_convex [where K = "{}"]) + apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \convex S\ + fcd_ge closed_segment_subset abc a'b' a'c) + by (metis \d \ a' = k\ b'c closed_segment_commute convex_contains_segment + convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl) + then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" + by (rule has_chain_integral_chain_integral3) + show ?thesis + using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) +qed + +lemma hol_pal_lem3: + assumes S: "convex S" "open S" + and abc: "a \ S" "b \ S" "c \ S" + and "d \ 0" and lek: "d \ a \ k" + and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" + and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "d \ b \ k") + case True show ?thesis + by (rule hol_pal_lem2 [OF S abc \d \ 0\ lek True holf1 holf2 contf]) +next + case False + show ?thesis + proof (cases "d \ c \ k") + case True + have "contour_integral (linepath c a) f + + contour_integral (linepath a b) f + + contour_integral (linepath b c) f = 0" + by (rule hol_pal_lem2 [OF S \c \ S\ \a \ S\ \b \ S\ \d \ 0\ \d \ c \ k\ lek holf1 holf2 contf]) + then show ?thesis + by (simp add: algebra_simps) + next + case False + have "contour_integral (linepath b c) f + + contour_integral (linepath c a) f + + contour_integral (linepath a b) f = 0" + apply (rule hol_pal_lem2 [OF S \b \ S\ \c \ S\ \a \ S\, of "-d" "-k"]) + using \d \ 0\ \\ d \ b \ k\ False by (simp_all add: holf1 holf2 contf) + then show ?thesis + by (simp add: algebra_simps) + qed +qed + +lemma hol_pal_lem4: + assumes S: "convex S" "open S" + and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" + and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" + and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" + and contf: "continuous_on S f" + shows "contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0" +proof (cases "d \ a \ k") + case True show ?thesis + by (rule hol_pal_lem3 [OF S abc \d \ 0\ True holf1 holf2 contf]) +next + case False + show ?thesis + apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"]) + using \d \ 0\ False by (simp_all add: holf1 holf2 contf) +qed + +lemma holomorphic_on_paste_across_line: + assumes S: "open S" and "d \ 0" + and holf1: "f holomorphic_on (S \ {z. d \ z < k})" + and holf2: "f holomorphic_on (S \ {z. k < d \ z})" + and contf: "continuous_on S f" + shows "f holomorphic_on S" +proof - + have *: "\t. open t \ p \ t \ continuous_on t f \ + (\a b c. convex hull {a, b, c} \ t \ + contour_integral (linepath a b) f + + contour_integral (linepath b c) f + + contour_integral (linepath c a) f = 0)" + if "p \ S" for p + proof - + obtain e where "e>0" and e: "ball p e \ S" + using \p \ S\ openE S by blast + then have "continuous_on (ball p e) f" + using contf continuous_on_subset by blast + moreover have "f holomorphic_on {z. dist p z < e \ d \ z < k}" + apply (rule holomorphic_on_subset [OF holf1]) + using e by auto + moreover have "f holomorphic_on {z. dist p z < e \ k < d \ z}" + apply (rule holomorphic_on_subset [OF holf2]) + using e by auto + ultimately show ?thesis + apply (rule_tac x="ball p e" in exI) + using \e > 0\ e \d \ 0\ + apply (simp add:, clarify) + apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k]) + apply (auto simp: subset_hull) + done + qed + show ?thesis + by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) +qed + +proposition Schwarz_reflection: + assumes "open S" and cnjs: "cnj ` S \ S" + and holf: "f holomorphic_on (S \ {z. 0 < Im z})" + and contf: "continuous_on (S \ {z. 0 \ Im z}) f" + and f: "\z. \z \ S; z \ \\ \ (f z) \ \" + shows "(\z. if 0 \ Im z then f z else cnj(f(cnj z))) holomorphic_on S" +proof - + have 1: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. 0 < Im z})" + by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) + have cont_cfc: "continuous_on (S \ {z. Im z \ 0}) (cnj o f o cnj)" + apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) + using cnjs apply auto + done + have "cnj \ f \ cnj field_differentiable at x within S \ {z. Im z < 0}" + if "x \ S" "Im x < 0" "f field_differentiable at (cnj x) within S \ {z. 0 < Im z}" for x + using that + apply (simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm, clarify) + apply (rule_tac x="cnj f'" in exI) + apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) + apply (drule_tac x="cnj xa" in bspec) + using cnjs apply force + apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) + done + then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \ {z. Im z < 0})" + using holf cnjs + by (force simp: holomorphic_on_def) + have 2: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. Im z < 0})" + apply (rule iffD1 [OF holomorphic_cong [OF refl]]) + using hol_cfc by auto + have [simp]: "(S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0}) = S" + by force + have "continuous_on ((S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0})) + (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" + apply (rule continuous_on_cases_local) + using cont_cfc contf + apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) + using f Reals_cnj_iff complex_is_Real_iff apply auto + done + then have 3: "continuous_on S (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" + by force + show ?thesis + apply (rule holomorphic_on_paste_across_line [OF \open S\, of "- \" _ 0]) + using 1 2 3 + apply auto + done +qed + +subsection\Bloch's theorem\ + +lemma Bloch_lemma_0: + assumes holf: "f holomorphic_on cball 0 r" and "0 < r" + and [simp]: "f 0 = 0" + and le: "\z. norm z < r \ norm(deriv f z) \ 2 * norm(deriv f 0)" + shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \ f ` ball 0 r" +proof - + have "sqrt 2 < 3/2" + by (rule real_less_lsqrt) (auto simp: power2_eq_square) + then have sq3: "0 < 3 - 2 * sqrt 2" by simp + show ?thesis + proof (cases "deriv f 0 = 0") + case True then show ?thesis by simp + next + case False + define C where "C = 2 * norm(deriv f 0)" + have "0 < C" using False by (simp add: C_def) + have holf': "f holomorphic_on ball 0 r" using holf + using ball_subset_cball holomorphic_on_subset by blast + then have holdf': "deriv f holomorphic_on ball 0 r" + by (rule holomorphic_deriv [OF _ open_ball]) + have "Le1": "norm(deriv f z - deriv f 0) \ norm z / (r - norm z) * C" + if "norm z < r" for z + proof - + have T1: "norm(deriv f z - deriv f 0) \ norm z / (R - norm z) * C" + if R: "norm z < R" "R < r" for R + proof - + have "0 < R" using R + by (metis less_trans norm_zero zero_less_norm_iff) + have df_le: "\x. norm x < r \ norm (deriv f x) \ C" + using le by (simp add: C_def) + have hol_df: "deriv f holomorphic_on cball 0 R" + apply (rule holomorphic_on_subset) using R holdf' by auto + have *: "((\w. deriv f w / (w - z)) has_contour_integral 2 * pi * \ * deriv f z) (circlepath 0 R)" + if "norm z < R" for z + using \0 < R\ that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] + by (force simp: winding_number_circlepath) + have **: "((\x. deriv f x / (x - z) - deriv f x / x) has_contour_integral + of_real (2 * pi) * \ * (deriv f z - deriv f 0)) + (circlepath 0 R)" + using has_contour_integral_diff [OF * [of z] * [of 0]] \0 < R\ that + by (simp add: algebra_simps) + have [simp]: "\x. norm x = R \ x \ z" using that(1) by blast + have "norm (deriv f x / (x - z) - deriv f x / x) + \ C * norm z / (R * (R - norm z))" + if "norm x = R" for x + proof - + have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = + norm (deriv f x) * norm z" + by (simp add: norm_mult right_diff_distrib') + show ?thesis + using \0 < R\ \0 < C\ R that + apply (simp add: norm_mult norm_divide divide_simps) + using df_le norm_triangle_ineq2 \0 < C\ apply (auto intro!: mult_mono) + done + qed + then show ?thesis + using has_contour_integral_bound_circlepath + [OF **, of "C * norm z/(R*(R - norm z))"] + \0 < R\ \0 < C\ R + apply (simp add: norm_mult norm_divide) + apply (simp add: divide_simps mult.commute) + done + qed + obtain r' where r': "norm z < r'" "r' < r" + using Rats_dense_in_real [of "norm z" r] \norm z < r\ by blast + then have [simp]: "closure {r'<.. norm(f z)" + if r: "norm z < r" for z + proof - + have 1: "\x. x \ ball 0 r \ + ((\z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) + (at x within ball 0 r)" + by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ + have 2: "closed_segment 0 z \ ball 0 r" + by (metis \0 < r\ convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) + have 3: "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" + apply (rule integrable_on_cmult_right [where 'b=real, simplified]) + apply (rule integrable_on_cdivide [where 'b=real, simplified]) + apply (rule integrable_on_cmult_left [where 'b=real, simplified]) + apply (rule ident_integrable_on) + done + have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm z * norm z * x * C / (r - norm z)" + if x: "0 \ x" "x \ 1" for x + proof - + have [simp]: "x * norm z < r" + using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) + have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \ norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" + apply (rule Le1) using r x \0 < r\ by simp + also have "... \ norm (x *\<^sub>R z) / (r - norm z) * C" + using r x \0 < r\ + apply (simp add: field_split_simps) + by (simp add: \0 < C\ mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) + finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm (x *\<^sub>R z) / (r - norm z) * C * norm z" + by (rule mult_right_mono) simp + with x show ?thesis by (simp add: algebra_simps) + qed + have le_norm: "abc \ norm d - e \ norm(f - d) \ e \ abc \ norm f" for abc d e and f::complex + by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) + have "norm (integral {0..1} (\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) + \ integral {0..1} (\t. (norm z)\<^sup>2 * t / (r - norm z) * C)" + apply (rule integral_norm_bound_integral) + using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 + apply (simp add: has_contour_integral_linepath has_integral_integrable_integral) + apply (rule 3) + apply (simp add: norm_mult power2_eq_square 4) + done + then have int_le: "norm (f z - deriv f 0 * z) \ (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" + using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 + apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) + done + show ?thesis + apply (rule le_norm [OF _ int_le]) + using \norm z < r\ + apply (simp add: power2_eq_square divide_simps C_def norm_mult) + proof - + have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" + by (simp add: algebra_simps) + then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" + by (simp add: algebra_simps) + qed + qed + have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" + by (auto simp: sqrt2_less_2) + have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" + apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) + apply (subst closure_ball) + using \0 < r\ mult_pos_pos sq201 + apply (auto simp: cball_subset_cball_iff) + done + have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" + apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force) + using \0 < r\ mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff) + using False \0 < r\ centre_in_ball holf' holomorphic_nonconstant by blast + have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = + ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" + by simp + also have "... \ f ` ball 0 ((1 - sqrt 2 / 2) * r)" + proof - + have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \ norm (f z)" + if "norm z = (1 - sqrt 2 / 2) * r" for z + apply (rule order_trans [OF _ *]) + using \0 < r\ + apply (simp_all add: field_simps power2_eq_square that) + apply (simp add: mult.assoc [symmetric]) + done + show ?thesis + apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) + using \0 < r\ sq201 3 apply simp_all + using C_def \0 < C\ sq3 apply force + done + qed + also have "... \ f ` ball 0 r" + apply (rule image_subsetI [OF imageI], simp) + apply (erule less_le_trans) + using \0 < r\ apply (auto simp: field_simps) + done + finally show ?thesis . + qed +qed + +lemma Bloch_lemma: + assumes holf: "f holomorphic_on cball a r" and "0 < r" + and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" + shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" +proof - + have fz: "(\z. f (a + z)) = f o (\z. (a + z))" + by (simp add: o_def) + have hol0: "(\z. f (a + z)) holomorphic_on cball 0 r" + unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ + then have [simp]: "\x. norm x < r \ (\z. f (a + z)) field_differentiable at x" + by (metis open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) + have [simp]: "\z. norm z < r \ f field_differentiable at (a + z)" + by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute) + then have [simp]: "f field_differentiable at a" + by (metis add.comm_neutral \0 < r\ norm_eq_zero) + have hol1: "(\z. f (a + z) - f a) holomorphic_on cball 0 r" + by (intro holomorphic_intros hol0) + then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\z. f (a + z) - f a) 0)) + \ (\z. f (a + z) - f a) ` ball 0 r" + apply (rule Bloch_lemma_0) + apply (simp_all add: \0 < r\) + apply (simp add: fz deriv_chain) + apply (simp add: dist_norm le) + done + then show ?thesis + apply clarify + apply (drule_tac c="x - f a" in subsetD) + apply (force simp: fz \0 < r\ dist_norm deriv_chain field_differentiable_compose)+ + done +qed + +proposition Bloch_unit: + assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" + obtains b r where "1/12 < r" and "ball b r \ f ` (ball a 1)" +proof - + define r :: real where "r = 249/256" + have "0 < r" "r < 1" by (auto simp: r_def) + define g where "g z = deriv f z * of_real(r - norm(z - a))" for z + have "deriv f holomorphic_on ball a 1" + by (rule holomorphic_deriv [OF holf open_ball]) + then have "continuous_on (ball a 1) (deriv f)" + using holomorphic_on_imp_continuous_on by blast + then have "continuous_on (cball a r) (deriv f)" + by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \r < 1\) + then have "continuous_on (cball a r) g" + by (simp add: g_def continuous_intros) + then have 1: "compact (g ` cball a r)" + by (rule compact_continuous_image [OF _ compact_cball]) + have 2: "g ` cball a r \ {}" + using \r > 0\ by auto + obtain p where pr: "p \ cball a r" + and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" + using distance_attains_sup [OF 1 2, of 0] by force + define t where "t = (r - norm(p - a)) / 2" + have "norm (p - a) \ r" + using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) + then have "norm (p - a) < r" using pr + by (simp add: norm_minus_commute dist_norm) + then have "0 < t" + by (simp add: t_def) + have cpt: "cball p t \ ball a r" + using \0 < t\ by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) + have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" + if "y \ cball a r" for y + proof - + have [simp]: "norm (y - a) \ r" + using that by (simp add: dist_norm norm_minus_commute) + have "norm (g y) \ norm (g p)" + using pge [OF that] by simp + then have "norm (deriv f y) * abs (r - norm (y - a)) \ norm (deriv f p) * abs (r - norm (p - a))" + by (simp only: dist_norm g_def norm_mult norm_of_real) + with that \norm (p - a) < r\ show ?thesis + by (simp add: dist_norm field_split_simps) + qed + have le_norm_dfp: "r / (r - norm (p - a)) \ norm (deriv f p)" + using gen_le_dfp [of a] \r > 0\ by auto + have 1: "f holomorphic_on cball p t" + apply (rule holomorphic_on_subset [OF holf]) + using cpt \r < 1\ order_subst1 subset_ball by auto + have 2: "norm (deriv f z) \ 2 * norm (deriv f p)" if "z \ ball p t" for z + proof - + have z: "z \ cball a r" + by (meson ball_subset_cball subsetD cpt that) + then have "norm(z - a) < r" + by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) + have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" + using gen_le_dfp [OF z] by simp + with \norm (z - a) < r\ \norm (p - a) < r\ + have "norm (deriv f z) \ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" + by (simp add: field_simps) + also have "... \ 2 * norm (deriv f p)" + apply (rule mult_right_mono) + using that \norm (p - a) < r\ \norm(z - a) < r\ + apply (simp_all add: field_simps t_def dist_norm [symmetric]) + using dist_triangle3 [of z a p] by linarith + finally show ?thesis . + qed + have sqrt2: "sqrt 2 < 2113/1494" + by (rule real_less_lsqrt) (auto simp: power2_eq_square) + then have sq3: "0 < 3 - 2 * sqrt 2" by simp + have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" + using sq3 sqrt2 by (auto simp: field_simps r_def) + also have "... \ cmod (deriv f p) * (r - cmod (p - a))" + using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) + finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" + using pos_divide_less_eq half_gt_zero_iff sq3 by blast + then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" + using sq3 by (simp add: mult.commute t_def) + have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball p t" + by (rule Bloch_lemma [OF 1 \0 < t\ 2]) + also have "... \ f ` ball a 1" + apply (rule image_mono) + apply (rule order_trans [OF ball_subset_cball]) + apply (rule order_trans [OF cpt]) + using \0 < t\ \r < 1\ apply (simp add: ball_subset_ball_iff dist_norm) + done + finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball a 1" . + with ** show ?thesis + by (rule that) +qed + +theorem Bloch: + assumes holf: "f holomorphic_on ball a r" and "0 < r" + and r': "r' \ r * norm (deriv f a) / 12" + obtains b where "ball b r' \ f ` (ball a r)" +proof (cases "deriv f a = 0") + case True with r' show ?thesis + using ball_eq_empty that by fastforce +next + case False + define C where "C = deriv f a" + have "0 < norm C" using False by (simp add: C_def) + have dfa: "f field_differentiable at a" + apply (rule holomorphic_on_imp_differentiable_at [OF holf]) + using \0 < r\ by auto + have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" + by (simp add: o_def) + have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" + apply (rule holomorphic_on_subset [OF holf]) + using \0 < r\ apply (force simp: dist_norm norm_mult) + done + have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" + apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ + using \0 < r\ by (simp add: C_def False) + have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative + (deriv f (a + of_real r * z) / C)) (at z)" + if "norm z < 1" for z + proof - + have *: "((\x. f (a + of_real r * x)) has_field_derivative + (deriv f (a + of_real r * z) * of_real r)) (at z)" + apply (simp add: fo) + apply (rule DERIV_chain [OF field_differentiable_derivI]) + apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) + using \0 < r\ apply (simp add: dist_norm norm_mult that) + apply (rule derivative_eq_intros | simp)+ + done + show ?thesis + apply (rule derivative_eq_intros * | simp)+ + using \0 < r\ by (auto simp: C_def False) + qed + have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" + apply (subst deriv_cdivide_right) + apply (simp add: field_differentiable_def fo) + apply (rule exI) + apply (rule DERIV_chain [OF field_differentiable_derivI]) + apply (simp add: dfa) + apply (rule derivative_eq_intros | simp add: C_def False fo)+ + using \0 < r\ + apply (simp add: C_def False fo) + apply (simp add: derivative_intros dfa deriv_chain) + done + have sb1: "(*) (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 + \ f ` ball a r" + using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) + have sb2: "ball (C * r * b) r' \ (*) (C * r) ` ball b t" + if "1 / 12 < t" for b t + proof - + have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" + using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right + by auto + show ?thesis + apply clarify + apply (rule_tac x="x / (C * r)" in image_eqI) + using \0 < r\ + apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) + apply (erule less_le_trans) + apply (rule order_trans [OF r' *]) + done + qed + show ?thesis + apply (rule Bloch_unit [OF 1 2]) + apply (rename_tac t) + apply (rule_tac b="(C * of_real r) * b" in that) + apply (drule image_mono [where f = "\z. (C * of_real r) * z"]) + using sb1 sb2 + apply force + done +qed + +corollary Bloch_general: + assumes holf: "f holomorphic_on s" and "a \ s" + and tle: "\z. z \ frontier s \ t \ dist a z" + and rle: "r \ t * norm(deriv f a) / 12" + obtains b where "ball b r \ f ` s" +proof - + consider "r \ 0" | "0 < t * norm(deriv f a) / 12" using rle by force + then show ?thesis + proof cases + case 1 then show ?thesis + by (simp add: ball_empty that) + next + case 2 + show ?thesis + proof (cases "deriv f a = 0") + case True then show ?thesis + using rle by (simp add: ball_empty that) + next + case False + then have "t > 0" + using 2 by (force simp: zero_less_mult_iff) + have "\ ball a t \ s \ ball a t \ frontier s \ {}" + apply (rule connected_Int_frontier [of "ball a t" s], simp_all) + using \0 < t\ \a \ s\ centre_in_ball apply blast + done + with tle have *: "ball a t \ s" by fastforce + then have 1: "f holomorphic_on ball a t" + using holf using holomorphic_on_subset by blast + show ?thesis + apply (rule Bloch [OF 1 \t > 0\ rle]) + apply (rule_tac b=b in that) + using * apply force + done + qed + qed +qed + +subsection \Cauchy's residue theorem\ + +text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. + Interactive Theorem Proving\ + +definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where + "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" + +lemma Eps_cong: + assumes "\x. P x = Q x" + shows "Eps P = Eps Q" + using ext[of P Q, OF assms] by simp + +lemma residue_cong: + assumes eq: "eventually (\z. f z = g z) (at z)" and "z = z'" + shows "residue f z = residue g z'" +proof - + from assms have eq': "eventually (\z. g z = f z) (at z)" + by (simp add: eq_commute) + let ?P = "\f c e. (\\>0. \ < e \ + (f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \))" + have "residue f z = residue g z" unfolding residue_def + proof (rule Eps_cong) + fix c :: complex + have "\e>0. ?P g c e" + if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g + proof - + from that(1) obtain e where e: "e > 0" "?P f c e" + by blast + from that(2) obtain e' where e': "e' > 0" "\z'. z' \ z \ dist z' z < e' \ f z' = g z'" + unfolding eventually_at by blast + have "?P g c (min e e')" + proof (intro allI exI impI, goal_cases) + case (1 \) + hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" + using e(2) by auto + thus ?case + proof (rule has_contour_integral_eq) + fix z' assume "z' \ path_image (circlepath z \)" + hence "dist z' z < e'" and "z' \ z" + using 1 by (auto simp: dist_commute) + with e'(2)[of z'] show "f z' = g z'" by simp + qed + qed + moreover from e and e' have "min e e' > 0" by auto + ultimately show ?thesis by blast + qed + from this[OF _ eq] and this[OF _ eq'] + show "(\e>0. ?P f c e) \ (\e>0. ?P g c e)" + by blast + qed + with assms show ?thesis by simp +qed + +lemma contour_integral_circlepath_eq: + assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" + and e2_cball:"cball z e2 \ s" + shows + "f contour_integrable_on circlepath z e1" + "f contour_integrable_on circlepath z e2" + "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" +proof - + define l where "l \ linepath (z+e2) (z+e1)" + have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto + have "e2>0" using \e1>0\ \e1\e2\ by auto + have zl_img:"z\path_image l" + proof + assume "z \ path_image l" + then have "e2 \ cmod (e2 - e1)" + using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def + by (auto simp add:closed_segment_commute) + thus False using \e2>0\ \e1>0\ \e1\e2\ + apply (subst (asm) norm_of_real) + by auto + qed + define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" + show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" + proof - + show "f contour_integrable_on circlepath z e2" + apply (intro contour_integrable_continuous_circlepath[OF + continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) + using \e2>0\ e2_cball by auto + show "f contour_integrable_on (circlepath z e1)" + apply (intro contour_integrable_continuous_circlepath[OF + continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) + using \e1>0\ \e1\e2\ e2_cball by auto + qed + have [simp]:"f contour_integrable_on l" + proof - + have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ + by (intro closed_segment_subset,auto simp add:dist_norm) + hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def + by auto + then show "f contour_integrable_on l" unfolding l_def + apply (intro contour_integrable_continuous_linepath[OF + continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) + by auto + qed + let ?ig="\g. contour_integral g f" + have "(f has_contour_integral 0) g" + proof (rule Cauchy_theorem_global[OF _ f_holo]) + show "open (s - {z})" using \open s\ by auto + show "valid_path g" unfolding g_def l_def by auto + show "pathfinish g = pathstart g" unfolding g_def l_def by auto + next + have path_img:"path_image g \ cball z e2" + proof - + have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ + by (intro closed_segment_subset,auto simp add:dist_norm) + moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto + ultimately show ?thesis unfolding g_def l_def using \e2>0\ + by (simp add: path_image_join closed_segment_commute) + qed + show "path_image g \ s - {z}" + proof - + have "z\path_image g" using zl_img + unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) + moreover note \cball z e2 \ s\ and path_img + ultimately show ?thesis by auto + qed + show "winding_number g w = 0" when"w \ s - {z}" for w + proof - + have "winding_number g w = 0" when "w\s" using that e2_cball + apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) + by (auto simp add:g_def l_def) + moreover have "winding_number g z=0" + proof - + let ?Wz="\g. winding_number g z" + have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + + ?Wz (reversepath l)" + using \e2>0\ \e1>0\ zl_img unfolding g_def l_def + by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ + also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" + using zl_img + apply (subst (2) winding_number_reversepath) + by (auto simp add:l_def closed_segment_commute) + also have "... = 0" + proof - + have "?Wz (circlepath z e2) = 1" using \e2>0\ + by (auto intro: winding_number_circlepath_centre) + moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \e1>0\ + apply (subst winding_number_reversepath) + by (auto intro: winding_number_circlepath_centre) + ultimately show ?thesis by auto + qed + finally show ?thesis . + qed + ultimately show ?thesis using that by auto + qed + qed + then have "0 = ?ig g" using contour_integral_unique by simp + also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + + ?ig (reversepath l)" + unfolding g_def + by (auto simp add:contour_integrable_reversepath_eq) + also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" + by (auto simp add:contour_integral_reversepath) + finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" + by simp +qed + +lemma base_residue: + assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" + and r_cball:"cball z r \ s" + shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" +proof - + obtain e where "e>0" and e_cball:"cball z e \ s" + using open_contains_cball[of s] \open s\ \z\s\ by auto + define c where "c \ 2 * pi * \" + define i where "i \ contour_integral (circlepath z e) f / c" + have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ + proof - + have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" + "f contour_integrable_on circlepath z \" + "f contour_integrable_on circlepath z e" + using \\ + by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ + then show ?thesis unfolding i_def c_def + by (auto intro:has_contour_integral_integral) + qed + then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" + unfolding residue_def c_def + apply (rule_tac someI[of _ i],intro exI[where x=e]) + by (auto simp add:\e>0\ c_def) + then obtain e' where "e'>0" + and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" + by auto + let ?int="\e. contour_integral (circlepath z e) f" + define \ where "\ \ Min {r,e'} / 2" + have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto + have "(f has_contour_integral c * (residue f z)) (circlepath z \)" + using e'_def[rule_format,OF \\>0\ \\] . + then show ?thesis unfolding c_def + using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] + by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) +qed + +lemma residue_holo: + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s" + shows "residue f z = 0" +proof - + define c where "c \ 2 * pi * \" + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + using open_contains_cball_eq by blast + have "(f has_contour_integral c*residue f z) (circlepath z e)" + using f_holo + by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) + moreover have "(f has_contour_integral 0) (circlepath z e)" + using f_holo e_cball \e>0\ + by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) + ultimately have "c*residue f z =0" + using has_contour_integral_unique by blast + thus ?thesis unfolding c_def by auto +qed + +lemma residue_const:"residue (\_. c) z = 0" + by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) + +lemma residue_add: + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + and g_holo:"g holomorphic_on s - {z}" + shows "residue (\z. f z + g z) z= residue f z + residue g z" +proof - + define c where "c \ 2 * pi * \" + define fg where "fg \ (\z. f z+g z)" + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + using open_contains_cball_eq by blast + have "(fg has_contour_integral c * residue fg z) (circlepath z e)" + unfolding fg_def using f_holo g_holo + apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) + by (auto intro:holomorphic_intros) + moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" + unfolding fg_def using f_holo g_holo + by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) + ultimately have "c*(residue f z + residue g z) = c * residue fg z" + using has_contour_integral_unique by (auto simp add:distrib_left) + thus ?thesis unfolding fg_def + by (auto simp add:c_def) +qed + +lemma residue_lmul: + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\z. c * (f z)) z= c * residue f z" +proof (cases "c=0") + case True + thus ?thesis using residue_const by auto +next + case False + define c' where "c' \ 2 * pi * \" + define f' where "f' \ (\z. c * (f z))" + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + using open_contains_cball_eq by blast + have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" + unfolding f'_def using f_holo + apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) + by (auto intro:holomorphic_intros) + moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" + unfolding f'_def using f_holo + by (auto intro: has_contour_integral_lmul + base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) + ultimately have "c' * residue f' z = c * (c' * residue f z)" + using has_contour_integral_unique by auto + thus ?thesis unfolding f'_def c'_def using False + by (auto simp add:field_simps) +qed + +lemma residue_rmul: + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\z. (f z) * c) z= residue f z * c" +using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) + +lemma residue_div: + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\z. (f z) / c) z= residue f z / c " +using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) + +lemma residue_neg: + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + shows "residue (\z. - (f z)) z= - residue f z" +using residue_lmul[OF assms,of "-1"] by auto + +lemma residue_diff: + assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" + and g_holo:"g holomorphic_on s - {z}" + shows "residue (\z. f z - g z) z= residue f z - residue g z" +using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] +by (auto intro:holomorphic_intros g_holo) + +lemma residue_simple: + assumes "open s" "z\s" and f_holo:"f holomorphic_on s" + shows "residue (\w. f w / (w - z)) z = f z" +proof - + define c where "c \ 2 * pi * \" + define f' where "f' \ \w. f w / (w - z)" + obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ + using open_contains_cball_eq by blast + have "(f' has_contour_integral c * f z) (circlepath z e)" + unfolding f'_def c_def using \e>0\ f_holo e_cball + by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) + moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" + unfolding f'_def using f_holo + apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) + by (auto intro!:holomorphic_intros) + ultimately have "c * f z = c * residue f' z" + using has_contour_integral_unique by blast + thus ?thesis unfolding c_def f'_def by auto +qed + +lemma residue_simple': + assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" + and lim: "((\w. f w * (w - z)) \ c) (at z)" + shows "residue f z = c" +proof - + define g where "g = (\w. if w = z then c else f w * (w - z))" + from holo have "(\w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P") + by (force intro: holomorphic_intros) + also have "?P \ g holomorphic_on (s - {z})" + by (intro holomorphic_cong refl) (simp_all add: g_def) + finally have *: "g holomorphic_on (s - {z})" . + + note lim + also have "(\w. f w * (w - z)) \z\ c \ g \z\ g z" + by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter) + finally have **: "g \z\ g z" . + + have g_holo: "g holomorphic_on s" + by (rule no_isolated_singularity'[where K = "{z}"]) + (insert assms * **, simp_all add: at_within_open_NO_MATCH) + from s and this have "residue (\w. g w / (w - z)) z = g z" + by (rule residue_simple) + also have "\\<^sub>F za in at z. g za / (za - z) = f za" + unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def) + hence "residue (\w. g w / (w - z)) z = residue f z" + by (intro residue_cong refl) + finally show ?thesis + by (simp add: g_def) +qed + +lemma residue_holomorphic_over_power: + assumes "open A" "z0 \ A" "f holomorphic_on A" + shows "residue (\z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n" +proof - + let ?f = "\z. f z / (z - z0) ^ Suc n" + from assms(1,2) obtain r where r: "r > 0" "cball z0 r \ A" + by (auto simp: open_contains_cball) + have "(?f has_contour_integral 2 * pi * \ * residue ?f z0) (circlepath z0 r)" + using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros) + moreover have "(?f has_contour_integral 2 * pi * \ / fact n * (deriv ^^ n) f z0) (circlepath z0 r)" + using assms r + by (intro Cauchy_has_contour_integral_higher_derivative_circlepath) + (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on) + ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" + by (rule has_contour_integral_unique) + thus ?thesis by (simp add: field_simps) +qed + +lemma residue_holomorphic_over_power': + assumes "open A" "0 \ A" "f holomorphic_on A" + shows "residue (\z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" + using residue_holomorphic_over_power[OF assms] by simp + +theorem residue_fps_expansion_over_power_at_0: + assumes "f has_fps_expansion F" + shows "residue (\z. f z / z ^ Suc n) 0 = fps_nth F n" +proof - + from has_fps_expansion_imp_holomorphic[OF assms] guess s . note s = this + have "residue (\z. f z / (z - 0) ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" + using assms s unfolding has_fps_expansion_def + by (intro residue_holomorphic_over_power[of s]) (auto simp: zero_ereal_def) + also from assms have "\ = fps_nth F n" + by (subst fps_nth_fps_expansion) auto + finally show ?thesis by simp +qed + +lemma get_integrable_path: + assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\s-pts" "b\s-pts" + obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" + "path_image g \ s-pts" "f contour_integrable_on g" using assms +proof (induct arbitrary:s thesis a rule:finite_induct[OF \finite pts\]) + case 1 + obtain g where "valid_path g" "path_image g \ s" "pathstart g = a" "pathfinish g = b" + using connected_open_polynomial_connected[OF \open s\,of a b ] \connected (s - {})\ + valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto + moreover have "f contour_integrable_on g" + using contour_integrable_holomorphic_simple[OF _ \open s\ \valid_path g\ \path_image g \ s\,of f] + \f holomorphic_on s - {}\ + by auto + ultimately show ?case using "1"(1)[of g] by auto +next + case idt:(2 p pts) + obtain e where "e>0" and e:"\w\ball a e. w \ s \ (w \ a \ w \ insert p pts)" + using finite_ball_avoid[OF \open s\ \finite (insert p pts)\, of a] + \a \ s - insert p pts\ + by auto + define a' where "a' \ a+e/2" + have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ + by (auto simp add:dist_complex_def a'_def) + then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" + "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" + using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) + by (metis Diff_insert2 open_delete) + define g where "g \ linepath a a' +++ g'" + have "valid_path g" unfolding g_def by (auto intro: valid_path_join) + moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto + moreover have "path_image g \ s - insert p pts" unfolding g_def + proof (rule subset_path_image_join) + have "closed_segment a a' \ ball a e" using \e>0\ + by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) + then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) + by auto + next + show "path_image g' \ s - insert p pts" using g'(4) by blast + qed + moreover have "f contour_integrable_on g" + proof - + have "closed_segment a a' \ ball a e" using \e>0\ + by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) + then have "continuous_on (closed_segment a a') f" + using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] + apply (elim continuous_on_subset) + by auto + then have "f contour_integrable_on linepath a a'" + using contour_integrable_continuous_linepath by auto + then show ?thesis unfolding g_def + apply (rule contour_integrable_joinI) + by (auto simp add: \e>0\) + qed + ultimately show ?case using idt.prems(1)[of g] by auto +qed + +lemma Cauchy_theorem_aux: + assumes "open s" "connected (s-pts)" "finite pts" "pts \ s" "f holomorphic_on s-pts" + "valid_path g" "pathfinish g = pathstart g" "path_image g \ s-pts" + "\z. (z \ s) \ winding_number g z = 0" + "\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" + shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" + using assms +proof (induct arbitrary:s g rule:finite_induct[OF \finite pts\]) + case 1 + then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) +next + case (2 p pts) + note fin[simp] = \finite (insert p pts)\ + and connected = \connected (s - insert p pts)\ + and valid[simp] = \valid_path g\ + and g_loop[simp] = \pathfinish g = pathstart g\ + and holo[simp]= \f holomorphic_on s - insert p pts\ + and path_img = \path_image g \ s - insert p pts\ + and winding = \\z. z \ s \ winding_number g z = 0\ + and h = \\pa\s. 0 < h pa \ (\w\cball pa (h pa). w \ s \ (w \ pa \ w \ insert p pts))\ + have "h p>0" and "p\s" + and h_p: "\w\cball p (h p). w \ s \ (w \ p \ w \ insert p pts)" + using h \insert p pts \ s\ by auto + obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" + "path_image pg \ s-insert p pts" "f contour_integrable_on pg" + proof - + have "p + h p\cball p (h p)" using h[rule_format,of p] + by (simp add: \p \ s\ dist_norm) + then have "p + h p \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ + by fastforce + moreover have "pathstart g \ s - insert p pts " using path_img by auto + ultimately show ?thesis + using get_integrable_path[OF \open s\ connected fin holo,of "pathstart g" "p+h p"] that + by blast + qed + obtain n::int where "n=winding_number g p" + using integer_winding_number[OF _ g_loop,of p] valid path_img + by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) + define p_circ where "p_circ \ circlepath p (h p)" + define p_circ_pt where "p_circ_pt \ linepath (p+h p) (p+h p)" + define n_circ where "n_circ \ \n. ((+++) p_circ ^^ n) p_circ_pt" + define cp where "cp \ if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" + have n_circ:"valid_path (n_circ k)" + "winding_number (n_circ k) p = k" + "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" + "path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" + "p \ path_image (n_circ k)" + "\p'. p'\s - pts \ winding_number (n_circ k) p'=0 \ p'\path_image (n_circ k)" + "f contour_integrable_on (n_circ k)" + "contour_integral (n_circ k) f = k * contour_integral p_circ f" + for k + proof (induct k) + case 0 + show "valid_path (n_circ 0)" + and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" + and "winding_number (n_circ 0) p = of_nat 0" + and "pathstart (n_circ 0) = p + h p" + and "pathfinish (n_circ 0) = p + h p" + and "p \ path_image (n_circ 0)" + unfolding n_circ_def p_circ_pt_def using \h p > 0\ + by (auto simp add: dist_norm) + show "winding_number (n_circ 0) p'=0 \ p'\path_image (n_circ 0)" when "p'\s- pts" for p' + unfolding n_circ_def p_circ_pt_def + apply (auto intro!:winding_number_trivial) + by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ + show "f contour_integrable_on (n_circ 0)" + unfolding n_circ_def p_circ_pt_def + by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing) + show "contour_integral (n_circ 0) f = of_nat 0 * contour_integral p_circ f" + unfolding n_circ_def p_circ_pt_def by auto + next + case (Suc k) + have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto + have pcirc:"p \ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" + using Suc(3) unfolding p_circ_def using \h p > 0\ by (auto simp add: p_circ_def) + have pcirc_image:"path_image p_circ \ s - insert p pts" + proof - + have "path_image p_circ \ cball p (h p)" using \0 < h p\ p_circ_def by auto + then show ?thesis using h_p pcirc(1) by auto + qed + have pcirc_integrable:"f contour_integrable_on p_circ" + by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] + contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on + holomorphic_on_subset[OF holo]) + show "valid_path (n_circ (Suc k))" + using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto + show "path_image (n_circ (Suc k)) + = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" + proof - + have "path_image p_circ = sphere p (h p)" + unfolding p_circ_def using \0 < h p\ by auto + then show ?thesis unfolding n_Suc using Suc.hyps(5) \h p>0\ + by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) + qed + then show "p \ path_image (n_circ (Suc k))" using \h p>0\ by auto + show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" + proof - + have "winding_number p_circ p = 1" + by (simp add: \h p > 0\ p_circ_def winding_number_circlepath_centre) + moreover have "p \ path_image (n_circ k)" using Suc(5) \h p>0\ by auto + then have "winding_number (p_circ +++ n_circ k) p + = winding_number p_circ p + winding_number (n_circ k) p" + using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc + apply (intro winding_number_join) + by auto + ultimately show ?thesis using Suc(2) unfolding n_circ_def + by auto + qed + show "pathstart (n_circ (Suc k)) = p + h p" + by (simp add: n_circ_def p_circ_def) + show "pathfinish (n_circ (Suc k)) = p + h p" + using Suc(4) unfolding n_circ_def by auto + show "winding_number (n_circ (Suc k)) p'=0 \ p'\path_image (n_circ (Suc k))" when "p'\s-pts" for p' + proof - + have " p' \ path_image p_circ" using \p \ s\ h p_circ_def that using pcirc_image by blast + moreover have "p' \ path_image (n_circ k)" + using Suc.hyps(7) that by blast + moreover have "winding_number p_circ p' = 0" + proof - + have "path_image p_circ \ cball p (h p)" + using h unfolding p_circ_def using \p \ s\ by fastforce + moreover have "p'\cball p (h p)" using \p \ s\ h that "2.hyps"(2) by fastforce + ultimately show ?thesis unfolding p_circ_def + apply (intro winding_number_zero_outside) + by auto + qed + ultimately show ?thesis + unfolding n_Suc + apply (subst winding_number_join) + by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that]) + qed + show "f contour_integrable_on (n_circ (Suc k))" + unfolding n_Suc + by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) + show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" + unfolding n_Suc + by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] + Suc(9) algebra_simps) + qed + have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" + "valid_path cp" "path_image cp \ s - insert p pts" + "winding_number cp p = - n" + "\p'. p'\s - pts \ winding_number cp p'=0 \ p' \ path_image cp" + "f contour_integrable_on cp" + "contour_integral cp f = - n * contour_integral p_circ f" + proof - + show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" + using n_circ unfolding cp_def by auto + next + have "sphere p (h p) \ s - insert p pts" + using h[rule_format,of p] \insert p pts \ s\ by force + moreover have "p + complex_of_real (h p) \ s - insert p pts" + using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) + ultimately show "path_image cp \ s - insert p pts" unfolding cp_def + using n_circ(5) by auto + next + show "winding_number cp p = - n" + unfolding cp_def using winding_number_reversepath n_circ \h p>0\ + by (auto simp: valid_path_imp_path) + next + show "winding_number cp p'=0 \ p' \ path_image cp" when "p'\s - pts" for p' + unfolding cp_def + apply (auto) + apply (subst winding_number_reversepath) + by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1)) + next + show "f contour_integrable_on cp" unfolding cp_def + using contour_integrable_reversepath_eq n_circ(1,8) by auto + next + show "contour_integral cp f = - n * contour_integral p_circ f" + unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) + by auto + qed + define g' where "g' \ g +++ pg +++ cp +++ (reversepath pg)" + have "contour_integral g' f = (\p\pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" + proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) + show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) + show "open (s - {p})" using \open s\ by auto + show " pts \ s - {p}" using \insert p pts \ s\ \ p \ pts\ by blast + show "f holomorphic_on s - {p} - pts" using holo \p \ pts\ by (metis Diff_insert2) + show "valid_path g'" + unfolding g'_def cp_def using n_circ valid pg g_loop + by (auto intro!:valid_path_join ) + show "pathfinish g' = pathstart g'" + unfolding g'_def cp_def using pg(2) by simp + show "path_image g' \ s - {p} - pts" + proof - + define s' where "s' \ s - {p} - pts" + have s':"s' = s-insert p pts " unfolding s'_def by auto + then show ?thesis using path_img pg(4) cp(4) + unfolding g'_def + apply (fold s'_def s') + apply (intro subset_path_image_join) + by auto + qed + note path_join_imp[simp] + show "\z. z \ s - {p} \ winding_number g' z = 0" + proof clarify + fix z assume z:"z\s - {p}" + have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z + + winding_number (pg +++ cp +++ (reversepath pg)) z" + proof (rule winding_number_join) + show "path g" using \valid_path g\ by (simp add: valid_path_imp_path) + show "z \ path_image g" using z path_img by auto + show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp + by (simp add: valid_path_imp_path) + next + have "path_image (pg +++ cp +++ reversepath pg) \ s - insert p pts" + using pg(4) cp(4) by (auto simp:subset_path_image_join) + then show "z \ path_image (pg +++ cp +++ reversepath pg)" using z by auto + next + show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto + qed + also have "... = winding_number g z + (winding_number pg z + + winding_number (cp +++ (reversepath pg)) z)" + proof (subst add_left_cancel,rule winding_number_join) + show "path pg" and "path (cp +++ reversepath pg)" + and "pathfinish pg = pathstart (cp +++ reversepath pg)" + by (auto simp add: valid_path_imp_path) + show "z \ path_image pg" using pg(4) z by blast + show "z \ path_image (cp +++ reversepath pg)" using z + by (metis Diff_iff \z \ path_image pg\ contra_subsetD cp(4) insertI1 + not_in_path_image_join path_image_reversepath singletonD) + qed + also have "... = winding_number g z + (winding_number pg z + + (winding_number cp z + winding_number (reversepath pg) z))" + apply (auto intro!:winding_number_join simp: valid_path_imp_path) + apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) + by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) + also have "... = winding_number g z + winding_number cp z" + apply (subst winding_number_reversepath) + apply (auto simp: valid_path_imp_path) + by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) + finally have "winding_number g' z = winding_number g z + winding_number cp z" + unfolding g'_def . + moreover have "winding_number g z + winding_number cp z = 0" + using winding z \n=winding_number g p\ by auto + ultimately show "winding_number g' z = 0" unfolding g'_def by auto + qed + show "\pa\s - {p}. 0 < h pa \ (\w\cball pa (h pa). w \ s - {p} \ (w \ pa \ w \ pts))" + using h by fastforce + qed + moreover have "contour_integral g' f = contour_integral g f + - winding_number g p * contour_integral p_circ f" + proof - + have "contour_integral g' f = contour_integral g f + + contour_integral (pg +++ cp +++ reversepath pg) f" + unfolding g'_def + apply (subst contour_integral_join) + by (auto simp add:open_Diff[OF \open s\,OF finite_imp_closed[OF fin]] + intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] + contour_integrable_reversepath) + also have "... = contour_integral g f + contour_integral pg f + + contour_integral (cp +++ reversepath pg) f" + apply (subst contour_integral_join) + by (auto simp add:contour_integrable_reversepath) + also have "... = contour_integral g f + contour_integral pg f + + contour_integral cp f + contour_integral (reversepath pg) f" + apply (subst contour_integral_join) + by (auto simp add:contour_integrable_reversepath) + also have "... = contour_integral g f + contour_integral cp f" + using contour_integral_reversepath + by (auto simp add:contour_integrable_reversepath) + also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f" + using \n=winding_number g p\ by auto + finally show ?thesis . + qed + moreover have "winding_number g' p' = winding_number g p'" when "p'\pts" for p' + proof - + have [simp]: "p' \ path_image g" "p' \ path_image pg" "p'\path_image cp" + using "2.prems"(8) that + apply blast + apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) + by (meson DiffD2 cp(4) rev_subsetD subset_insertI that) + have "winding_number g' p' = winding_number g p' + + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def + apply (subst winding_number_join) + apply (simp_all add: valid_path_imp_path) + apply (intro not_in_path_image_join) + by auto + also have "... = winding_number g p' + winding_number pg p' + + winding_number (cp +++ reversepath pg) p'" + apply (subst winding_number_join) + apply (simp_all add: valid_path_imp_path) + apply (intro not_in_path_image_join) + by auto + also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p' + + winding_number (reversepath pg) p'" + apply (subst winding_number_join) + by (simp_all add: valid_path_imp_path) + also have "... = winding_number g p' + winding_number cp p'" + apply (subst winding_number_reversepath) + by (simp_all add: valid_path_imp_path) + also have "... = winding_number g p'" using that by auto + finally show ?thesis . + qed + ultimately show ?case unfolding p_circ_def + apply (subst (asm) sum.cong[OF refl, + of pts _ "\p. winding_number g p * contour_integral (circlepath p (h p)) f"]) + by (auto simp add:sum.insert[OF \finite pts\ \p\pts\] algebra_simps) +qed + +lemma Cauchy_theorem_singularities: + assumes "open s" "connected s" "finite pts" and + holo:"f holomorphic_on s-pts" and + "valid_path g" and + loop:"pathfinish g = pathstart g" and + "path_image g \ s-pts" and + homo:"\z. (z \ s) \ winding_number g z = 0" and + avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" + shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" + (is "?L=?R") +proof - + define circ where "circ \ \p. winding_number g p * contour_integral (circlepath p (h p)) f" + define pts1 where "pts1 \ pts \ s" + define pts2 where "pts2 \ pts - pts1" + have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" + unfolding pts1_def pts2_def by auto + have "contour_integral g f = (\p\pts1. circ p)" unfolding circ_def + proof (rule Cauchy_theorem_aux[OF \open s\ _ _ \pts1\s\ _ \valid_path g\ loop _ homo]) + have "finite pts1" unfolding pts1_def using \finite pts\ by auto + then show "connected (s - pts1)" + using \open s\ \connected s\ connected_open_delete_finite[of s] by auto + next + show "finite pts1" using \pts = pts1 \ pts2\ assms(3) by auto + show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) + show "path_image g \ s - pts1" using assms(7) pts1_def by auto + show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ pts1))" + by (simp add: avoid pts1_def) + qed + moreover have "sum circ pts2=0" + proof - + have "winding_number g p=0" when "p\pts2" for p + using \pts2 \ s={}\ that homo[rule_format,of p] by auto + thus ?thesis unfolding circ_def + apply (intro sum.neutral) + by auto + qed + moreover have "?R=sum circ pts1 + sum circ pts2" + unfolding circ_def + using sum.union_disjoint[OF _ _ \pts1 \ pts2 = {}\] \finite pts\ \pts=pts1 \ pts2\ + by blast + ultimately show ?thesis + apply (fold circ_def) + by auto +qed + +theorem Residue_theorem: + fixes s pts::"complex set" and f::"complex \ complex" + and g::"real \ complex" + assumes "open s" "connected s" "finite pts" and + holo:"f holomorphic_on s-pts" and + "valid_path g" and + loop:"pathfinish g = pathstart g" and + "path_image g \ s-pts" and + homo:"\z. (z \ s) \ winding_number g z = 0" + shows "contour_integral g f = 2 * pi * \ *(\p\pts. winding_number g p * residue f p)" +proof - + define c where "c \ 2 * pi * \" + obtain h where avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" + using finite_cball_avoid[OF \open s\ \finite pts\] by metis + have "contour_integral g f + = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" + using Cauchy_theorem_singularities[OF assms avoid] . + also have "... = (\p\pts. c * winding_number g p * residue f p)" + proof (intro sum.cong) + show "pts = pts" by simp + next + fix x assume "x \ pts" + show "winding_number g x * contour_integral (circlepath x (h x)) f + = c * winding_number g x * residue f x" + proof (cases "x\s") + case False + then have "winding_number g x=0" using homo by auto + thus ?thesis by auto + next + case True + have "contour_integral (circlepath x (h x)) f = c* residue f x" + using \x\pts\ \finite pts\ avoid[rule_format,OF True] + apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) + by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \open s\ finite_imp_closed]) + then show ?thesis by auto + qed + qed + also have "... = c * (\p\pts. winding_number g p * residue f p)" + by (simp add: sum_distrib_left algebra_simps) + finally show ?thesis unfolding c_def . +qed + +subsection \Non-essential singular points\ + +definition\<^marker>\tag important\ is_pole :: + "('a::topological_space \ 'b::real_normed_vector) \ 'a \ bool" where + "is_pole f a = (LIM x (at a). f x :> at_infinity)" + +lemma is_pole_cong: + assumes "eventually (\x. f x = g x) (at a)" "a=b" + shows "is_pole f a \ is_pole g b" + unfolding is_pole_def using assms by (intro filterlim_cong,auto) + +lemma is_pole_transform: + assumes "is_pole f a" "eventually (\x. f x = g x) (at a)" "a=b" + shows "is_pole g b" + using is_pole_cong assms by auto + +lemma is_pole_tendsto: + fixes f::"('a::topological_space \ 'b::real_normed_div_algebra)" + shows "is_pole f x \ ((inverse o f) \ 0) (at x)" +unfolding is_pole_def +by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) + +lemma is_pole_inverse_holomorphic: + assumes "open s" + and f_holo:"f holomorphic_on (s-{z})" + and pole:"is_pole f z" + and non_z:"\x\s-{z}. f x\0" + shows "(\x. if x=z then 0 else inverse (f x)) holomorphic_on s" +proof - + define g where "g \ \x. if x=z then 0 else inverse (f x)" + have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] + apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \ f"]) + by (simp_all add:g_def) + moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto + hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def + by (auto elim!:continuous_on_inverse simp add:non_z) + hence "continuous_on (s-{z}) g" unfolding g_def + apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) + by auto + ultimately have "continuous_on s g" using open_delete[OF \open s\] \open s\ + by (auto simp add:continuous_on_eq_continuous_at) + moreover have "(inverse o f) holomorphic_on (s-{z})" + unfolding comp_def using f_holo + by (auto elim!:holomorphic_on_inverse simp add:non_z) + hence "g holomorphic_on (s-{z})" + apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) + by (auto simp add:g_def) + ultimately show ?thesis unfolding g_def using \open s\ + by (auto elim!: no_isolated_singularity) +qed + +lemma not_is_pole_holomorphic: + assumes "open A" "x \ A" "f holomorphic_on A" + shows "\is_pole f x" +proof - + have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact + with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at) + hence "f \x\ f x" by (simp add: isCont_def) + thus "\is_pole f x" unfolding is_pole_def + using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto +qed + +lemma is_pole_inverse_power: "n > 0 \ is_pole (\z::complex. 1 / (z - a) ^ n) a" + unfolding is_pole_def inverse_eq_divide [symmetric] + by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros) + (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros) + +lemma is_pole_inverse: "is_pole (\z::complex. 1 / (z - a)) a" + using is_pole_inverse_power[of 1 a] by simp + +lemma is_pole_divide: + fixes f :: "'a :: t2_space \ 'b :: real_normed_field" + assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \ 0" + shows "is_pole (\z. f z / g z) z" +proof - + have "filterlim (\z. f z * inverse (g z)) at_infinity (at z)" + by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"] + filterlim_compose[OF filterlim_inverse_at_infinity])+ + (insert assms, auto simp: isCont_def) + thus ?thesis by (simp add: field_split_simps is_pole_def) +qed + +lemma is_pole_basic: + assumes "f holomorphic_on A" "open A" "z \ A" "f z \ 0" "n > 0" + shows "is_pole (\w. f w / (w - z) ^ n) z" +proof (rule is_pole_divide) + have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact + with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at) + have "filterlim (\w. (w - z) ^ n) (nhds 0) (at z)" + using assms by (auto intro!: tendsto_eq_intros) + thus "filterlim (\w. (w - z) ^ n) (at 0) (at z)" + by (intro filterlim_atI tendsto_eq_intros) + (insert assms, auto simp: eventually_at_filter) +qed fact+ + +lemma is_pole_basic': + assumes "f holomorphic_on A" "open A" "0 \ A" "f 0 \ 0" "n > 0" + shows "is_pole (\w. f w / w ^ n) 0" + using is_pole_basic[of f A 0] assms by simp + +text \The proposition + \<^term>\\x. ((f::complex\complex) \ x) (at z) \ is_pole f z\ +can be interpreted as the complex function \<^term>\f\ has a non-essential singularity at \<^term>\z\ +(i.e. the singularity is either removable or a pole).\ +definition not_essential::"[complex \ complex, complex] \ bool" where + "not_essential f z = (\x. f\z\x \ is_pole f z)" + +definition isolated_singularity_at::"[complex \ complex, complex] \ bool" where + "isolated_singularity_at f z = (\r>0. f analytic_on ball z r-{z})" + +named_theorems singularity_intros "introduction rules for singularities" + +lemma holomorphic_factor_unique: + fixes f::"complex \ complex" and z::complex and r::real and m n::int + assumes "r>0" "g z\0" "h z\0" + and asm:"\w\ball z r-{z}. f w = g w * (w-z) powr n \ g w\0 \ f w = h w * (w - z) powr m \ h w\0" + and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" + shows "n=m" +proof - + have [simp]:"at z within ball z r \ bot" using \r>0\ + by (auto simp add:at_within_ball_bot_iff) + have False when "n>m" + proof - + have "(h \ 0) (at z within ball z r)" + proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (n - m) * g w"]) + have "\w\ball z r-{z}. h w = (w-z)powr(n-m) * g w" + using \n>m\ asm \r>0\ + apply (auto simp add:field_simps powr_diff) + by force + then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ + \ (x' - z) powr (n - m) * g x' = h x'" for x' by auto + next + define F where "F \ at z within ball z r" + define f' where "f' \ \x. (x - z) powr (n-m)" + have "f' z=0" using \n>m\ unfolding f'_def by auto + moreover have "continuous F f'" unfolding f'_def F_def continuous_def + apply (subst Lim_ident_at) + using \n>m\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) + ultimately have "(f' \ 0) F" unfolding F_def + by (simp add: continuous_within) + moreover have "(g \ g z) F" + using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \r>0\ + unfolding F_def by auto + ultimately show " ((\w. f' w * g w) \ 0) F" using tendsto_mult by fastforce + qed + moreover have "(h \ h z) (at z within ball z r)" + using holomorphic_on_imp_continuous_on[OF h_holo] + by (auto simp add:continuous_on_def \r>0\) + ultimately have "h z=0" by (auto intro!: tendsto_unique) + thus False using \h z\0\ by auto + qed + moreover have False when "m>n" + proof - + have "(g \ 0) (at z within ball z r)" + proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (m - n) * h w"]) + have "\w\ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \m>n\ asm + apply (auto simp add:field_simps powr_diff) + by force + then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ + \ (x' - z) powr (m - n) * h x' = g x'" for x' by auto + next + define F where "F \ at z within ball z r" + define f' where "f' \\x. (x - z) powr (m-n)" + have "f' z=0" using \m>n\ unfolding f'_def by auto + moreover have "continuous F f'" unfolding f'_def F_def continuous_def + apply (subst Lim_ident_at) + using \m>n\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) + ultimately have "(f' \ 0) F" unfolding F_def + by (simp add: continuous_within) + moreover have "(h \ h z) F" + using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \r>0\ + unfolding F_def by auto + ultimately show " ((\w. f' w * h w) \ 0) F" using tendsto_mult by fastforce + qed + moreover have "(g \ g z) (at z within ball z r)" + using holomorphic_on_imp_continuous_on[OF g_holo] + by (auto simp add:continuous_on_def \r>0\) + ultimately have "g z=0" by (auto intro!: tendsto_unique) + thus False using \g z\0\ by auto + qed + ultimately show "n=m" by fastforce +qed + +lemma holomorphic_factor_puncture: + assumes f_iso:"isolated_singularity_at f z" + and "not_essential f z" \ \\<^term>\f\ has either a removable singularity or a pole at \<^term>\z\\ + and non_zero:"\\<^sub>Fw in (at z). f w\0" \ \\<^term>\f\ will not be constantly zero in a neighbour of \<^term>\z\\ + shows "\!n::int. \g r. 0 < r \ g holomorphic_on cball z r \ g z\0 + \ (\w\cball z r-{z}. f w = g w * (w-z) powr n \ g w\0)" +proof - + define P where "P = (\f n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 + \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" + have imp_unique:"\!n::int. \g r. P f n g r" when "\n g r. P f n g r" + proof (rule ex_ex1I[OF that]) + fix n1 n2 :: int + assume g1_asm:"\g1 r1. P f n1 g1 r1" and g2_asm:"\g2 r2. P f n2 g2 r2" + define fac where "fac \ \n g r. \w\cball z r-{z}. f w = g w * (w - z) powr (of_int n) \ g w \ 0" + obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\0" + and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto + obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\0" + and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto + define r where "r \ min r1 r2" + have "r>0" using \r1>0\ \r2>0\ unfolding r_def by auto + moreover have "\w\ball z r-{z}. f w = g1 w * (w-z) powr n1 \ g1 w\0 + \ f w = g2 w * (w - z) powr n2 \ g2 w\0" + using \fac n1 g1 r1\ \fac n2 g2 r2\ unfolding fac_def r_def + by fastforce + ultimately show "n1=n2" using g1_holo g2_holo \g1 z\0\ \g2 z\0\ + apply (elim holomorphic_factor_unique) + by (auto simp add:r_def) + qed + + have P_exist:"\ n g r. P h n g r" when + "\z'. (h \ z') (at z)" "isolated_singularity_at h z" "\\<^sub>Fw in (at z). h w\0" + for h + proof - + from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}" + unfolding isolated_singularity_at_def by auto + obtain z' where "(h \ z') (at z)" using \\z'. (h \ z') (at z)\ by auto + define h' where "h'=(\x. if x=z then z' else h x)" + have "h' holomorphic_on ball z r" + apply (rule no_isolated_singularity'[of "{z}"]) + subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \h \z\ z'\ empty_iff h'_def insert_iff) + subgoal using \h analytic_on ball z r - {z}\ analytic_imp_holomorphic h'_def holomorphic_transform + by fastforce + by auto + have ?thesis when "z'=0" + proof - + have "h' z=0" using that unfolding h'_def by auto + moreover have "\ h' constant_on ball z r" + using \\\<^sub>Fw in (at z). h w\0\ unfolding constant_on_def frequently_def eventually_at h'_def + apply simp + by (metis \0 < r\ centre_in_ball dist_commute mem_ball that) + moreover note \h' holomorphic_on ball z r\ + ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \ ball z r" and + g:"g holomorphic_on ball z r1" + "\w. w \ ball z r1 \ h' w = (w - z) ^ n * g w" + "\w. w \ ball z r1 \ g w \ 0" + using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified, + OF \h' holomorphic_on ball z r\ \r>0\ \h' z=0\ \\ h' constant_on ball z r\] + by (auto simp add:dist_commute) + define rr where "rr=r1/2" + have "P h' n g rr" + unfolding P_def rr_def + using \n>0\ \r1>0\ g by (auto simp add:powr_nat) + then have "P h n g rr" + unfolding h'_def P_def by auto + then show ?thesis unfolding P_def by blast + qed + moreover have ?thesis when "z'\0" + proof - + have "h' z\0" using that unfolding h'_def by auto + obtain r1 where "r1>0" "cball z r1 \ ball z r" "\x\cball z r1. h' x\0" + proof - + have "isCont h' z" "h' z\0" + by (auto simp add: Lim_cong_within \h \z\ z'\ \z'\0\ continuous_at h'_def) + then obtain r2 where r2:"r2>0" "\x\ball z r2. h' x\0" + using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto + define r1 where "r1=min r2 r / 2" + have "0 < r1" "cball z r1 \ ball z r" + using \r2>0\ \r>0\ unfolding r1_def by auto + moreover have "\x\cball z r1. h' x \ 0" + using r2 unfolding r1_def by simp + ultimately show ?thesis using that by auto + qed + then have "P h' 0 h' r1" using \h' holomorphic_on ball z r\ unfolding P_def by auto + then have "P h 0 h' r1" unfolding P_def h'_def by auto + then show ?thesis unfolding P_def by blast + qed + ultimately show ?thesis by auto + qed + + have ?thesis when "\x. (f \ x) (at z)" + apply (rule_tac imp_unique[unfolded P_def]) + using P_exist[OF that(1) f_iso non_zero] unfolding P_def . + moreover have ?thesis when "is_pole f z" + proof (rule imp_unique[unfolded P_def]) + obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\x\ball z e-{z}. f x\0" + proof - + have "\\<^sub>F z in at z. f z \ 0" + using \is_pole f z\ filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def + by auto + then obtain e1 where e1:"e1>0" "\x\ball z e1-{z}. f x\0" + using that eventually_at[of "\x. f x\0" z UNIV,simplified] by (auto simp add:dist_commute) + obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}" + using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto + define e where "e=min e1 e2" + show ?thesis + apply (rule that[of e]) + using e1 e2 unfolding e_def by auto + qed + + define h where "h \ \x. inverse (f x)" + + have "\n g r. P h n g r" + proof - + have "h \z\ 0" + using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce + moreover have "\\<^sub>Fw in (at z). h w\0" + using non_zero + apply (elim frequently_rev_mp) + unfolding h_def eventually_at by (auto intro:exI[where x=1]) + moreover have "isolated_singularity_at h z" + unfolding isolated_singularity_at_def h_def + apply (rule exI[where x=e]) + using e_holo e_nz \e>0\ by (metis open_ball analytic_on_open + holomorphic_on_inverse open_delete) + ultimately show ?thesis + using P_exist[of h] by auto + qed + then obtain n g r + where "0 < r" and + g_holo:"g holomorphic_on cball z r" and "g z\0" and + g_fac:"(\w\cball z r-{z}. h w = g w * (w - z) powr of_int n \ g w \ 0)" + unfolding P_def by auto + have "P f (-n) (inverse o g) r" + proof - + have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\cball z r - {z}" for w + using g_fac[rule_format,of w] that unfolding h_def + apply (auto simp add:powr_minus ) + by (metis inverse_inverse_eq inverse_mult_distrib) + then show ?thesis + unfolding P_def comp_def + using \r>0\ g_holo g_fac \g z\0\ by (auto intro:holomorphic_intros) + qed + then show "\x g r. 0 < r \ g holomorphic_on cball z r \ g z \ 0 + \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int x \ g w \ 0)" + unfolding P_def by blast + qed + ultimately show ?thesis using \not_essential f z\ unfolding not_essential_def by presburger +qed + +lemma not_essential_transform: + assumes "not_essential g z" + assumes "\\<^sub>F w in (at z). g w = f w" + shows "not_essential f z" + using assms unfolding not_essential_def + by (simp add: filterlim_cong is_pole_cong) + +lemma isolated_singularity_at_transform: + assumes "isolated_singularity_at g z" + assumes "\\<^sub>F w in (at z). g w = f w" + shows "isolated_singularity_at f z" +proof - + obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}" + using assms(1) unfolding isolated_singularity_at_def by auto + obtain r2 where "r2>0" and r2:" \x. x \ z \ dist x z < r2 \ g x = f x" + using assms(2) unfolding eventually_at by auto + define r3 where "r3=min r1 r2" + have "r3>0" unfolding r3_def using \r1>0\ \r2>0\ by auto + moreover have "f analytic_on ball z r3 - {z}" + proof - + have "g holomorphic_on ball z r3 - {z}" + using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto) + then have "f holomorphic_on ball z r3 - {z}" + using r2 unfolding r3_def + by (auto simp add:dist_commute elim!:holomorphic_transform) + then show ?thesis by (subst analytic_on_open,auto) + qed + ultimately show ?thesis unfolding isolated_singularity_at_def by auto +qed + +lemma not_essential_powr[singularity_intros]: + assumes "LIM w (at z). f w :> (at x)" + shows "not_essential (\w. (f w) powr (of_int n)) z" +proof - + define fp where "fp=(\w. (f w) powr (of_int n))" + have ?thesis when "n>0" + proof - + have "(\w. (f w) ^ (nat n)) \z\ x ^ nat n" + using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros) + then have "fp \z\ x ^ nat n" unfolding fp_def + apply (elim Lim_transform_within[where d=1],simp) + by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power) + then show ?thesis unfolding not_essential_def fp_def by auto + qed + moreover have ?thesis when "n=0" + proof - + have "fp \z\ 1 " + apply (subst tendsto_cong[where g="\_.1"]) + using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto + then show ?thesis unfolding fp_def not_essential_def by auto + qed + moreover have ?thesis when "n<0" + proof (cases "x=0") + case True + have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity" + apply (subst filterlim_inverse_at_iff[symmetric],simp) + apply (rule filterlim_atI) + subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros) + subgoal using filterlim_at_within_not_equal[OF assms,of 0] + by (eventually_elim,insert that,auto) + done + then have "LIM w (at z). fp w :> at_infinity" + proof (elim filterlim_mono_eventually) + show "\\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x" + using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def + apply eventually_elim + using powr_of_int that by auto + qed auto + then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto + next + case False + let ?xx= "inverse (x ^ (nat (-n)))" + have "(\w. inverse ((f w) ^ (nat (-n)))) \z\?xx" + using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros) + then have "fp \z\?xx" + apply (elim Lim_transform_within[where d=1],simp) + unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less + not_le power_eq_0_iff powr_0 powr_of_int that) + then show ?thesis unfolding fp_def not_essential_def by auto + qed + ultimately show ?thesis by linarith +qed + +lemma isolated_singularity_at_powr[singularity_intros]: + assumes "isolated_singularity_at f z" "\\<^sub>F w in (at z). f w\0" + shows "isolated_singularity_at (\w. (f w) powr (of_int n)) z" +proof - + obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}" + using assms(1) unfolding isolated_singularity_at_def by auto + then have r1:"f holomorphic_on ball z r1 - {z}" + using analytic_on_open[of "ball z r1-{z}" f] by blast + obtain r2 where "r2>0" and r2:"\w. w \ z \ dist w z < r2 \ f w \ 0" + using assms(2) unfolding eventually_at by auto + define r3 where "r3=min r1 r2" + have "(\w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}" + apply (rule holomorphic_on_powr_of_int) + subgoal unfolding r3_def using r1 by auto + subgoal unfolding r3_def using r2 by (auto simp add:dist_commute) + done + moreover have "r3>0" unfolding r3_def using \0 < r1\ \0 < r2\ by linarith + ultimately show ?thesis unfolding isolated_singularity_at_def + apply (subst (asm) analytic_on_open[symmetric]) + by auto +qed + +lemma non_zero_neighbour: + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" + and f_nconst:"\\<^sub>Fw in (at z). f w\0" + shows "\\<^sub>F w in (at z). f w\0" +proof - + obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" + and fr: "fp holomorphic_on cball z fr" + "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" + using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto + have "f w \ 0" when " w \ z" "dist w z < fr" for w + proof - + have "f w = fp w * (w - z) powr of_int fn" "fp w \ 0" + using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute) + moreover have "(w - z) powr of_int fn \0" + unfolding powr_eq_0_iff using \w\z\ by auto + ultimately show ?thesis by auto + qed + then show ?thesis using \fr>0\ unfolding eventually_at by auto +qed + +lemma non_zero_neighbour_pole: + assumes "is_pole f z" + shows "\\<^sub>F w in (at z). f w\0" + using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0] + unfolding is_pole_def by auto + +lemma non_zero_neighbour_alt: + assumes holo: "f holomorphic_on S" + and "open S" "connected S" "z \ S" "\ \ S" "f \ \ 0" + shows "\\<^sub>F w in (at z). f w\0 \ w\S" +proof (cases "f z = 0") + case True + from isolated_zeros[OF holo \open S\ \connected S\ \z \ S\ True \\ \ S\ \f \ \ 0\] + obtain r where "0 < r" "ball z r \ S" "\w \ ball z r - {z}.f w \ 0" by metis + then show ?thesis unfolding eventually_at + apply (rule_tac x=r in exI) + by (auto simp add:dist_commute) +next + case False + obtain r1 where r1:"r1>0" "\y. dist z y < r1 \ f y \ 0" + using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at + holo holomorphic_on_imp_continuous_on by blast + obtain r2 where r2:"r2>0" "ball z r2 \ S" + using assms(2) assms(4) openE by blast + show ?thesis unfolding eventually_at + apply (rule_tac x="min r1 r2" in exI) + using r1 r2 by (auto simp add:dist_commute) +qed + +lemma not_essential_times[singularity_intros]: + assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" + assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" + shows "not_essential (\w. f w * g w) z" +proof - + define fg where "fg = (\w. f w * g w)" + have ?thesis when "\ ((\\<^sub>Fw in (at z). f w\0) \ (\\<^sub>Fw in (at z). g w\0))" + proof - + have "\\<^sub>Fw in (at z). fg w=0" + using that[unfolded frequently_def, simplified] unfolding fg_def + by (auto elim: eventually_rev_mp) + from tendsto_cong[OF this] have "fg \z\0" by auto + then show ?thesis unfolding not_essential_def fg_def by auto + qed + moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" and g_nconst:"\\<^sub>Fw in (at z). g w\0" + proof - + obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" + and fr: "fp holomorphic_on cball z fr" + "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" + using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto + obtain gn gp gr where [simp]:"gp z \ 0" and "gr > 0" + and gr: "gp holomorphic_on cball z gr" + "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" + using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto + + define r1 where "r1=(min fr gr)" + have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto + have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" + when "w\ball z r1 - {z}" for w + proof - + have "f w = fp w * (w - z) powr of_int fn" "fp w\0" + using fr(2)[rule_format,of w] that unfolding r1_def by auto + moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" + using gr(2)[rule_format, of w] that unfolding r1_def by auto + ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" + unfolding fg_def by (auto simp add:powr_add) + qed + + have [intro]: "fp \z\fp z" "gp \z\gp z" + using fr(1) \fr>0\ gr(1) \gr>0\ + by (meson open_ball ball_subset_cball centre_in_ball + continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on + holomorphic_on_subset)+ + have ?thesis when "fn+gn>0" + proof - + have "(\w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \z\0" + using that by (auto intro!:tendsto_eq_intros) + then have "fg \z\ 0" + apply (elim Lim_transform_within[OF _ \r1>0\]) + by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self + eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int + that) + then show ?thesis unfolding not_essential_def fg_def by auto + qed + moreover have ?thesis when "fn+gn=0" + proof - + have "(\w. fp w * gp w) \z\fp z*gp z" + using that by (auto intro!:tendsto_eq_intros) + then have "fg \z\ fp z*gp z" + apply (elim Lim_transform_within[OF _ \r1>0\]) + apply (subst fg_times) + by (auto simp add:dist_commute that) + then show ?thesis unfolding not_essential_def fg_def by auto + qed + moreover have ?thesis when "fn+gn<0" + proof - + have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity" + apply (rule filterlim_divide_at_infinity) + apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI) + using eventually_at_topological by blast + then have "is_pole fg z" unfolding is_pole_def + apply (elim filterlim_transform_within[OF _ _ \r1>0\],simp) + apply (subst fg_times,simp add:dist_commute) + apply (subst powr_of_int) + using that by (auto simp add:field_split_simps) + then show ?thesis unfolding not_essential_def fg_def by auto + qed + ultimately show ?thesis unfolding not_essential_def fg_def by fastforce + qed + ultimately show ?thesis by auto +qed + +lemma not_essential_inverse[singularity_intros]: + assumes f_ness:"not_essential f z" + assumes f_iso:"isolated_singularity_at f z" + shows "not_essential (\w. inverse (f w)) z" +proof - + define vf where "vf = (\w. inverse (f w))" + have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" + proof - + have "\\<^sub>Fw in (at z). f w=0" + using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) + then have "\\<^sub>Fw in (at z). vf w=0" + unfolding vf_def by auto + from tendsto_cong[OF this] have "vf \z\0" unfolding vf_def by auto + then show ?thesis unfolding not_essential_def vf_def by auto + qed + moreover have ?thesis when "is_pole f z" + proof - + have "vf \z\0" + using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast + then show ?thesis unfolding not_essential_def vf_def by auto + qed + moreover have ?thesis when "\x. f\z\x " and f_nconst:"\\<^sub>Fw in (at z). f w\0" + proof - + from that obtain fz where fz:"f\z\fz" by auto + have ?thesis when "fz=0" + proof - + have "(\w. inverse (vf w)) \z\0" + using fz that unfolding vf_def by auto + moreover have "\\<^sub>F w in at z. inverse (vf w) \ 0" + using non_zero_neighbour[OF f_iso f_ness f_nconst] + unfolding vf_def by auto + ultimately have "is_pole vf z" + using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto + then show ?thesis unfolding not_essential_def vf_def by auto + qed + moreover have ?thesis when "fz\0" + proof - + have "vf \z\inverse fz" + using fz that unfolding vf_def by (auto intro:tendsto_eq_intros) + then show ?thesis unfolding not_essential_def vf_def by auto + qed + ultimately show ?thesis by auto + qed + ultimately show ?thesis using f_ness unfolding not_essential_def by auto +qed + +lemma isolated_singularity_at_inverse[singularity_intros]: + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" + shows "isolated_singularity_at (\w. inverse (f w)) z" +proof - + define vf where "vf = (\w. inverse (f w))" + have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" + proof - + have "\\<^sub>Fw in (at z). f w=0" + using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) + then have "\\<^sub>Fw in (at z). vf w=0" + unfolding vf_def by auto + then obtain d1 where "d1>0" and d1:"\x. x \ z \ dist x z < d1 \ vf x = 0" + unfolding eventually_at by auto + then have "vf holomorphic_on ball z d1-{z}" + apply (rule_tac holomorphic_transform[of "\_. 0"]) + by (auto simp add:dist_commute) + then have "vf analytic_on ball z d1 - {z}" + by (simp add: analytic_on_open open_delete) + then show ?thesis using \d1>0\ unfolding isolated_singularity_at_def vf_def by auto + qed + moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" + proof - + have "\\<^sub>F w in at z. f w \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] . + then obtain d1 where d1:"d1>0" "\x. x \ z \ dist x z < d1 \ f x \ 0" + unfolding eventually_at by auto + obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}" + using f_iso unfolding isolated_singularity_at_def by auto + define d3 where "d3=min d1 d2" + have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto + moreover have "vf analytic_on ball z d3 - {z}" + unfolding vf_def + apply (rule analytic_on_inverse) + subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto + subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute) + done + ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto + qed + ultimately show ?thesis by auto +qed + +lemma not_essential_divide[singularity_intros]: + assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" + assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" + shows "not_essential (\w. f w / g w) z" +proof - + have "not_essential (\w. f w * inverse (g w)) z" + apply (rule not_essential_times[where g="\w. inverse (g w)"]) + using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse) + then show ?thesis by (simp add:field_simps) +qed + +lemma + assumes f_iso:"isolated_singularity_at f z" + and g_iso:"isolated_singularity_at g z" + shows isolated_singularity_at_times[singularity_intros]: + "isolated_singularity_at (\w. f w * g w) z" and + isolated_singularity_at_add[singularity_intros]: + "isolated_singularity_at (\w. f w + g w) z" +proof - + obtain d1 d2 where "d1>0" "d2>0" + and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}" + using f_iso g_iso unfolding isolated_singularity_at_def by auto + define d3 where "d3=min d1 d2" + have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto + + have "(\w. f w * g w) analytic_on ball z d3 - {z}" + apply (rule analytic_on_mult) + using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) + then show "isolated_singularity_at (\w. f w * g w) z" + using \d3>0\ unfolding isolated_singularity_at_def by auto + have "(\w. f w + g w) analytic_on ball z d3 - {z}" + apply (rule analytic_on_add) + using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) + then show "isolated_singularity_at (\w. f w + g w) z" + using \d3>0\ unfolding isolated_singularity_at_def by auto +qed + +lemma isolated_singularity_at_uminus[singularity_intros]: + assumes f_iso:"isolated_singularity_at f z" + shows "isolated_singularity_at (\w. - f w) z" + using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast + +lemma isolated_singularity_at_id[singularity_intros]: + "isolated_singularity_at (\w. w) z" + unfolding isolated_singularity_at_def by (simp add: gt_ex) + +lemma isolated_singularity_at_minus[singularity_intros]: + assumes f_iso:"isolated_singularity_at f z" + and g_iso:"isolated_singularity_at g z" + shows "isolated_singularity_at (\w. f w - g w) z" + using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\w. - g w"] + ,OF g_iso] by simp + +lemma isolated_singularity_at_divide[singularity_intros]: + assumes f_iso:"isolated_singularity_at f z" + and g_iso:"isolated_singularity_at g z" + and g_ness:"not_essential g z" + shows "isolated_singularity_at (\w. f w / g w) z" + using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso, + of "\w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps) + +lemma isolated_singularity_at_const[singularity_intros]: + "isolated_singularity_at (\w. c) z" + unfolding isolated_singularity_at_def by (simp add: gt_ex) + +lemma isolated_singularity_at_holomorphic: + assumes "f holomorphic_on s-{z}" "open s" "z\s" + shows "isolated_singularity_at f z" + using assms unfolding isolated_singularity_at_def + by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) + +subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ + + +definition\<^marker>\tag important\ zorder :: "(complex \ complex) \ complex \ int" where + "zorder f z = (THE n. (\h r. r>0 \ h holomorphic_on cball z r \ h z\0 + \ (\w\cball z r - {z}. f w = h w * (w-z) powr (of_int n) + \ h w \0)))" + +definition\<^marker>\tag important\ zor_poly + ::"[complex \ complex, complex] \ complex \ complex" where + "zor_poly f z = (SOME h. \r. r > 0 \ h holomorphic_on cball z r \ h z \ 0 + \ (\w\cball z r - {z}. f w = h w * (w - z) powr (zorder f z) + \ h w \0))" + +lemma zorder_exist: + fixes f::"complex \ complex" and z::complex + defines "n\zorder f z" and "g\zor_poly f z" + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" + and f_nconst:"\\<^sub>Fw in (at z). f w\0" + shows "g z\0 \ (\r. r>0 \ g holomorphic_on cball z r + \ (\w\cball z r - {z}. f w = g w * (w-z) powr n \ g w \0))" +proof - + define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 + \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" + have "\!n. \g r. P n g r" + using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto + then have "\g r. P n g r" + unfolding n_def P_def zorder_def + by (drule_tac theI',argo) + then have "\r. P n g r" + unfolding P_def zor_poly_def g_def n_def + by (drule_tac someI_ex,argo) + then obtain r1 where "P n g r1" by auto + then show ?thesis unfolding P_def by auto +qed + +lemma + fixes f::"complex \ complex" and z::complex + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" + and f_nconst:"\\<^sub>Fw in (at z). f w\0" + shows zorder_inverse: "zorder (\w. inverse (f w)) z = - zorder f z" + and zor_poly_inverse: "\\<^sub>Fw in (at z). zor_poly (\w. inverse (f w)) z w + = inverse (zor_poly f z w)" +proof - + define vf where "vf = (\w. inverse (f w))" + define fn vfn where + "fn = zorder f z" and "vfn = zorder vf z" + define fp vfp where + "fp = zor_poly f z" and "vfp = zor_poly vf z" + + obtain fr where [simp]:"fp z \ 0" and "fr > 0" + and fr: "fp holomorphic_on cball z fr" + "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" + using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def] + by auto + have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" + and fr_nz: "inverse (fp w)\0" + when "w\ball z fr - {z}" for w + proof - + have "f w = fp w * (w - z) powr of_int fn" "fp w\0" + using fr(2)[rule_format,of w] that by auto + then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\0" + unfolding vf_def by (auto simp add:powr_minus) + qed + obtain vfr where [simp]:"vfp z \ 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" + "(\w\cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0)" + proof - + have "isolated_singularity_at vf z" + using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . + moreover have "not_essential vf z" + using not_essential_inverse[OF f_ness f_iso] unfolding vf_def . + moreover have "\\<^sub>F w in at z. vf w \ 0" + using f_nconst unfolding vf_def by (auto elim:frequently_elim1) + ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto + qed + + + define r1 where "r1 = min fr vfr" + have "r1>0" using \fr>0\ \vfr>0\ unfolding r1_def by simp + show "vfn = - fn" + apply (rule holomorphic_factor_unique[of r1 vfp z "\w. inverse (fp w)" vf]) + subgoal using \r1>0\ by simp + subgoal by simp + subgoal by simp + subgoal + proof (rule ballI) + fix w assume "w \ ball z r1 - {z}" + then have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" unfolding r1_def by auto + from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] + show "vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0 + \ vf w = inverse (fp w) * (w - z) powr of_int (- fn) \ inverse (fp w) \ 0" by auto + qed + subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) + subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros) + done + + have "vfp w = inverse (fp w)" when "w\ball z r1-{z}" for w + proof - + have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" "w\z" using that unfolding r1_def by auto + from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \vfn = - fn\ \w\z\ + show ?thesis by auto + qed + then show "\\<^sub>Fw in (at z). vfp w = inverse (fp w)" + unfolding eventually_at using \r1>0\ + apply (rule_tac x=r1 in exI) + by (auto simp add:dist_commute) +qed + +lemma + fixes f g::"complex \ complex" and z::complex + assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" + and f_ness:"not_essential f z" and g_ness:"not_essential g z" + and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" + shows zorder_times:"zorder (\w. f w * g w) z = zorder f z + zorder g z" and + zor_poly_times:"\\<^sub>Fw in (at z). zor_poly (\w. f w * g w) z w + = zor_poly f z w *zor_poly g z w" +proof - + define fg where "fg = (\w. f w * g w)" + define fn gn fgn where + "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z" + define fp gp fgp where + "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z" + have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" + using fg_nconst by (auto elim!:frequently_elim1) + obtain fr where [simp]:"fp z \ 0" and "fr > 0" + and fr: "fp holomorphic_on cball z fr" + "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" + using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto + obtain gr where [simp]:"gp z \ 0" and "gr > 0" + and gr: "gp holomorphic_on cball z gr" + "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" + using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto + define r1 where "r1=min fr gr" + have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto + have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" + when "w\ball z r1 - {z}" for w + proof - + have "f w = fp w * (w - z) powr of_int fn" "fp w\0" + using fr(2)[rule_format,of w] that unfolding r1_def by auto + moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" + using gr(2)[rule_format, of w] that unfolding r1_def by auto + ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" + unfolding fg_def by (auto simp add:powr_add) + qed + + obtain fgr where [simp]:"fgp z \ 0" and "fgr > 0" + and fgr: "fgp holomorphic_on cball z fgr" + "\w\cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0" + proof - + have "fgp z \ 0 \ (\r>0. fgp holomorphic_on cball z r + \ (\w\cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0))" + apply (rule zorder_exist[of fg z, folded fgn_def fgp_def]) + subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] . + subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] . + subgoal unfolding fg_def using fg_nconst . + done + then show ?thesis using that by blast + qed + define r2 where "r2 = min fgr r1" + have "r2>0" using \r1>0\ \fgr>0\ unfolding r2_def by simp + show "fgn = fn + gn " + apply (rule holomorphic_factor_unique[of r2 fgp z "\w. fp w * gp w" fg]) + subgoal using \r2>0\ by simp + subgoal by simp + subgoal by simp + subgoal + proof (rule ballI) + fix w assume "w \ ball z r2 - {z}" + then have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" unfolding r2_def by auto + from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] + show "fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0 + \ fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \ fp w * gp w \ 0" by auto + qed + subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) + subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) + done + + have "fgp w = fp w *gp w" when "w\ball z r2-{z}" for w + proof - + have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" "w\z" using that unfolding r2_def by auto + from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \fgn = fn + gn\ \w\z\ + show ?thesis by auto + qed + then show "\\<^sub>Fw in (at z). fgp w = fp w * gp w" + using \r2>0\ unfolding eventually_at by (auto simp add:dist_commute) +qed + +lemma + fixes f g::"complex \ complex" and z::complex + assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" + and f_ness:"not_essential f z" and g_ness:"not_essential g z" + and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" + shows zorder_divide:"zorder (\w. f w / g w) z = zorder f z - zorder g z" and + zor_poly_divide:"\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w + = zor_poly f z w / zor_poly g z w" +proof - + have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" + using fg_nconst by (auto elim!:frequently_elim1) + define vg where "vg=(\w. inverse (g w))" + have "zorder (\w. f w * vg w) z = zorder f z + zorder vg z" + apply (rule zorder_times[OF f_iso _ f_ness,of vg]) + subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . + subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . + subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) + done + then show "zorder (\w. f w / g w) z = zorder f z - zorder g z" + using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def + by (auto simp add:field_simps) + + have "\\<^sub>F w in at z. zor_poly (\w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w" + apply (rule zor_poly_times[OF f_iso _ f_ness,of vg]) + subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . + subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . + subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) + done + then show "\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" + using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def + apply eventually_elim + by (auto simp add:field_simps) +qed + +lemma zorder_exist_zero: + fixes f::"complex \ complex" and z::complex + defines "n\zorder f z" and "g\zor_poly f z" + assumes holo: "f holomorphic_on s" and + "open s" "connected s" "z\s" + and non_const: "\w\s. f w \ 0" + shows "(if f z=0 then n > 0 else n=0) \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r + \ (\w\cball z r. f w = g w * (w-z) ^ nat n \ g w \0))" +proof - + obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" + "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + proof - + have "g z \ 0 \ (\r>0. g holomorphic_on cball z r + \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" + proof (rule zorder_exist[of f z,folded g_def n_def]) + show "isolated_singularity_at f z" unfolding isolated_singularity_at_def + using holo assms(4,6) + by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE) + show "not_essential f z" unfolding not_essential_def + using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on + by fastforce + have "\\<^sub>F w in at z. f w \ 0 \ w\s" + proof - + obtain w where "w\s" "f w\0" using non_const by auto + then show ?thesis + by (rule non_zero_neighbour_alt[OF holo \open s\ \connected s\ \z\s\]) + qed + then show "\\<^sub>F w in at z. f w \ 0" + apply (elim eventually_frequentlyE) + by auto + qed + then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" + "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + by auto + obtain r2 where r2: "r2>0" "cball z r2 \ s" + using assms(4,6) open_contains_cball_eq by blast + define r3 where "r3=min r1 r2" + have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto + moreover have "g holomorphic_on cball z r3" + using r1(1) unfolding r3_def by auto + moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + using r1(2) unfolding r3_def by auto + ultimately show ?thesis using that[of r3] \g z\0\ by auto + qed + + have if_0:"if f z=0 then n > 0 else n=0" + proof - + have "f\ z \ f z" + by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on) + then have "(\w. g w * (w - z) powr of_int n) \z\ f z" + apply (elim Lim_transform_within_open[where s="ball z r"]) + using r by auto + moreover have "g \z\g z" + by (metis (mono_tags, lifting) open_ball at_within_open_subset + ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE) + ultimately have "(\w. (g w * (w - z) powr of_int n) / g w) \z\ f z/g z" + apply (rule_tac tendsto_divide) + using \g z\0\ by auto + then have powr_tendsto:"(\w. (w - z) powr of_int n) \z\ f z/g z" + apply (elim Lim_transform_within_open[where s="ball z r"]) + using r by auto + + have ?thesis when "n\0" "f z=0" + proof - + have "(\w. (w - z) ^ nat n) \z\ f z/g z" + using powr_tendsto + apply (elim Lim_transform_within[where d=r]) + by (auto simp add: powr_of_int \n\0\ \r>0\) + then have *:"(\w. (w - z) ^ nat n) \z\ 0" using \f z=0\ by simp + moreover have False when "n=0" + proof - + have "(\w. (w - z) ^ nat n) \z\ 1" + using \n=0\ by auto + then show False using * using LIM_unique zero_neq_one by blast + qed + ultimately show ?thesis using that by fastforce + qed + moreover have ?thesis when "n\0" "f z\0" + proof - + have False when "n>0" + proof - + have "(\w. (w - z) ^ nat n) \z\ f z/g z" + using powr_tendsto + apply (elim Lim_transform_within[where d=r]) + by (auto simp add: powr_of_int \n\0\ \r>0\) + moreover have "(\w. (w - z) ^ nat n) \z\ 0" + using \n>0\ by (auto intro!:tendsto_eq_intros) + ultimately show False using \f z\0\ \g z\0\ using LIM_unique divide_eq_0_iff by blast + qed + then show ?thesis using that by force + qed + moreover have False when "n<0" + proof - + have "(\w. inverse ((w - z) ^ nat (- n))) \z\ f z/g z" + "(\w.((w - z) ^ nat (- n))) \z\ 0" + subgoal using powr_tendsto powr_of_int that + by (elim Lim_transform_within_open[where s=UNIV],auto) + subgoal using that by (auto intro!:tendsto_eq_intros) + done + from tendsto_mult[OF this,simplified] + have "(\x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \z\ 0" . + then have "(\x. 1::complex) \z\ 0" + by (elim Lim_transform_within_open[where s=UNIV],auto) + then show False using LIM_const_eq by fastforce + qed + ultimately show ?thesis by fastforce + qed + moreover have "f w = g w * (w-z) ^ nat n \ g w \0" when "w\cball z r" for w + proof (cases "w=z") + case True + then have "f \z\f w" + using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce + then have "(\w. g w * (w-z) ^ nat n) \z\f w" + proof (elim Lim_transform_within[OF _ \r>0\]) + fix x assume "0 < dist x z" "dist x z < r" + then have "x \ cball z r - {z}" "x\z" + unfolding cball_def by (auto simp add: dist_commute) + then have "f x = g x * (x - z) powr of_int n" + using r(4)[rule_format,of x] by simp + also have "... = g x * (x - z) ^ nat n" + apply (subst powr_of_int) + using if_0 \x\z\ by (auto split:if_splits) + finally show "f x = g x * (x - z) ^ nat n" . + qed + moreover have "(\w. g w * (w-z) ^ nat n) \z\ g w * (w-z) ^ nat n" + using True apply (auto intro!:tendsto_eq_intros) + by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball + continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that) + ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast + then show ?thesis using \g z\0\ True by auto + next + case False + then have "f w = g w * (w - z) powr of_int n \ g w \ 0" + using r(4) that by auto + then show ?thesis using False if_0 powr_of_int by (auto split:if_splits) + qed + ultimately show ?thesis using r by auto +qed + +lemma zorder_exist_pole: + fixes f::"complex \ complex" and z::complex + defines "n\zorder f z" and "g\zor_poly f z" + assumes holo: "f holomorphic_on s-{z}" and + "open s" "z\s" + and "is_pole f z" + shows "n < 0 \ g z\0 \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r + \ (\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0))" +proof - + obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" + "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + proof - + have "g z \ 0 \ (\r>0. g holomorphic_on cball z r + \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" + proof (rule zorder_exist[of f z,folded g_def n_def]) + show "isolated_singularity_at f z" unfolding isolated_singularity_at_def + using holo assms(4,5) + by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) + show "not_essential f z" unfolding not_essential_def + using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on + by fastforce + from non_zero_neighbour_pole[OF \is_pole f z\] show "\\<^sub>F w in at z. f w \ 0" + apply (elim eventually_frequentlyE) + by auto + qed + then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" + "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + by auto + obtain r2 where r2: "r2>0" "cball z r2 \ s" + using assms(4,5) open_contains_cball_eq by metis + define r3 where "r3=min r1 r2" + have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto + moreover have "g holomorphic_on cball z r3" + using r1(1) unfolding r3_def by auto + moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + using r1(2) unfolding r3_def by auto + ultimately show ?thesis using that[of r3] \g z\0\ by auto + qed + + have "n<0" + proof (rule ccontr) + assume " \ n < 0" + define c where "c=(if n=0 then g z else 0)" + have [simp]:"g \z\ g z" + by (metis open_ball at_within_open ball_subset_cball centre_in_ball + continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) ) + have "\\<^sub>F x in at z. f x = g x * (x - z) ^ nat n" + unfolding eventually_at_topological + apply (rule_tac exI[where x="ball z r"]) + using r powr_of_int \\ n < 0\ by auto + moreover have "(\x. g x * (x - z) ^ nat n) \z\c" + proof (cases "n=0") + case True + then show ?thesis unfolding c_def by simp + next + case False + then have "(\x. (x - z) ^ nat n) \z\ 0" using \\ n < 0\ + by (auto intro!:tendsto_eq_intros) + from tendsto_mult[OF _ this,of g "g z",simplified] + show ?thesis unfolding c_def using False by simp + qed + ultimately have "f \z\c" using tendsto_cong by fast + then show False using \is_pole f z\ at_neq_bot not_tendsto_and_filterlim_at_infinity + unfolding is_pole_def by blast + qed + moreover have "\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0" + using r(4) \n<0\ powr_of_int + by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le) + ultimately show ?thesis using r(1-3) \g z\0\ by auto +qed + +lemma zorder_eqI: + assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" + assumes fg_eq:"\w. \w \ s;w\z\ \ f w = g w * (w - z) powr n" + shows "zorder f z = n" +proof - + have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact + moreover have "open (-{0::complex})" by auto + ultimately have "open ((g -` (-{0})) \ s)" + unfolding continuous_on_open_vimage[OF \open s\] by blast + moreover from assms have "z \ (g -` (-{0})) \ s" by auto + ultimately obtain r where r: "r > 0" "cball z r \ s \ (g -` (-{0}))" + unfolding open_contains_cball by blast + + let ?gg= "(\w. g w * (w - z) powr n)" + define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 + \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" + have "P n g r" + unfolding P_def using r assms(3,4,5) by auto + then have "\g r. P n g r" by auto + moreover have unique: "\!n. \g r. P n g r" unfolding P_def + proof (rule holomorphic_factor_puncture) + have "ball z r-{z} \ s" using r using ball_subset_cball by blast + then have "?gg holomorphic_on ball z r-{z}" + using \g holomorphic_on s\ r by (auto intro!: holomorphic_intros) + then have "f holomorphic_on ball z r - {z}" + apply (elim holomorphic_transform) + using fg_eq \ball z r-{z} \ s\ by auto + then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def + using analytic_on_open open_delete r(1) by blast + next + have "not_essential ?gg z" + proof (intro singularity_intros) + show "not_essential g z" + by (meson \continuous_on s g\ assms(1) assms(2) continuous_on_eq_continuous_at + isCont_def not_essential_def) + show " \\<^sub>F w in at z. w - z \ 0" by (simp add: eventually_at_filter) + then show "LIM w at z. w - z :> at 0" + unfolding filterlim_at by (auto intro:tendsto_eq_intros) + show "isolated_singularity_at g z" + by (meson Diff_subset open_ball analytic_on_holomorphic + assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE) + qed + then show "not_essential f z" + apply (elim not_essential_transform) + unfolding eventually_at using assms(1,2) assms(5)[symmetric] + by (metis dist_commute mem_ball openE subsetCE) + show "\\<^sub>F w in at z. f w \ 0" unfolding frequently_at + proof (rule,rule) + fix d::real assume "0 < d" + define z' where "z'=z+min d r / 2" + have "z' \ z" " dist z' z < d " + unfolding z'_def using \d>0\ \r>0\ + by (auto simp add:dist_norm) + moreover have "f z' \ 0" + proof (subst fg_eq[OF _ \z'\z\]) + have "z' \ cball z r" unfolding z'_def using \r>0\ \d>0\ by (auto simp add:dist_norm) + then show " z' \ s" using r(2) by blast + show "g z' * (z' - z) powr of_int n \ 0" + using P_def \P n g r\ \z' \ cball z r\ calculation(1) by auto + qed + ultimately show "\x\UNIV. x \ z \ dist x z < d \ f x \ 0" by auto + qed + qed + ultimately have "(THE n. \g r. P n g r) = n" + by (rule_tac the1_equality) + then show ?thesis unfolding zorder_def P_def by blast +qed + +lemma residue_pole_order: + fixes f::"complex \ complex" and z::complex + defines "n \ nat (- zorder f z)" and "h \ zor_poly f z" + assumes f_iso:"isolated_singularity_at f z" + and pole:"is_pole f z" + shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" +proof - + define g where "g \ \x. if x=z then 0 else inverse (f x)" + obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" + using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast + obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ ball z e" and h_holo: "h holomorphic_on cball z r" + and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" + proof - + obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" + "(\w\cball z r - {z}. f w = h w / (w - z) ^ n \ h w \ 0)" + using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\,folded n_def h_def] by auto + have "n>0" using \zorder f z < 0\ unfolding n_def by simp + moreover have "(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" + using \h z\0\ r(6) by blast + ultimately show ?thesis using r(3,4,5) that by blast + qed + have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" + using h_divide by simp + define c where "c \ 2 * pi * \" + define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" + define h' where "h' \ \u. h u / (u - z) ^ n" + have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" + unfolding h'_def + proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", + folded c_def Suc_pred'[OF \n>0\]]) + show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp + show "h holomorphic_on ball z r" using h_holo by auto + show " z \ ball z r" using \r>0\ by auto + qed + then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto + then have "(f has_contour_integral c * der_f) (circlepath z r)" + proof (elim has_contour_integral_eq) + fix x assume "x \ path_image (circlepath z r)" + hence "x\cball z r - {z}" using \r>0\ by auto + then show "h' x = f x" using h_divide unfolding h'_def by auto + qed + moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" + using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] + unfolding c_def by simp + ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast + hence "der_f = residue f z" unfolding c_def by auto + thus ?thesis unfolding der_f_def by auto +qed + +lemma simple_zeroI: + assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" + assumes "\w. w \ s \ f w = g w * (w - z)" + shows "zorder f z = 1" + using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto) + +lemma higher_deriv_power: + shows "(deriv ^^ j) (\w. (w - z) ^ n) w = + pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)" +proof (induction j arbitrary: w) + case 0 + thus ?case by auto +next + case (Suc j w) + have "(deriv ^^ Suc j) (\w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\w. (w - z) ^ n)) w" + by simp + also have "(deriv ^^ j) (\w. (w - z) ^ n) = + (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))" + using Suc by (intro Suc.IH ext) + also { + have "(\ has_field_derivative of_nat (n - j) * + pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)" + using Suc.prems by (auto intro!: derivative_eq_intros) + also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = + pochhammer (of_nat (Suc n - Suc j)) (Suc j)" + by (cases "Suc j \ n", subst pochhammer_rec) + (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left) + finally have "deriv (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w = + \ * (w - z) ^ (n - Suc j)" + by (rule DERIV_imp_deriv) + } + finally show ?case . +qed + +lemma zorder_zero_eqI: + assumes f_holo:"f holomorphic_on s" and "open s" "z \ s" + assumes zero: "\i. i < nat n \ (deriv ^^ i) f z = 0" + assumes nz: "(deriv ^^ nat n) f z \ 0" and "n\0" + shows "zorder f z = n" +proof - + obtain r where [simp]:"r>0" and "ball z r \ s" + using \open s\ \z\s\ openE by blast + have nz':"\w\ball z r. f w \ 0" + proof (rule ccontr) + assume "\ (\w\ball z r. f w \ 0)" + then have "eventually (\u. f u = 0) (nhds z)" + using \r>0\ unfolding eventually_nhds + apply (rule_tac x="ball z r" in exI) + by auto + then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\_. 0) z" + by (intro higher_deriv_cong_ev) auto + also have "(deriv ^^ nat n) (\_. 0) z = 0" + by (induction n) simp_all + finally show False using nz by contradiction + qed + + define zn g where "zn = zorder f z" and "g = zor_poly f z" + obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and + [simp]:"e>0" and "cball z e \ ball z r" and + g_holo:"g holomorphic_on cball z e" and + e_fac:"(\w\cball z e. f w = g w * (w - z) ^ nat zn \ g w \ 0)" + proof - + have "f holomorphic_on ball z r" + using f_holo \ball z r \ s\ by auto + from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def] + show ?thesis by blast + qed + from this(1,2,5) have "zn\0" "g z\0" + subgoal by (auto split:if_splits) + subgoal using \0 < e\ ball_subset_cball centre_in_ball e_fac by blast + done + + define A where "A = (\i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)" + have deriv_A:"(deriv ^^ i) f z = (if zn \ int i then A i else 0)" for i + proof - + have "eventually (\w. w \ ball z e) (nhds z)" + using \cball z e \ ball z r\ \e>0\ by (intro eventually_nhds_in_open) auto + hence "eventually (\w. f w = (w - z) ^ (nat zn) * g w) (nhds z)" + apply eventually_elim + by (use e_fac in auto) + hence "(deriv ^^ i) f z = (deriv ^^ i) (\w. (w - z) ^ nat zn * g w) z" + by (intro higher_deriv_cong_ev) auto + also have "\ = (\j=0..i. of_nat (i choose j) * + (deriv ^^ j) (\w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)" + using g_holo \e>0\ + by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros) + also have "\ = (\j=0..i. if j = nat zn then + of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" + proof (intro sum.cong refl, goal_cases) + case (1 j) + have "(deriv ^^ j) (\w. (w - z) ^ nat zn) z = + pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)" + by (subst higher_deriv_power) auto + also have "\ = (if j = nat zn then fact j else 0)" + by (auto simp: not_less pochhammer_0_left pochhammer_fact) + also have "of_nat (i choose j) * \ * (deriv ^^ (i - j)) g z = + (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) + * (deriv ^^ (i - nat zn)) g z else 0)" + by simp + finally show ?case . + qed + also have "\ = (if i \ zn then A i else 0)" + by (auto simp: A_def) + finally show "(deriv ^^ i) f z = \" . + qed + + have False when "nn\0\ by auto + with nz show False by auto + qed + moreover have "n\zn" + proof - + have "g z \ 0" using e_fac[rule_format,of z] \e>0\ by simp + then have "(deriv ^^ nat zn) f z \ 0" + using deriv_A[of "nat zn"] by(auto simp add:A_def) + then have "nat zn \ nat n" using zero[of "nat zn"] by linarith + moreover have "zn\0" using e_if by (auto split:if_splits) + ultimately show ?thesis using nat_le_eq_zle by blast + qed + ultimately show ?thesis unfolding zn_def by fastforce +qed + +lemma + assumes "eventually (\z. f z = g z) (at z)" "z = z'" + shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'" +proof - + define P where "P = (\ff n h r. 0 < r \ h holomorphic_on cball z r \ h z\0 + \ (\w\cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \ h w\0))" + have "(\r. P f n h r) = (\r. P g n h r)" for n h + proof - + have *: "\r. P g n h r" if "\r. P f n h r" and "eventually (\x. f x = g x) (at z)" for f g + proof - + from that(1) obtain r1 where r1_P:"P f n h r1" by auto + from that(2) obtain r2 where "r2>0" and r2_dist:"\x. x \ z \ dist x z \ r2 \ f x = g x" + unfolding eventually_at_le by auto + define r where "r=min r1 r2" + have "r>0" "h z\0" using r1_P \r2>0\ unfolding r_def P_def by auto + moreover have "h holomorphic_on cball z r" + using r1_P unfolding P_def r_def by auto + moreover have "g w = h w * (w - z) powr of_int n \ h w \ 0" when "w\cball z r - {z}" for w + proof - + have "f w = h w * (w - z) powr of_int n \ h w \ 0" + using r1_P that unfolding P_def r_def by auto + moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def + by (simp add: dist_commute) + ultimately show ?thesis by simp + qed + ultimately show ?thesis unfolding P_def by auto + qed + from assms have eq': "eventually (\z. g z = f z) (at z)" + by (simp add: eq_commute) + show ?thesis + by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']]) + qed + then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'" + using \z=z'\ unfolding P_def zorder_def zor_poly_def by auto +qed + +lemma zorder_nonzero_div_power: + assumes "open s" "z \ s" "f holomorphic_on s" "f z \ 0" "n > 0" + shows "zorder (\w. f w / (w - z) ^ n) z = - n" + apply (rule zorder_eqI[OF \open s\ \z\s\ \f holomorphic_on s\ \f z\0\]) + apply (subst powr_of_int) + using \n>0\ by (auto simp add:field_simps) + +lemma zor_poly_eq: + assumes "isolated_singularity_at f z" "not_essential f z" "\\<^sub>F w in at z. f w \ 0" + shows "eventually (\w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)" +proof - + obtain r where r:"r>0" + "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))" + using zorder_exist[OF assms] by blast + then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" + by (auto simp: field_simps powr_minus) + have "eventually (\w. w \ ball z r - {z}) (at z)" + using r eventually_at_ball'[of r z UNIV] by auto + thus ?thesis by eventually_elim (insert *, auto) +qed + +lemma zor_poly_zero_eq: + assumes "f holomorphic_on s" "open s" "connected s" "z \ s" "\w\s. f w \ 0" + shows "eventually (\w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)" +proof - + obtain r where r:"r>0" + "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))" + using zorder_exist_zero[OF assms] by auto + then have *: "\w\ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" + by (auto simp: field_simps powr_minus) + have "eventually (\w. w \ ball z r - {z}) (at z)" + using r eventually_at_ball'[of r z UNIV] by auto + thus ?thesis by eventually_elim (insert *, auto) +qed + +lemma zor_poly_pole_eq: + assumes f_iso:"isolated_singularity_at f z" "is_pole f z" + shows "eventually (\w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)" +proof - + obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" + using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast + obtain r where r:"r>0" + "(\w\cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))" + using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\] by auto + then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" + by (auto simp: field_simps) + have "eventually (\w. w \ ball z r - {z}) (at z)" + using r eventually_at_ball'[of r z UNIV] by auto + thus ?thesis by eventually_elim (insert *, auto) +qed + +lemma zor_poly_eqI: + fixes f :: "complex \ complex" and z0 :: complex + defines "n \ zorder f z0" + assumes "isolated_singularity_at f z0" "not_essential f z0" "\\<^sub>F w in at z0. f w \ 0" + assumes lim: "((\x. f (g x) * (g x - z0) powr - n) \ c) F" + assumes g: "filterlim g (at z0) F" and "F \ bot" + shows "zor_poly f z0 z0 = c" +proof - + from zorder_exist[OF assms(2-4)] obtain r where + r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" + "\w. w \ cball z0 r - {z0} \ f w = zor_poly f z0 w * (w - z0) powr n" + unfolding n_def by blast + from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" + using eventually_at_ball'[of r z0 UNIV] by auto + hence "eventually (\w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)" + by eventually_elim (insert r, auto simp: field_simps powr_minus) + moreover have "continuous_on (ball z0 r) (zor_poly f z0)" + using r by (intro holomorphic_on_imp_continuous_on) auto + with r(1,2) have "isCont (zor_poly f z0) z0" + by (auto simp: continuous_on_eq_continuous_at) + hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" + unfolding isCont_def . + ultimately have "((\w. f w * (w - z0) powr - n) \ zor_poly f z0 z0) (at z0)" + by (blast intro: Lim_transform_eventually) + hence "((\x. f (g x) * (g x - z0) powr - n) \ zor_poly f z0 z0) F" + by (rule filterlim_compose[OF _ g]) + from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . +qed + +lemma zor_poly_zero_eqI: + fixes f :: "complex \ complex" and z0 :: complex + defines "n \ zorder f z0" + assumes "f holomorphic_on A" "open A" "connected A" "z0 \ A" "\z\A. f z \ 0" + assumes lim: "((\x. f (g x) / (g x - z0) ^ nat n) \ c) F" + assumes g: "filterlim g (at z0) F" and "F \ bot" + shows "zor_poly f z0 z0 = c" +proof - + from zorder_exist_zero[OF assms(2-6)] obtain r where + r: "r > 0" "cball z0 r \ A" "zor_poly f z0 holomorphic_on cball z0 r" + "\w. w \ cball z0 r \ f w = zor_poly f z0 w * (w - z0) ^ nat n" + unfolding n_def by blast + from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" + using eventually_at_ball'[of r z0 UNIV] by auto + hence "eventually (\w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)" + by eventually_elim (insert r, auto simp: field_simps) + moreover have "continuous_on (ball z0 r) (zor_poly f z0)" + using r by (intro holomorphic_on_imp_continuous_on) auto + with r(1,2) have "isCont (zor_poly f z0) z0" + by (auto simp: continuous_on_eq_continuous_at) + hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" + unfolding isCont_def . + ultimately have "((\w. f w / (w - z0) ^ nat n) \ zor_poly f z0 z0) (at z0)" + by (blast intro: Lim_transform_eventually) + hence "((\x. f (g x) / (g x - z0) ^ nat n) \ zor_poly f z0 z0) F" + by (rule filterlim_compose[OF _ g]) + from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . +qed + +lemma zor_poly_pole_eqI: + fixes f :: "complex \ complex" and z0 :: complex + defines "n \ zorder f z0" + assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0" + assumes lim: "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ c) F" + assumes g: "filterlim g (at z0) F" and "F \ bot" + shows "zor_poly f z0 z0 = c" +proof - + obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" + proof - + have "\\<^sub>F w in at z0. f w \ 0" + using non_zero_neighbour_pole[OF \is_pole f z0\] by (auto elim:eventually_frequentlyE) + moreover have "not_essential f z0" unfolding not_essential_def using \is_pole f z0\ by simp + ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto + qed + from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" + using eventually_at_ball'[of r z0 UNIV] by auto + have "eventually (\w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)" + using zor_poly_pole_eq[OF f_iso \is_pole f z0\] unfolding n_def . + moreover have "continuous_on (ball z0 r) (zor_poly f z0)" + using r by (intro holomorphic_on_imp_continuous_on) auto + with r(1,2) have "isCont (zor_poly f z0) z0" + by (auto simp: continuous_on_eq_continuous_at) + hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" + unfolding isCont_def . + ultimately have "((\w. f w * (w - z0) ^ nat (-n)) \ zor_poly f z0 z0) (at z0)" + by (blast intro: Lim_transform_eventually) + hence "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ zor_poly f z0 z0) F" + by (rule filterlim_compose[OF _ g]) + from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . +qed + +lemma residue_simple_pole: + assumes "isolated_singularity_at f z0" + assumes "is_pole f z0" "zorder f z0 = - 1" + shows "residue f z0 = zor_poly f z0 z0" + using assms by (subst residue_pole_order) simp_all + +lemma residue_simple_pole_limit: + assumes "isolated_singularity_at f z0" + assumes "is_pole f z0" "zorder f z0 = - 1" + assumes "((\x. f (g x) * (g x - z0)) \ c) F" + assumes "filterlim g (at z0) F" "F \ bot" + shows "residue f z0 = c" +proof - + have "residue f z0 = zor_poly f z0 z0" + by (rule residue_simple_pole assms)+ + also have "\ = c" + apply (rule zor_poly_pole_eqI) + using assms by auto + finally show ?thesis . +qed + +lemma lhopital_complex_simple: + assumes "(f has_field_derivative f') (at z)" + assumes "(g has_field_derivative g') (at z)" + assumes "f z = 0" "g z = 0" "g' \ 0" "f' / g' = c" + shows "((\w. f w / g w) \ c) (at z)" +proof - + have "eventually (\w. w \ z) (at z)" + by (auto simp: eventually_at_filter) + hence "eventually (\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)" + by eventually_elim (simp add: assms field_split_simps) + moreover have "((\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \ f' / g') (at z)" + by (intro tendsto_divide has_field_derivativeD assms) + ultimately have "((\w. f w / g w) \ f' / g') (at z)" + by (blast intro: Lim_transform_eventually) + with assms show ?thesis by simp +qed + +lemma + assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" + and "open s" "connected s" "z \ s" + assumes g_deriv:"(g has_field_derivative g') (at z)" + assumes "f z \ 0" "g z = 0" "g' \ 0" + shows porder_simple_pole_deriv: "zorder (\w. f w / g w) z = - 1" + and residue_simple_pole_deriv: "residue (\w. f w / g w) z = f z / g'" +proof - + have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z" + using isolated_singularity_at_holomorphic[OF _ \open s\ \z\s\] f_holo g_holo + by (meson Diff_subset holomorphic_on_subset)+ + have [simp]:"not_essential f z" "not_essential g z" + unfolding not_essential_def using f_holo g_holo assms(3,5) + by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+ + have g_nconst:"\\<^sub>F w in at z. g w \0 " + proof (rule ccontr) + assume "\ (\\<^sub>F w in at z. g w \ 0)" + then have "\\<^sub>F w in nhds z. g w = 0" + unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ + by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball) + then have "deriv g z = deriv (\_. 0) z" + by (intro deriv_cong_ev) auto + then have "deriv g z = 0" by auto + then have "g' = 0" using g_deriv DERIV_imp_deriv by blast + then show False using \g'\0\ by auto + qed + + have "zorder (\w. f w / g w) z = zorder f z - zorder g z" + proof - + have "\\<^sub>F w in at z. f w \0 \ w\s" + apply (rule non_zero_neighbour_alt) + using assms by auto + with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" + by (elim frequently_rev_mp eventually_rev_mp,auto) + then show ?thesis using zorder_divide[of f z g] by auto + qed + moreover have "zorder f z=0" + apply (rule zorder_zero_eqI[OF f_holo \open s\ \z\s\]) + using \f z\0\ by auto + moreover have "zorder g z=1" + apply (rule zorder_zero_eqI[OF g_holo \open s\ \z\s\]) + subgoal using assms(8) by auto + subgoal using DERIV_imp_deriv assms(9) g_deriv by auto + subgoal by simp + done + ultimately show "zorder (\w. f w / g w) z = - 1" by auto + + show "residue (\w. f w / g w) z = f z / g'" + proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified]) + show "zorder (\w. f w / g w) z = - 1" by fact + show "isolated_singularity_at (\w. f w / g w) z" + by (auto intro: singularity_intros) + show "is_pole (\w. f w / g w) z" + proof (rule is_pole_divide) + have "\\<^sub>F x in at z. g x \ 0" + apply (rule non_zero_neighbour) + using g_nconst by auto + moreover have "g \z\ 0" + using DERIV_isCont assms(8) continuous_at g_deriv by force + ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp + show "isCont f z" + using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on + by auto + show "f z \ 0" by fact + qed + show "filterlim id (at z) (at z)" by (simp add: filterlim_iff) + have "((\w. (f w * (w - z)) / g w) \ f z / g') (at z)" + proof (rule lhopital_complex_simple) + show "((\w. f w * (w - z)) has_field_derivative f z) (at z)" + using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo]) + show "(g has_field_derivative g') (at z)" by fact + qed (insert assms, auto) + then show "((\w. (f w / g w) * (w - z)) \ f z / g') (at z)" + by (simp add: field_split_simps) + qed +qed + +subsection \The argument principle\ + +theorem argument_principle: + fixes f::"complex \ complex" and poles s:: "complex set" + defines "pz \ {w. f w = 0 \ w \ poles}" \ \\<^term>\pz\ is the set of poles and zeros\ + assumes "open s" and + "connected s" and + f_holo:"f holomorphic_on s-poles" and + h_holo:"h holomorphic_on s" and + "valid_path g" and + loop:"pathfinish g = pathstart g" and + path_img:"path_image g \ s - pz" and + homo:"\z. (z \ s) \ winding_number g z = 0" and + finite:"finite pz" and + poles:"\p\poles. is_pole f p" + shows "contour_integral g (\x. deriv f x * h x / f x) = 2 * pi * \ * + (\p\pz. winding_number g p * h p * zorder f p)" + (is "?L=?R") +proof - + define c where "c \ 2 * complex_of_real pi * \ " + define ff where "ff \ (\x. deriv f x * h x / f x)" + define cont where "cont \ \ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" + define avoid where "avoid \ \p e. \w\cball p e. w \ s \ (w \ p \ w \ pz)" + + have "\e>0. avoid p e \ (p\pz \ cont ff p e)" when "p\s" for p + proof - + obtain e1 where "e1>0" and e1_avoid:"avoid p e1" + using finite_cball_avoid[OF \open s\ finite] \p\s\ unfolding avoid_def by auto + have "\e2>0. cball p e2 \ ball p e1 \ cont ff p e2" when "p\pz" + proof - + define po where "po \ zorder f p" + define pp where "pp \ zor_poly f p" + define f' where "f' \ \w. pp w * (w - p) powr po" + define ff' where "ff' \ (\x. deriv f' x * h x / f' x)" + obtain r where "pp p\0" "r>0" and + "rw\cball p r-{p}. f w = pp w * (w - p) powr po \ pp w \ 0)" + proof - + have "isolated_singularity_at f p" + proof - + have "f holomorphic_on ball p e1 - {p}" + apply (intro holomorphic_on_subset[OF f_holo]) + using e1_avoid \p\pz\ unfolding avoid_def pz_def by force + then show ?thesis unfolding isolated_singularity_at_def + using \e1>0\ analytic_on_open open_delete by blast + qed + moreover have "not_essential f p" + proof (cases "is_pole f p") + case True + then show ?thesis unfolding not_essential_def by auto + next + case False + then have "p\s-poles" using \p\s\ poles unfolding pz_def by auto + moreover have "open (s-poles)" + using \open s\ + apply (elim open_Diff) + apply (rule finite_imp_closed) + using finite unfolding pz_def by simp + ultimately have "isCont f p" + using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at + by auto + then show ?thesis unfolding isCont_def not_essential_def by auto + qed + moreover have "\\<^sub>F w in at p. f w \ 0 " + proof (rule ccontr) + assume "\ (\\<^sub>F w in at p. f w \ 0)" + then have "\\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto + then obtain rr where "rr>0" "\w\ball p rr - {p}. f w =0" + unfolding eventually_at by (auto simp add:dist_commute) + then have "ball p rr - {p} \ {w\ball p rr-{p}. f w=0}" by blast + moreover have "infinite (ball p rr - {p})" using \rr>0\ using finite_imp_not_open by fastforce + ultimately have "infinite {w\ball p rr-{p}. f w=0}" using infinite_super by blast + then have "infinite pz" + unfolding pz_def infinite_super by auto + then show False using \finite pz\ by auto + qed + ultimately obtain r where "pp p \ 0" and r:"r>0" "pp holomorphic_on cball p r" + "(\w\cball p r - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" + using zorder_exist[of f p,folded po_def pp_def] by auto + define r1 where "r1=min r e1 / 2" + have "r1e1>0\ \r>0\ by auto + moreover have "r1>0" "pp holomorphic_on cball p r1" + "(\w\cball p r1 - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" + unfolding r1_def using \e1>0\ r by auto + ultimately show ?thesis using that \pp p\0\ by auto + qed + + define e2 where "e2 \ r/2" + have "e2>0" using \r>0\ unfolding e2_def by auto + define anal where "anal \ \w. deriv pp w * h w / pp w" + define prin where "prin \ \w. po * h w / (w - p)" + have "((\w. prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)" + proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) + have "ball p r \ s" + using \r avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq) + then have "cball p e2 \ s" + using \r>0\ unfolding e2_def by auto + then have "(\w. po * h w) holomorphic_on cball p e2" + using h_holo by (auto intro!: holomorphic_intros) + then show "(prin has_contour_integral c * po * h p ) (circlepath p e2)" + using Cauchy_integral_circlepath_simple[folded c_def, of "\w. po * h w"] \e2>0\ + unfolding prin_def by (auto simp add: mult.assoc) + have "anal holomorphic_on ball p r" unfolding anal_def + using pp_holo h_holo pp_po \ball p r \ s\ \pp p\0\ + by (auto intro!: holomorphic_intros) + then show "(anal has_contour_integral 0) (circlepath p e2)" + using e2_def \r>0\ + by (auto elim!: Cauchy_theorem_disc_simple) + qed + then have "cont ff' p e2" unfolding cont_def po_def + proof (elim has_contour_integral_eq) + fix w assume "w \ path_image (circlepath p e2)" + then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto + define wp where "wp \ w-p" + have "wp\0" and "pp w \0" + unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto + moreover have der_f':"deriv f' w = po * pp w * (w-p) powr (po - 1) + deriv pp w * (w-p) powr po" + proof (rule DERIV_imp_deriv) + have "(pp has_field_derivative (deriv pp w)) (at w)" + using DERIV_deriv_iff_has_field_derivative pp_holo \w\p\ + by (meson open_ball \w \ ball p r\ ball_subset_cball holomorphic_derivI holomorphic_on_subset) + then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1) + + deriv pp w * (w - p) powr of_int po) (at w)" + unfolding f'_def using \w\p\ + by (auto intro!: derivative_eq_intros DERIV_cong[OF has_field_derivative_powr_of_int]) + qed + ultimately show "prin w + anal w = ff' w" + unfolding ff'_def prin_def anal_def + apply simp + apply (unfold f'_def) + apply (fold wp_def) + apply (auto simp add:field_simps) + by (metis (no_types, lifting) diff_add_cancel mult.commute powr_add powr_to_1) + qed + then have "cont ff p e2" unfolding cont_def + proof (elim has_contour_integral_eq) + fix w assume "w \ path_image (circlepath p e2)" + then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto + have "deriv f' w = deriv f w" + proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) + show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo + by (auto intro!: holomorphic_intros) + next + have "ball p e1 - {p} \ s - poles" + using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def + by auto + then have "ball p r - {p} \ s - poles" + apply (elim dual_order.trans) + using \r by auto + then show "f holomorphic_on ball p r - {p}" using f_holo + by auto + next + show "open (ball p r - {p})" by auto + show "w \ ball p r - {p}" using \w\ball p r\ \w\p\ by auto + next + fix x assume "x \ ball p r - {p}" + then show "f' x = f x" + using pp_po unfolding f'_def by auto + qed + moreover have " f' w = f w " + using \w \ ball p r\ ball_subset_cball subset_iff pp_po \w\p\ + unfolding f'_def by auto + ultimately show "ff' w = ff w" + unfolding ff'_def ff_def by simp + qed + moreover have "cball p e2 \ ball p e1" + using \0 < r\ \r e2_def by auto + ultimately show ?thesis using \e2>0\ by auto + qed + then obtain e2 where e2:"p\pz \ e2>0 \ cball p e2 \ ball p e1 \ cont ff p e2" + by auto + define e4 where "e4 \ if p\pz then e2 else e1" + have "e4>0" using e2 \e1>0\ unfolding e4_def by auto + moreover have "avoid p e4" using e2 \e1>0\ e1_avoid unfolding e4_def avoid_def by auto + moreover have "p\pz \ cont ff p e4" + by (auto simp add: e2 e4_def) + ultimately show ?thesis by auto + qed + then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) + \ (p\pz \ cont ff p (get_e p))" + by metis + define ci where "ci \ \p. contour_integral (circlepath p (get_e p)) ff" + define w where "w \ \p. winding_number g p" + have "contour_integral g ff = (\p\pz. w p * ci p)" unfolding ci_def w_def + proof (rule Cauchy_theorem_singularities[OF \open s\ \connected s\ finite _ \valid_path g\ loop + path_img homo]) + have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto + then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo + by (auto intro!: holomorphic_intros simp add:pz_def) + next + show "\p\s. 0 < get_e p \ (\w\cball p (get_e p). w \ s \ (w \ p \ w \ pz))" + using get_e using avoid_def by blast + qed + also have "... = (\p\pz. c * w p * h p * zorder f p)" + proof (rule sum.cong[of pz pz,simplified]) + fix p assume "p \ pz" + show "w p * ci p = c * w p * h p * (zorder f p)" + proof (cases "p\s") + assume "p \ s" + have "ci p = c * h p * (zorder f p)" unfolding ci_def + apply (rule contour_integral_unique) + using get_e \p\s\ \p\pz\ unfolding cont_def by (metis mult.assoc mult.commute) + thus ?thesis by auto + next + assume "p\s" + then have "w p=0" using homo unfolding w_def by auto + then show ?thesis by auto + qed + qed + also have "... = c*(\p\pz. w p * h p * zorder f p)" + unfolding sum_distrib_left by (simp add:algebra_simps) + finally have "contour_integral g ff = c * (\p\pz. w p * h p * of_int (zorder f p))" . + then show ?thesis unfolding ff_def c_def w_def by simp +qed + +subsection \Rouche's theorem \ + +theorem Rouche_theorem: + fixes f g::"complex \ complex" and s:: "complex set" + defines "fg\(\p. f p + g p)" + defines "zeros_fg\{p. fg p = 0}" and "zeros_f\{p. f p = 0}" + assumes + "open s" and "connected s" and + "finite zeros_fg" and + "finite zeros_f" and + f_holo:"f holomorphic_on s" and + g_holo:"g holomorphic_on s" and + "valid_path \" and + loop:"pathfinish \ = pathstart \" and + path_img:"path_image \ \ s " and + path_less:"\z\path_image \. cmod(f z) > cmod(g z)" and + homo:"\z. (z \ s) \ winding_number \ z = 0" + shows "(\p\zeros_fg. winding_number \ p * zorder fg p) + = (\p\zeros_f. winding_number \ p * zorder f p)" +proof - + have path_fg:"path_image \ \ s - zeros_fg" + proof - + have False when "z\path_image \" and "f z + g z=0" for z + proof - + have "cmod (f z) > cmod (g z)" using \z\path_image \\ path_less by auto + moreover have "f z = - g z" using \f z + g z =0\ by (simp add: eq_neg_iff_add_eq_0) + then have "cmod (f z) = cmod (g z)" by auto + ultimately show False by auto + qed + then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto + qed + have path_f:"path_image \ \ s - zeros_f" + proof - + have False when "z\path_image \" and "f z =0" for z + proof - + have "cmod (g z) < cmod (f z) " using \z\path_image \\ path_less by auto + then have "cmod (g z) < 0" using \f z=0\ by auto + then show False by auto + qed + then show ?thesis unfolding zeros_f_def using path_img by auto + qed + define w where "w \ \p. winding_number \ p" + define c where "c \ 2 * complex_of_real pi * \" + define h where "h \ \p. g p / f p + 1" + obtain spikes + where "finite spikes" and spikes: "\x\{0..1} - spikes. \ differentiable at x" + using \valid_path \\ + by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) + have h_contour:"((\x. deriv h x / h x) has_contour_integral 0) \" + proof - + have outside_img:"0 \ outside (path_image (h o \))" + proof - + have "h p \ ball 1 1" when "p\path_image \" for p + proof - + have "cmod (g p/f p) <1" using path_less[rule_format,OF that] + apply (cases "cmod (f p) = 0") + by (auto simp add: norm_divide) + then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) + qed + then have "path_image (h o \) \ ball 1 1" + by (simp add: image_subset_iff path_image_compose) + moreover have " (0::complex) \ ball 1 1" by (simp add: dist_norm) + ultimately show "?thesis" + using convex_in_outside[of "ball 1 1" 0] outside_mono by blast + qed + have valid_h:"valid_path (h \ \)" + proof (rule valid_path_compose_holomorphic[OF \valid_path \\ _ _ path_f]) + show "h holomorphic_on s - zeros_f" + unfolding h_def using f_holo g_holo + by (auto intro!: holomorphic_intros simp add:zeros_f_def) + next + show "open (s - zeros_f)" using \finite zeros_f\ \open s\ finite_imp_closed + by auto + qed + have "((\z. 1/z) has_contour_integral 0) (h \ \)" + proof - + have "0 \ path_image (h \ \)" using outside_img by (simp add: outside_def) + then have "((\z. 1/z) has_contour_integral c * winding_number (h \ \) 0) (h \ \)" + using has_contour_integral_winding_number[of "h o \" 0,simplified] valid_h + unfolding c_def by auto + moreover have "winding_number (h o \) 0 = 0" + proof - + have "0 \ outside (path_image (h \ \))" using outside_img . + moreover have "path (h o \)" + using valid_h by (simp add: valid_path_imp_path) + moreover have "pathfinish (h o \) = pathstart (h o \)" + by (simp add: loop pathfinish_compose pathstart_compose) + ultimately show ?thesis using winding_number_zero_in_outside by auto + qed + ultimately show ?thesis by auto + qed + moreover have "vector_derivative (h \ \) (at x) = vector_derivative \ (at x) * deriv h (\ x)" + when "x\{0..1} - spikes" for x + proof (rule vector_derivative_chain_at_general) + show "\ differentiable at x" using that \valid_path \\ spikes by auto + next + define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" + define t where "t \ \ x" + have "f t\0" unfolding zeros_f_def t_def + by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) + moreover have "t\s" + using contra_subsetD path_image_def path_fg t_def that by fastforce + ultimately have "(h has_field_derivative der t) (at t)" + unfolding h_def der_def using g_holo f_holo \open s\ + by (auto intro!: holomorphic_derivI derivative_eq_intros) + then show "h field_differentiable at (\ x)" + unfolding t_def field_differentiable_def by blast + qed + then have " ((/) 1 has_contour_integral 0) (h \ \) + = ((\x. deriv h x / h x) has_contour_integral 0) \" + unfolding has_contour_integral + apply (intro has_integral_spike_eq[OF negligible_finite, OF \finite spikes\]) + by auto + ultimately show ?thesis by auto + qed + then have "contour_integral \ (\x. deriv h x / h x) = 0" + using contour_integral_unique by simp + moreover have "contour_integral \ (\x. deriv fg x / fg x) = contour_integral \ (\x. deriv f x / f x) + + contour_integral \ (\p. deriv h p / h p)" + proof - + have "(\p. deriv f p / f p) contour_integrable_on \" + proof (rule contour_integrable_holomorphic_simple[OF _ _ \valid_path \\ path_f]) + show "open (s - zeros_f)" using finite_imp_closed[OF \finite zeros_f\] \open s\ + by auto + then show "(\p. deriv f p / f p) holomorphic_on s - zeros_f" + using f_holo + by (auto intro!: holomorphic_intros simp add:zeros_f_def) + qed + moreover have "(\p. deriv h p / h p) contour_integrable_on \" + using h_contour + by (simp add: has_contour_integral_integrable) + ultimately have "contour_integral \ (\x. deriv f x / f x + deriv h x / h x) = + contour_integral \ (\p. deriv f p / f p) + contour_integral \ (\p. deriv h p / h p)" + using contour_integral_add[of "(\p. deriv f p / f p)" \ "(\p. deriv h p / h p)" ] + by auto + moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" + when "p\ path_image \" for p + proof - + have "fg p\0" and "f p\0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def + by auto + have "h p\0" + proof (rule ccontr) + assume "\ h p \ 0" + then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) + then have "cmod (g p/f p) = 1" by auto + moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] + apply (cases "cmod (f p) = 0") + by (auto simp add: norm_divide) + ultimately show False by auto + qed + have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def + using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \open s\] path_img that + by auto + have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" + proof - + define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" + have "p\s" using path_img that by auto + then have "(h has_field_derivative der p) (at p)" + unfolding h_def der_def using g_holo f_holo \open s\ \f p\0\ + by (auto intro!: derivative_eq_intros holomorphic_derivI) + then show ?thesis unfolding der_def using DERIV_imp_deriv by auto + qed + show ?thesis + apply (simp only:der_fg der_h) + apply (auto simp add:field_simps \h p\0\ \f p\0\ \fg p\0\) + by (auto simp add:field_simps h_def \f p\0\ fg_def) + qed + then have "contour_integral \ (\p. deriv fg p / fg p) + = contour_integral \ (\p. deriv f p / f p + deriv h p / h p)" + by (elim contour_integral_eq) + ultimately show ?thesis by auto + qed + moreover have "contour_integral \ (\x. deriv fg x / fg x) = c * (\p\zeros_fg. w p * zorder fg p)" + unfolding c_def zeros_fg_def w_def + proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo + , of _ "{}" "\_. 1",simplified]) + show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto + show "path_image \ \ s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def . + show " finite {p. fg p = 0}" using \finite zeros_fg\ unfolding zeros_fg_def . + qed + moreover have "contour_integral \ (\x. deriv f x / f x) = c * (\p\zeros_f. w p * zorder f p)" + unfolding c_def zeros_f_def w_def + proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo + , of _ "{}" "\_. 1",simplified]) + show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto + show "path_image \ \ s - {p. f p = 0}" using path_f unfolding zeros_f_def . + show " finite {p. f p = 0}" using \finite zeros_f\ unfolding zeros_f_def . + qed + ultimately have " c* (\p\zeros_fg. w p * (zorder fg p)) = c* (\p\zeros_f. w p * (zorder f p))" + by auto + then show ?thesis unfolding c_def using w_def by auto +qed + + +subsection \Poles and residues of some well-known functions\ + +(* TODO: add more material here for other functions *) +lemma is_pole_Gamma: "is_pole Gamma (-of_nat n)" + unfolding is_pole_def using Gamma_poles . + +lemma Gamme_residue: + "residue Gamma (-of_nat n) = (-1) ^ n / fact n" +proof (rule residue_simple') + show "open (- (\\<^sub>\\<^sub>0 - {-of_nat n}) :: complex set)" + by (intro open_Compl closed_subset_Ints) auto + show "Gamma holomorphic_on (- (\\<^sub>\\<^sub>0 - {-of_nat n}) - {- of_nat n})" + by (rule holomorphic_Gamma) auto + show "(\w. Gamma w * (w - (-of_nat n))) \(-of_nat n)\ (- 1) ^ n / fact n" + using Gamma_residues[of n] by simp +qed auto + +end diff --git a/src/HOL/Complex_Analysis/Great_Picard.thy b/src/HOL/Complex_Analysis/Great_Picard.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/Great_Picard.thy @@ -0,0 +1,1862 @@ +section \The Great Picard Theorem and its Applications\ + +text\Ported from HOL Light (cauchy.ml) by L C Paulson, 2017\ + +theory Great_Picard + imports Conformal_Mappings + +begin + +subsection\Schottky's theorem\ + +lemma Schottky_lemma0: + assumes holf: "f holomorphic_on S" and cons: "contractible S" and "a \ S" + and f: "\z. z \ S \ f z \ 1 \ f z \ -1" + obtains g where "g holomorphic_on S" + "norm(g a) \ 1 + norm(f a) / 3" + "\z. z \ S \ f z = cos(of_real pi * g z)" +proof - + obtain g where holg: "g holomorphic_on S" and g: "norm(g a) \ pi + norm(f a)" + and f_eq_cos: "\z. z \ S \ f z = cos(g z)" + using contractible_imp_holomorphic_arccos_bounded [OF assms] + by blast + show ?thesis + proof + show "(\z. g z / pi) holomorphic_on S" + by (auto intro: holomorphic_intros holg) + have "3 \ pi" + using pi_approx by force + have "3 * norm(g a) \ 3 * (pi + norm(f a))" + using g by auto + also have "... \ pi * 3 + pi * cmod (f a)" + using \3 \ pi\ by (simp add: mult_right_mono algebra_simps) + finally show "cmod (g a / complex_of_real pi) \ 1 + cmod (f a) / 3" + by (simp add: field_simps norm_divide) + show "\z. z \ S \ f z = cos (complex_of_real pi * (g z / complex_of_real pi))" + by (simp add: f_eq_cos) + qed +qed + + +lemma Schottky_lemma1: + fixes n::nat + assumes "0 < n" + shows "0 < n + sqrt(real n ^ 2 - 1)" +proof - + have "(n-1)^2 \ n^2 - 1" + using assms by (simp add: algebra_simps power2_eq_square) + then have "real (n - 1) \ sqrt (real (n\<^sup>2 - 1))" + by (metis of_nat_le_iff of_nat_power real_le_rsqrt) + then have "n-1 \ sqrt(real n ^ 2 - 1)" + by (simp add: Suc_leI assms of_nat_diff) + then show ?thesis + using assms by linarith +qed + + +lemma Schottky_lemma2: + fixes x::real + assumes "0 \ x" + obtains n where "0 < n" "\x - ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi\ < 1/2" +proof - + obtain n0::nat where "0 < n0" "ln(n0 + sqrt(real n0 ^ 2 - 1)) / pi \ x" + proof + show "ln(real 1 + sqrt(real 1 ^ 2 - 1))/pi \ x" + by (auto simp: assms) + qed auto + moreover + obtain M::nat where "\n. \0 < n; ln(n + sqrt(real n ^ 2 - 1)) / pi \ x\ \ n \ M" + proof + fix n::nat + assume "0 < n" "ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi \ x" + then have "ln (n + sqrt ((real n)\<^sup>2 - 1)) \ x * pi" + by (simp add: field_split_simps) + then have *: "exp (ln (n + sqrt ((real n)\<^sup>2 - 1))) \ exp (x * pi)" + by blast + have 0: "0 \ sqrt ((real n)\<^sup>2 - 1)" + using \0 < n\ by auto + have "n + sqrt ((real n)\<^sup>2 - 1) = exp (ln (n + sqrt ((real n)\<^sup>2 - 1)))" + by (simp add: Suc_leI \0 < n\ add_pos_nonneg real_of_nat_ge_one_iff) + also have "... \ exp (x * pi)" + using "*" by blast + finally have "real n \ exp (x * pi)" + using 0 by linarith + then show "n \ nat (ceiling (exp(x * pi)))" + by linarith + qed + ultimately obtain n where + "0 < n" and le_x: "ln(n + sqrt(real n ^ 2 - 1)) / pi \ x" + and le_n: "\k. \0 < k; ln(k + sqrt(real k ^ 2 - 1)) / pi \ x\ \ k \ n" + using bounded_Max_nat [of "\n. 0 ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi \ x"] by metis + define a where "a \ ln(n + sqrt(real n ^ 2 - 1)) / pi" + define b where "b \ ln (1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / pi" + have le_xa: "a \ x" + and le_na: "\k. \0 < k; ln(k + sqrt(real k ^ 2 - 1)) / pi \ x\ \ k \ n" + using le_x le_n by (auto simp: a_def) + moreover have "x < b" + using le_n [of "Suc n"] by (force simp: b_def) + moreover have "b - a < 1" + proof - + have "ln (1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) - ln (real n + sqrt ((real n)\<^sup>2 - 1)) = + ln ((1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / (real n + sqrt ((real n)\<^sup>2 - 1)))" + by (simp add: \0 < n\ Schottky_lemma1 add_pos_nonneg ln_div [symmetric]) + also have "... \ 3" + proof (cases "n = 1") + case True + have "sqrt 3 \ 2" + by (simp add: real_le_lsqrt) + then have "(2 + sqrt 3) \ 4" + by simp + also have "... \ exp 3" + using exp_ge_add_one_self [of "3::real"] by simp + finally have "ln (2 + sqrt 3) \ 3" + by (metis add_nonneg_nonneg add_pos_nonneg dbl_def dbl_inc_def dbl_inc_simps(3) + dbl_simps(3) exp_gt_zero ln_exp ln_le_cancel_iff real_sqrt_ge_0_iff zero_le_one zero_less_one) + then show ?thesis + by (simp add: True) + next + case False with \0 < n\ have "1 < n" "2 \ n" + by linarith+ + then have 1: "1 \ real n * real n" + by (metis less_imp_le_nat one_le_power power2_eq_square real_of_nat_ge_one_iff) + have *: "4 + (m+2) * 2 \ (m+2) * ((m+2) * 3)" for m::nat + by simp + have "4 + n * 2 \ n * (n * 3)" + using * [of "n-2"] \2 \ n\ + by (metis le_add_diff_inverse2) + then have **: "4 + real n * 2 \ real n * (real n * 3)" + by (metis (mono_tags, hide_lams) of_nat_le_iff of_nat_add of_nat_mult of_nat_numeral) + have "sqrt ((1 + real n)\<^sup>2 - 1) \ 2 * sqrt ((real n)\<^sup>2 - 1)" + by (auto simp: real_le_lsqrt power2_eq_square algebra_simps 1 **) + then + have "((1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / (real n + sqrt ((real n)\<^sup>2 - 1))) \ 2" + using Schottky_lemma1 \0 < n\ by (simp add: field_split_simps) + then have "ln ((1 + real n + sqrt ((1 + real n)\<^sup>2 - 1)) / (real n + sqrt ((real n)\<^sup>2 - 1))) \ ln 2" + apply (subst ln_le_cancel_iff) + using Schottky_lemma1 \0 < n\ by auto (force simp: field_split_simps) + also have "... \ 3" + using ln_add_one_self_le_self [of 1] by auto + finally show ?thesis . + qed + also have "... < pi" + using pi_approx by simp + finally show ?thesis + by (simp add: a_def b_def field_split_simps) + qed + ultimately have "\x - a\ < 1/2 \ \x - b\ < 1/2" + by (auto simp: abs_if) + then show thesis + proof + assume "\x - a\ < 1 / 2" + then show ?thesis + by (rule_tac n=n in that) (auto simp: a_def \0 < n\) + next + assume "\x - b\ < 1 / 2" + then show ?thesis + by (rule_tac n="Suc n" in that) (auto simp: b_def \0 < n\) + qed +qed + + +lemma Schottky_lemma3: + fixes z::complex + assumes "z \ (\m \ Ints. \n \ {0<..}. {Complex m (ln(n + sqrt(real n ^ 2 - 1)) / pi)}) + \ (\m \ Ints. \n \ {0<..}. {Complex m (-ln(n + sqrt(real n ^ 2 - 1)) / pi)})" + shows "cos(pi * cos(pi * z)) = 1 \ cos(pi * cos(pi * z)) = -1" +proof - + have sqrt2 [simp]: "complex_of_real (sqrt x) * complex_of_real (sqrt x) = x" if "x \ 0" for x::real + by (metis abs_of_nonneg of_real_mult real_sqrt_mult_self that) + have 1: "\k. exp (\ * (of_int m * complex_of_real pi) - + (ln (real n + sqrt ((real n)\<^sup>2 - 1)))) + + inverse + (exp (\ * (of_int m * complex_of_real pi) - + (ln (real n + sqrt ((real n)\<^sup>2 - 1))))) = of_int k * 2" + if "n > 0" for m n + proof - + have eeq: "e \ 0 \ e + inverse e = n*2 \ inverse e^2 - 2 * n*inverse e + 1 = 0" for n e::complex + by (auto simp: field_simps power2_eq_square) + have [simp]: "1 \ real n * real n" + by (metis One_nat_def add.commute nat_less_real_le of_nat_1 of_nat_Suc one_le_power power2_eq_square that) + show ?thesis + apply (simp add: eeq) + using Schottky_lemma1 [OF that] + apply (auto simp: eeq exp_diff exp_Euler exp_of_real algebra_simps sin_int_times_real cos_int_times_real) + apply (rule_tac x="int n" in exI) + apply (auto simp: power2_eq_square algebra_simps) + apply (rule_tac x="- int n" in exI) + apply (auto simp: power2_eq_square algebra_simps) + done + qed + have 2: "\k. exp (\ * (of_int m * complex_of_real pi) + + (ln (real n + sqrt ((real n)\<^sup>2 - 1)))) + + inverse + (exp (\ * (of_int m * complex_of_real pi) + + (ln (real n + sqrt ((real n)\<^sup>2 - 1))))) = of_int k * 2" + if "n > 0" for m n + proof - + have eeq: "e \ 0 \ e + inverse e = n*2 \ e^2 - 2 * n*e + 1 = 0" for n e::complex + by (auto simp: field_simps power2_eq_square) + have [simp]: "1 \ real n * real n" + by (metis One_nat_def add.commute nat_less_real_le of_nat_1 of_nat_Suc one_le_power power2_eq_square that) + show ?thesis + apply (simp add: eeq) + using Schottky_lemma1 [OF that] + apply (auto simp: exp_add exp_Euler exp_of_real algebra_simps sin_int_times_real cos_int_times_real) + apply (rule_tac x="int n" in exI) + apply (auto simp: power2_eq_square algebra_simps) + apply (rule_tac x="- int n" in exI) + apply (auto simp: power2_eq_square algebra_simps) + done + qed + have "\x. cos (complex_of_real pi * z) = of_int x" + using assms + apply safe + apply (auto simp: Ints_def cos_exp_eq exp_minus Complex_eq) + apply (auto simp: algebra_simps dest: 1 2) + done + then have "sin(pi * cos(pi * z)) ^ 2 = 0" + by (simp add: Complex_Transcendental.sin_eq_0) + then have "1 - cos(pi * cos(pi * z)) ^ 2 = 0" + by (simp add: sin_squared_eq) + then show ?thesis + using power2_eq_1_iff by auto +qed + + +theorem Schottky: + assumes holf: "f holomorphic_on cball 0 1" + and nof0: "norm(f 0) \ r" + and not01: "\z. z \ cball 0 1 \ \(f z = 0 \ f z = 1)" + and "0 < t" "t < 1" "norm z \ t" + shows "norm(f z) \ exp(pi * exp(pi * (2 + 2 * r + 12 * t / (1 - t))))" +proof - + obtain h where holf: "h holomorphic_on cball 0 1" + and nh0: "norm (h 0) \ 1 + norm(2 * f 0 - 1) / 3" + and h: "\z. z \ cball 0 1 \ 2 * f z - 1 = cos(of_real pi * h z)" + proof (rule Schottky_lemma0 [of "\z. 2 * f z - 1" "cball 0 1" 0]) + show "(\z. 2 * f z - 1) holomorphic_on cball 0 1" + by (intro holomorphic_intros holf) + show "contractible (cball (0::complex) 1)" + by (auto simp: convex_imp_contractible) + show "\z. z \ cball 0 1 \ 2 * f z - 1 \ 1 \ 2 * f z - 1 \ - 1" + using not01 by force + qed auto + obtain g where holg: "g holomorphic_on cball 0 1" + and ng0: "norm(g 0) \ 1 + norm(h 0) / 3" + and g: "\z. z \ cball 0 1 \ h z = cos(of_real pi * g z)" + proof (rule Schottky_lemma0 [OF holf convex_imp_contractible, of 0]) + show "\z. z \ cball 0 1 \ h z \ 1 \ h z \ - 1" + using h not01 by fastforce+ + qed auto + have g0_2_f0: "norm(g 0) \ 2 + norm(f 0)" + proof - + have "cmod (2 * f 0 - 1) \ cmod (2 * f 0) + 1" + by (metis norm_one norm_triangle_ineq4) + also have "... \ 2 + cmod (f 0) * 3" + by simp + finally have "1 + norm(2 * f 0 - 1) / 3 \ (2 + norm(f 0) - 1) * 3" + apply (simp add: field_split_simps) + using norm_ge_zero [of "f 0 * 2 - 1"] + by linarith + with nh0 have "norm(h 0) \ (2 + norm(f 0) - 1) * 3" + by linarith + then have "1 + norm(h 0) / 3 \ 2 + norm(f 0)" + by simp + with ng0 show ?thesis + by auto + qed + have "z \ ball 0 1" + using assms by auto + have norm_g_12: "norm(g z - g 0) \ (12 * t) / (1 - t)" + proof - + obtain g' where g': "\x. x \ cball 0 1 \ (g has_field_derivative g' x) (at x within cball 0 1)" + using holg [unfolded holomorphic_on_def field_differentiable_def] by metis + have int_g': "(g' has_contour_integral g z - g 0) (linepath 0 z)" + using contour_integral_primitive [OF g' valid_path_linepath, of 0 z] + using \z \ ball 0 1\ segment_bound1 by fastforce + have "cmod (g' w) \ 12 / (1 - t)" if "w \ closed_segment 0 z" for w + proof - + have w: "w \ ball 0 1" + using segment_bound [OF that] \z \ ball 0 1\ by simp + have ttt: "\z. z \ frontier (cball 0 1) \ 1 - t \ dist w z" + using \norm z \ t\ segment_bound1 [OF \w \ closed_segment 0 z\] + apply (simp add: dist_complex_def) + by (metis diff_left_mono dist_commute dist_complex_def norm_triangle_ineq2 order_trans) + have *: "\\b. (\w \ T \ U. w \ ball b 1); \x. x \ D \ g x \ T \ U\ \ \b. ball b 1 \ g ` D" for T U D + by force + have "\b. ball b 1 \ g ` cball 0 1" + proof (rule *) + show "(\w \ (\m \ Ints. \n \ {0<..}. {Complex m (ln(n + sqrt(real n ^ 2 - 1)) / pi)}) \ + (\m \ Ints. \n \ {0<..}. {Complex m (-ln(n + sqrt(real n ^ 2 - 1)) / pi)}). w \ ball b 1)" for b + proof - + obtain m where m: "m \ \" "\Re b - m\ \ 1/2" + by (metis Ints_of_int abs_minus_commute of_int_round_abs_le) + show ?thesis + proof (cases "0::real" "Im b" rule: le_cases) + case le + then obtain n where "0 < n" and n: "\Im b - ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi\ < 1/2" + using Schottky_lemma2 [of "Im b"] by blast + have "dist b (Complex m (Im b)) \ 1/2" + by (metis cancel_comm_monoid_add_class.diff_cancel cmod_eq_Re complex.sel(1) complex.sel(2) dist_norm m(2) minus_complex.code) + moreover + have "dist (Complex m (Im b)) (Complex m (ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi)) < 1/2" + using n by (simp add: complex_norm cmod_eq_Re complex_diff dist_norm del: Complex_eq) + ultimately have "dist b (Complex m (ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)) < 1" + by (simp add: dist_triangle_lt [of b "Complex m (Im b)"] dist_commute) + with le m \0 < n\ show ?thesis + apply (rule_tac x = "Complex m (ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)" in bexI) + apply (simp_all del: Complex_eq greaterThan_0) + by blast + next + case ge + then obtain n where "0 < n" and n: "\- Im b - ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi\ < 1/2" + using Schottky_lemma2 [of "- Im b"] by auto + have "dist b (Complex m (Im b)) \ 1/2" + by (metis cancel_comm_monoid_add_class.diff_cancel cmod_eq_Re complex.sel(1) complex.sel(2) dist_norm m(2) minus_complex.code) + moreover + have "dist (Complex m (- ln (n + sqrt ((real n)\<^sup>2 - 1)) / pi)) (Complex m (Im b)) < 1/2" + using n + apply (simp add: complex_norm cmod_eq_Re complex_diff dist_norm del: Complex_eq) + by (metis add.commute add_uminus_conv_diff) + ultimately have "dist b (Complex m (- ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)) < 1" + by (simp add: dist_triangle_lt [of b "Complex m (Im b)"] dist_commute) + with ge m \0 < n\ show ?thesis + apply (rule_tac x = "Complex m (- ln (real n + sqrt ((real n)\<^sup>2 - 1)) / pi)" in bexI) + apply (simp_all del: Complex_eq greaterThan_0) + by blast + qed + qed + show "g v \ (\m \ Ints. \n \ {0<..}. {Complex m (ln(n + sqrt(real n ^ 2 - 1)) / pi)}) \ + (\m \ Ints. \n \ {0<..}. {Complex m (-ln(n + sqrt(real n ^ 2 - 1)) / pi)})" + if "v \ cball 0 1" for v + using not01 [OF that] + by (force simp: g [OF that, symmetric] h [OF that, symmetric] dest!: Schottky_lemma3 [of "g v"]) + qed + then have 12: "(1 - t) * cmod (deriv g w) / 12 < 1" + using Bloch_general [OF holg _ ttt, of 1] w by force + have "g field_differentiable at w within cball 0 1" + using holg w by (simp add: holomorphic_on_def) + then have "g field_differentiable at w within ball 0 1" + using ball_subset_cball field_differentiable_within_subset by blast + with w have der_gw: "(g has_field_derivative deriv g w) (at w)" + by (simp add: field_differentiable_within_open [of _ "ball 0 1"] field_differentiable_derivI) + with DERIV_unique [OF der_gw] g' w have "deriv g w = g' w" + by (metis open_ball at_within_open ball_subset_cball has_field_derivative_subset subsetCE) + then show "cmod (g' w) \ 12 / (1 - t)" + using g' 12 \t < 1\ by (simp add: field_simps) + qed + then have "cmod (g z - g 0) \ 12 / (1 - t) * cmod z" + using has_contour_integral_bound_linepath [OF int_g', of "12/(1 - t)"] assms + by simp + with \cmod z \ t\ \t < 1\ show ?thesis + by (simp add: field_split_simps) + qed + have fz: "f z = (1 + cos(of_real pi * h z)) / 2" + using h \z \ ball 0 1\ by (auto simp: field_simps) + have "cmod (f z) \ exp (cmod (complex_of_real pi * h z))" + by (simp add: fz mult.commute norm_cos_plus1_le) + also have "... \ exp (pi * exp (pi * (2 + 2 * r + 12 * t / (1 - t))))" + proof (simp add: norm_mult) + have "cmod (g z - g 0) \ 12 * t / (1 - t)" + using norm_g_12 \t < 1\ by (simp add: norm_mult) + then have "cmod (g z) - cmod (g 0) \ 12 * t / (1 - t)" + using norm_triangle_ineq2 order_trans by blast + then have *: "cmod (g z) \ 2 + 2 * r + 12 * t / (1 - t)" + using g0_2_f0 norm_ge_zero [of "f 0"] nof0 + by linarith + have "cmod (h z) \ exp (cmod (complex_of_real pi * g z))" + using \z \ ball 0 1\ by (simp add: g norm_cos_le) + also have "... \ exp (pi * (2 + 2 * r + 12 * t / (1 - t)))" + using \t < 1\ nof0 * by (simp add: norm_mult) + finally show "cmod (h z) \ exp (pi * (2 + 2 * r + 12 * t / (1 - t)))" . + qed + finally show ?thesis . +qed + + +subsection\The Little Picard Theorem\ + +theorem Landau_Picard: + obtains R + where "\z. 0 < R z" + "\f. \f holomorphic_on cball 0 (R(f 0)); + \z. norm z \ R(f 0) \ f z \ 0 \ f z \ 1\ \ norm(deriv f 0) < 1" +proof - + define R where "R \ \z. 3 * exp(pi * exp(pi*(2 + 2 * cmod z + 12)))" + show ?thesis + proof + show Rpos: "\z. 0 < R z" + by (auto simp: R_def) + show "norm(deriv f 0) < 1" + if holf: "f holomorphic_on cball 0 (R(f 0))" + and Rf: "\z. norm z \ R(f 0) \ f z \ 0 \ f z \ 1" for f + proof - + let ?r = "R(f 0)" + define g where "g \ f \ (\z. of_real ?r * z)" + have "0 < ?r" + using Rpos by blast + have holg: "g holomorphic_on cball 0 1" + unfolding g_def + apply (intro holomorphic_intros holomorphic_on_compose holomorphic_on_subset [OF holf]) + using Rpos by (auto simp: less_imp_le norm_mult) + have *: "norm(g z) \ exp(pi * exp(pi * (2 + 2 * norm (f 0) + 12 * t / (1 - t))))" + if "0 < t" "t < 1" "norm z \ t" for t z + proof (rule Schottky [OF holg]) + show "cmod (g 0) \ cmod (f 0)" + by (simp add: g_def) + show "\z. z \ cball 0 1 \ \ (g z = 0 \ g z = 1)" + using Rpos by (simp add: g_def less_imp_le norm_mult Rf) + qed (auto simp: that) + have C1: "g holomorphic_on ball 0 (1 / 2)" + by (rule holomorphic_on_subset [OF holg]) auto + have C2: "continuous_on (cball 0 (1 / 2)) g" + by (meson cball_divide_subset_numeral holg holomorphic_on_imp_continuous_on holomorphic_on_subset) + have C3: "cmod (g z) \ R (f 0) / 3" if "cmod (0 - z) = 1/2" for z + proof - + have "norm(g z) \ exp(pi * exp(pi * (2 + 2 * norm (f 0) + 12)))" + using * [of "1/2"] that by simp + also have "... = ?r / 3" + by (simp add: R_def) + finally show ?thesis . + qed + then have cmod_g'_le: "cmod (deriv g 0) * 3 \ R (f 0) * 2" + using Cauchy_inequality [OF C1 C2 _ C3, of 1] by simp + have holf': "f holomorphic_on ball 0 (R(f 0))" + by (rule holomorphic_on_subset [OF holf]) auto + then have fd0: "f field_differentiable at 0" + by (rule holomorphic_on_imp_differentiable_at [OF _ open_ball]) + (auto simp: Rpos [of "f 0"]) + have g_eq: "deriv g 0 = of_real ?r * deriv f 0" + apply (rule DERIV_imp_deriv) + apply (simp add: g_def) + by (metis DERIV_chain DERIV_cmult_Id fd0 field_differentiable_derivI mult.commute mult_zero_right) + show ?thesis + using cmod_g'_le Rpos [of "f 0"] by (simp add: g_eq norm_mult) + qed + qed +qed + +lemma little_Picard_01: + assumes holf: "f holomorphic_on UNIV" and f01: "\z. f z \ 0 \ f z \ 1" + obtains c where "f = (\x. c)" +proof - + obtain R + where Rpos: "\z. 0 < R z" + and R: "\h. \h holomorphic_on cball 0 (R(h 0)); + \z. norm z \ R(h 0) \ h z \ 0 \ h z \ 1\ \ norm(deriv h 0) < 1" + using Landau_Picard by metis + have contf: "continuous_on UNIV f" + by (simp add: holf holomorphic_on_imp_continuous_on) + show ?thesis + proof (cases "\x. deriv f x = 0") + case True + obtain c where "\x. f(x) = c" + apply (rule DERIV_zero_connected_constant [OF connected_UNIV open_UNIV finite.emptyI contf]) + apply (metis True DiffE holf holomorphic_derivI open_UNIV, auto) + done + then show ?thesis + using that by auto + next + case False + then obtain w where w: "deriv f w \ 0" by auto + define fw where "fw \ (f \ (\z. w + z / deriv f w))" + have norm_let1: "norm(deriv fw 0) < 1" + proof (rule R) + show "fw holomorphic_on cball 0 (R (fw 0))" + unfolding fw_def + by (intro holomorphic_intros holomorphic_on_compose w holomorphic_on_subset [OF holf] subset_UNIV) + show "fw z \ 0 \ fw z \ 1" if "cmod z \ R (fw 0)" for z + using f01 by (simp add: fw_def) + qed + have "(fw has_field_derivative deriv f w * inverse (deriv f w)) (at 0)" + apply (simp add: fw_def) + apply (rule DERIV_chain) + using holf holomorphic_derivI apply force + apply (intro derivative_eq_intros w) + apply (auto simp: field_simps) + done + then show ?thesis + using norm_let1 w by (simp add: DERIV_imp_deriv) + qed +qed + + +theorem little_Picard: + assumes holf: "f holomorphic_on UNIV" + and "a \ b" "range f \ {a,b} = {}" + obtains c where "f = (\x. c)" +proof - + let ?g = "\x. 1/(b - a)*(f x - b) + 1" + obtain c where "?g = (\x. c)" + proof (rule little_Picard_01) + show "?g holomorphic_on UNIV" + by (intro holomorphic_intros holf) + show "\z. ?g z \ 0 \ ?g z \ 1" + using assms by (auto simp: field_simps) + qed auto + then have "?g x = c" for x + by meson + then have "f x = c * (b-a) + a" for x + using assms by (auto simp: field_simps) + then show ?thesis + using that by blast +qed + + +text\A couple of little applications of Little Picard\ + +lemma holomorphic_periodic_fixpoint: + assumes holf: "f holomorphic_on UNIV" + and "p \ 0" and per: "\z. f(z + p) = f z" + obtains x where "f x = x" +proof - + have False if non: "\x. f x \ x" + proof - + obtain c where "(\z. f z - z) = (\z. c)" + proof (rule little_Picard) + show "(\z. f z - z) holomorphic_on UNIV" + by (simp add: holf holomorphic_on_diff) + show "range (\z. f z - z) \ {p,0} = {}" + using assms non by auto (metis add.commute diff_eq_eq) + qed (auto simp: assms) + with per show False + by (metis add.commute add_cancel_left_left \p \ 0\ diff_add_cancel) + qed + then show ?thesis + using that by blast +qed + + +lemma holomorphic_involution_point: + assumes holfU: "f holomorphic_on UNIV" and non: "\a. f \ (\x. a + x)" + obtains x where "f(f x) = x" +proof - + { assume non_ff [simp]: "\x. f(f x) \ x" + then have non_fp [simp]: "f z \ z" for z + by metis + have holf: "f holomorphic_on X" for X + using assms holomorphic_on_subset by blast + obtain c where c: "(\x. (f(f x) - x)/(f x - x)) = (\x. c)" + proof (rule little_Picard_01) + show "(\x. (f(f x) - x)/(f x - x)) holomorphic_on UNIV" + apply (intro holomorphic_intros holf holomorphic_on_compose [unfolded o_def, OF holf]) + using non_fp by auto + qed auto + then obtain "c \ 0" "c \ 1" + by (metis (no_types, hide_lams) non_ff diff_zero divide_eq_0_iff right_inverse_eq) + have eq: "f(f x) - c * f x = x*(1 - c)" for x + using fun_cong [OF c, of x] by (simp add: field_simps) + have df_times_dff: "deriv f z * (deriv f (f z) - c) = 1 * (1 - c)" for z + proof (rule DERIV_unique) + show "((\x. f (f x) - c * f x) has_field_derivative + deriv f z * (deriv f (f z) - c)) (at z)" + apply (intro derivative_eq_intros) + apply (rule DERIV_chain [unfolded o_def, of f]) + apply (auto simp: algebra_simps intro!: holomorphic_derivI [OF holfU]) + done + show "((\x. f (f x) - c * f x) has_field_derivative 1 * (1 - c)) (at z)" + by (simp add: eq mult_commute_abs) + qed + { fix z::complex + obtain k where k: "deriv f \ f = (\x. k)" + proof (rule little_Picard) + show "(deriv f \ f) holomorphic_on UNIV" + by (meson holfU holomorphic_deriv holomorphic_on_compose holomorphic_on_subset open_UNIV subset_UNIV) + obtain "deriv f (f x) \ 0" "deriv f (f x) \ c" for x + using df_times_dff \c \ 1\ eq_iff_diff_eq_0 + by (metis lambda_one mult_zero_left mult_zero_right) + then show "range (deriv f \ f) \ {0,c} = {}" + by force + qed (use \c \ 0\ in auto) + have "\ f constant_on UNIV" + by (meson UNIV_I non_ff constant_on_def) + with holf open_mapping_thm have "open(range f)" + by blast + obtain l where l: "\x. f x - k * x = l" + proof (rule DERIV_zero_connected_constant [of UNIV "{}" "\x. f x - k * x"], simp_all) + have "deriv f w - k = 0" for w + proof (rule analytic_continuation [OF _ open_UNIV connected_UNIV subset_UNIV, of "\z. deriv f z - k" "f z" "range f" w]) + show "(\z. deriv f z - k) holomorphic_on UNIV" + by (intro holomorphic_intros holf open_UNIV) + show "f z islimpt range f" + by (metis (no_types, lifting) IntI UNIV_I \open (range f)\ image_eqI inf.absorb_iff2 inf_aci(1) islimpt_UNIV islimpt_eq_acc_point open_Int top_greatest) + show "\z. z \ range f \ deriv f z - k = 0" + by (metis comp_def diff_self image_iff k) + qed auto + moreover + have "((\x. f x - k * x) has_field_derivative deriv f x - k) (at x)" for x + by (metis DERIV_cmult_Id Deriv.field_differentiable_diff UNIV_I field_differentiable_derivI holf holomorphic_on_def) + ultimately + show "\x. ((\x. f x - k * x) has_field_derivative 0) (at x)" + by auto + show "continuous_on UNIV (\x. f x - k * x)" + by (simp add: continuous_on_diff holf holomorphic_on_imp_continuous_on) + qed (auto simp: connected_UNIV) + have False + proof (cases "k=1") + case True + then have "\x. k * x + l \ a + x" for a + using l non [of a] ext [of f "(+) a"] + by (metis add.commute diff_eq_eq) + with True show ?thesis by auto + next + case False + have "\x. (1 - k) * x \ f 0" + using l [of 0] apply (simp add: algebra_simps) + by (metis diff_add_cancel l mult.commute non_fp) + then show False + by (metis False eq_iff_diff_eq_0 mult.commute nonzero_mult_div_cancel_right times_divide_eq_right) + qed + } + } + then show thesis + using that by blast +qed + + +subsection\The Arzelà--Ascoli theorem\ + +lemma subsequence_diagonalization_lemma: + fixes P :: "nat \ (nat \ 'a) \ bool" + assumes sub: "\i r. \k. strict_mono (k :: nat \ nat) \ P i (r \ k)" + and P_P: "\i r::nat \ 'a. \k1 k2 N. + \P i (r \ k1); \j. N \ j \ \j'. j \ j' \ k2 j = k1 j'\ \ P i (r \ k2)" + obtains k where "strict_mono (k :: nat \ nat)" "\i. P i (r \ k)" +proof - + obtain kk where "\i r. strict_mono (kk i r :: nat \ nat) \ P i (r \ (kk i r))" + using sub by metis + then have sub_kk: "\i r. strict_mono (kk i r)" and P_kk: "\i r. P i (r \ (kk i r))" + by auto + define rr where "rr \ rec_nat (kk 0 r) (\n x. x \ kk (Suc n) (r \ x))" + then have [simp]: "rr 0 = kk 0 r" "\n. rr(Suc n) = rr n \ kk (Suc n) (r \ rr n)" + by auto + show thesis + proof + have sub_rr: "strict_mono (rr i)" for i + using sub_kk by (induction i) (auto simp: strict_mono_def o_def) + have P_rr: "P i (r \ rr i)" for i + using P_kk by (induction i) (auto simp: o_def) + have "i \ i+d \ rr i n \ rr (i+d) n" for d i n + proof (induction d) + case 0 then show ?case + by simp + next + case (Suc d) then show ?case + apply simp + using seq_suble [OF sub_kk] order_trans strict_mono_less_eq [OF sub_rr] by blast + qed + then have "\i j n. i \ j \ rr i n \ rr j n" + by (metis le_iff_add) + show "strict_mono (\n. rr n n)" + apply (simp add: strict_mono_Suc_iff) + by (meson lessI less_le_trans seq_suble strict_monoD sub_kk sub_rr) + have "\j. i \ j \ rr (n+d) i = rr n j" for d n i + apply (induction d arbitrary: i, auto) + by (meson order_trans seq_suble sub_kk) + then have "\m n i. n \ m \ \j. i \ j \ rr m i = rr n j" + by (metis le_iff_add) + then show "P i (r \ (\n. rr n n))" for i + by (meson P_rr P_P) + qed +qed + +lemma function_convergent_subsequence: + fixes f :: "[nat,'a] \ 'b::{real_normed_vector,heine_borel}" + assumes "countable S" and M: "\n::nat. \x. x \ S \ norm(f n x) \ M" + obtains k where "strict_mono (k::nat\nat)" "\x. x \ S \ \l. (\n. f (k n) x) \ l" +proof (cases "S = {}") + case True + then show ?thesis + using strict_mono_id that by fastforce +next + case False + with \countable S\ obtain \ :: "nat \ 'a" where \: "S = range \" + using uncountable_def by blast + obtain k where "strict_mono k" and k: "\i. \l. (\n. (f \ k) n (\ i)) \ l" + proof (rule subsequence_diagonalization_lemma + [of "\i r. \l. ((\n. (f \ r) n (\ i)) \ l) sequentially" id]) + show "\k::nat\nat. strict_mono k \ (\l. (\n. (f \ (r \ k)) n (\ i)) \ l)" for i r + proof - + have "f (r n) (\ i) \ cball 0 M" for n + by (simp add: \ M) + then show ?thesis + using compact_def [of "cball (0::'b) M"] apply simp + apply (drule_tac x="(\n. f (r n) (\ i))" in spec) + apply (force simp: o_def) + done + qed + show "\i r k1 k2 N. + \\l. (\n. (f \ (r \ k1)) n (\ i)) \ l; \j. N \ j \ \j'\j. k2 j = k1 j'\ + \ \l. (\n. (f \ (r \ k2)) n (\ i)) \ l" + apply (simp add: lim_sequentially) + apply (erule ex_forward all_forward imp_forward)+ + apply auto + by (metis (no_types, hide_lams) le_cases order_trans) + qed auto + with \ that show ?thesis + by force +qed + + +theorem Arzela_Ascoli: + fixes \ :: "[nat,'a::euclidean_space] \ 'b::{real_normed_vector,heine_borel}" + assumes "compact S" + and M: "\n x. x \ S \ norm(\ n x) \ M" + and equicont: + "\x e. \x \ S; 0 < e\ + \ \d. 0 < d \ (\n y. y \ S \ norm(x - y) < d \ norm(\ n x - \ n y) < e)" + obtains g k where "continuous_on S g" "strict_mono (k :: nat \ nat)" + "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ norm(\(k n) x - g x) < e" +proof - + have UEQ: "\e. 0 < e \ \d. 0 < d \ (\n. \x \ S. \x' \ S. dist x' x < d \ dist (\ n x') (\ n x) < e)" + apply (rule compact_uniformly_equicontinuous [OF \compact S\, of "range \"]) + using equicont by (force simp: dist_commute dist_norm)+ + have "continuous_on S g" + if "\e. 0 < e \ \N. \n x. n \ N \ x \ S \ norm(\(r n) x - g x) < e" + for g:: "'a \ 'b" and r :: "nat \ nat" + proof (rule uniform_limit_theorem [of _ "\ \ r"]) + show "\\<^sub>F n in sequentially. continuous_on S ((\ \ r) n)" + apply (simp add: eventually_sequentially) + apply (rule_tac x=0 in exI) + using UEQ apply (force simp: continuous_on_iff) + done + show "uniform_limit S (\ \ r) g sequentially" + apply (simp add: uniform_limit_iff eventually_sequentially) + by (metis dist_norm that) + qed auto + moreover + obtain R where "countable R" "R \ S" and SR: "S \ closure R" + by (metis separable that) + obtain k where "strict_mono k" and k: "\x. x \ R \ \l. (\n. \ (k n) x) \ l" + apply (rule function_convergent_subsequence [OF \countable R\ M]) + using \R \ S\ apply force+ + done + then have Cauchy: "Cauchy ((\n. \ (k n) x))" if "x \ R" for x + using convergent_eq_Cauchy that by blast + have "\N. \m n x. N \ m \ N \ n \ x \ S \ dist ((\ \ k) m x) ((\ \ k) n x) < e" + if "0 < e" for e + proof - + obtain d where "0 < d" + and d: "\n. \x \ S. \x' \ S. dist x' x < d \ dist (\ n x') (\ n x) < e/3" + by (metis UEQ \0 < e\ divide_pos_pos zero_less_numeral) + obtain T where "T \ R" and "finite T" and T: "S \ (\c\T. ball c d)" + proof (rule compactE_image [OF \compact S\, of R "(\x. ball x d)"]) + have "closure R \ (\c\R. ball c d)" + apply clarsimp + using \0 < d\ closure_approachable by blast + with SR show "S \ (\c\R. ball c d)" + by auto + qed auto + have "\M. \m\M. \n\M. dist (\ (k m) x) (\ (k n) x) < e/3" if "x \ R" for x + using Cauchy \0 < e\ that unfolding Cauchy_def + by (metis less_divide_eq_numeral1(1) mult_zero_left) + then obtain MF where MF: "\x m n. \x \ R; m \ MF x; n \ MF x\ \ norm (\ (k m) x - \ (k n) x) < e/3" + using dist_norm by metis + have "dist ((\ \ k) m x) ((\ \ k) n x) < e" + if m: "Max (MF ` T) \ m" and n: "Max (MF ` T) \ n" "x \ S" for m n x + proof - + obtain t where "t \ T" and t: "x \ ball t d" + using \x \ S\ T by auto + have "norm(\ (k m) t - \ (k m) x) < e / 3" + by (metis \R \ S\ \T \ R\ \t \ T\ d dist_norm mem_ball subset_iff t \x \ S\) + moreover + have "norm(\ (k n) t - \ (k n) x) < e / 3" + by (metis \R \ S\ \T \ R\ \t \ T\ subsetD d dist_norm mem_ball t \x \ S\) + moreover + have "norm(\ (k m) t - \ (k n) t) < e / 3" + proof (rule MF) + show "t \ R" + using \T \ R\ \t \ T\ by blast + show "MF t \ m" "MF t \ n" + by (meson Max_ge \finite T\ \t \ T\ finite_imageI imageI le_trans m n)+ + qed + ultimately + show ?thesis + unfolding dist_norm [symmetric] o_def + by (metis dist_triangle_third dist_commute) + qed + then show ?thesis + by force + qed + then have "\g. \e>0. \N. \n\N. \x \ S. norm(\(k n) x - g x) < e" + using uniformly_convergent_eq_cauchy [of "\x. x \ S" "\ \ k"] + apply (simp add: o_def dist_norm) + by meson + ultimately show thesis + by (metis that \strict_mono k\) +qed + + + +subsubsection\<^marker>\tag important\\Montel's theorem\ + +text\a sequence of holomorphic functions uniformly bounded +on compact subsets of an open set S has a subsequence that converges to a +holomorphic function, and converges \emph{uniformly} on compact subsets of S.\ + + +theorem Montel: + fixes \ :: "[nat,complex] \ complex" + assumes "open S" + and \: "\h. h \ \ \ h holomorphic_on S" + and bounded: "\K. \compact K; K \ S\ \ \B. \h \ \. \ z \ K. norm(h z) \ B" + and rng_f: "range \ \ \" + obtains g r + where "g holomorphic_on S" "strict_mono (r :: nat \ nat)" + "\x. x \ S \ ((\n. \ (r n) x) \ g x) sequentially" + "\K. \compact K; K \ S\ \ uniform_limit K (\ \ r) g sequentially" +proof - + obtain K where comK: "\n. compact(K n)" and KS: "\n::nat. K n \ S" + and subK: "\X. \compact X; X \ S\ \ \N. \n\N. X \ K n" + using open_Union_compact_subsets [OF \open S\] by metis + then have "\i. \B. \h \ \. \ z \ K i. norm(h z) \ B" + by (simp add: bounded) + then obtain B where B: "\i h z. \h \ \; z \ K i\ \ norm(h z) \ B i" + by metis + have *: "\r g. strict_mono (r::nat\nat) \ (\e > 0. \N. \n\N. \x \ K i. norm((\ \ r) n x - g x) < e)" + if "\n. \ n \ \" for \ i + proof - + obtain g k where "continuous_on (K i) g" "strict_mono (k::nat\nat)" + "\e. 0 < e \ \N. \n\N. \x \ K i. norm(\(k n) x - g x) < e" + proof (rule Arzela_Ascoli [of "K i" "\" "B i"]) + show "\d>0. \n y. y \ K i \ cmod (z - y) < d \ cmod (\ n z - \ n y) < e" + if z: "z \ K i" and "0 < e" for z e + proof - + obtain r where "0 < r" and r: "cball z r \ S" + using z KS [of i] \open S\ by (force simp: open_contains_cball) + have "cball z (2 / 3 * r) \ cball z r" + using \0 < r\ by (simp add: cball_subset_cball_iff) + then have z23S: "cball z (2 / 3 * r) \ S" + using r by blast + obtain M where "0 < M" and M: "\n w. dist z w \ 2/3 * r \ norm(\ n w) \ M" + proof - + obtain N where N: "\n\N. cball z (2/3 * r) \ K n" + using subK compact_cball [of z "(2 / 3 * r)"] z23S by force + have "cmod (\ n w) \ \B N\ + 1" if "dist z w \ 2 / 3 * r" for n w + proof - + have "w \ K N" + using N mem_cball that by blast + then have "cmod (\ n w) \ B N" + using B \\n. \ n \ \\ by blast + also have "... \ \B N\ + 1" + by simp + finally show ?thesis . + qed + then show ?thesis + by (rule_tac M="\B N\ + 1" in that) auto + qed + have "cmod (\ n z - \ n y) < e" + if "y \ K i" and y_near_z: "cmod (z - y) < r/3" "cmod (z - y) < e * r / (6 * M)" + for n y + proof - + have "((\w. \ n w / (w - \)) has_contour_integral + (2 * pi) * \ * winding_number (circlepath z (2 / 3 * r)) \ * \ n \) + (circlepath z (2 / 3 * r))" + if "dist \ z < (2 / 3 * r)" for \ + proof (rule Cauchy_integral_formula_convex_simple) + have "\ n holomorphic_on S" + by (simp add: \ \\n. \ n \ \\) + with z23S show "\ n holomorphic_on cball z (2 / 3 * r)" + using holomorphic_on_subset by blast + qed (use that \0 < r\ in \auto simp: dist_commute\) + then have *: "((\w. \ n w / (w - \)) has_contour_integral (2 * pi) * \ * \ n \) + (circlepath z (2 / 3 * r))" + if "dist \ z < (2 / 3 * r)" for \ + using that by (simp add: winding_number_circlepath dist_norm) + have y: "((\w. \ n w / (w - y)) has_contour_integral (2 * pi) * \ * \ n y) + (circlepath z (2 / 3 * r))" + apply (rule *) + using that \0 < r\ by (simp only: dist_norm norm_minus_commute) + have z: "((\w. \ n w / (w - z)) has_contour_integral (2 * pi) * \ * \ n z) + (circlepath z (2 / 3 * r))" + apply (rule *) + using \0 < r\ by simp + have le_er: "cmod (\ n x / (x - y) - \ n x / (x - z)) \ e / r" + if "cmod (x - z) = r/3 + r/3" for x + proof - + have "\ (cmod (x - y) < r/3)" + using y_near_z(1) that \M > 0\ \r > 0\ + by (metis (full_types) norm_diff_triangle_less norm_minus_commute order_less_irrefl) + then have r4_le_xy: "r/4 \ cmod (x - y)" + using \r > 0\ by simp + then have neq: "x \ y" "x \ z" + using that \r > 0\ by (auto simp: field_split_simps norm_minus_commute) + have leM: "cmod (\ n x) \ M" + by (simp add: M dist_commute dist_norm that) + have "cmod (\ n x / (x - y) - \ n x / (x - z)) = cmod (\ n x) * cmod (1 / (x - y) - 1 / (x - z))" + by (metis (no_types, lifting) divide_inverse mult.left_neutral norm_mult right_diff_distrib') + also have "... = cmod (\ n x) * cmod ((y - z) / ((x - y) * (x - z)))" + using neq by (simp add: field_split_simps) + also have "... = cmod (\ n x) * (cmod (y - z) / (cmod(x - y) * (2/3 * r)))" + by (simp add: norm_mult norm_divide that) + also have "... \ M * (cmod (y - z) / (cmod(x - y) * (2/3 * r)))" + apply (rule mult_mono) + apply (rule leM) + using \r > 0\ \M > 0\ neq by auto + also have "... < M * ((e * r / (6 * M)) / (cmod(x - y) * (2/3 * r)))" + unfolding mult_less_cancel_left + using y_near_z(2) \M > 0\ \r > 0\ neq + apply (simp add: field_simps mult_less_0_iff norm_minus_commute) + done + also have "... \ e/r" + using \e > 0\ \r > 0\ r4_le_xy by (simp add: field_split_simps) + finally show ?thesis by simp + qed + have "(2 * pi) * cmod (\ n y - \ n z) = cmod ((2 * pi) * \ * \ n y - (2 * pi) * \ * \ n z)" + by (simp add: right_diff_distrib [symmetric] norm_mult) + also have "cmod ((2 * pi) * \ * \ n y - (2 * pi) * \ * \ n z) \ e / r * (2 * pi * (2 / 3 * r))" + apply (rule has_contour_integral_bound_circlepath [OF has_contour_integral_diff [OF y z], of "e/r"]) + using \e > 0\ \r > 0\ le_er by auto + also have "... = (2 * pi) * e * ((2 / 3))" + using \r > 0\ by (simp add: field_split_simps) + finally have "cmod (\ n y - \ n z) \ e * (2 / 3)" + by simp + also have "... < e" + using \e > 0\ by simp + finally show ?thesis by (simp add: norm_minus_commute) + qed + then show ?thesis + apply (rule_tac x="min (r/3) ((e * r)/(6 * M))" in exI) + using \0 < e\ \0 < r\ \0 < M\ by simp + qed + show "\n x. x \ K i \ cmod (\ n x) \ B i" + using B \\n. \ n \ \\ by blast + next + fix g :: "complex \ complex" and k :: "nat \ nat" + assume *: "\(g::complex\complex) (k::nat\nat). continuous_on (K i) g \ + strict_mono k \ + (\e. 0 < e \ \N. \n\N. \x\K i. cmod (\ (k n) x - g x) < e) \ thesis" + "continuous_on (K i) g" + "strict_mono k" + "\e. 0 < e \ \N. \n x. N \ n \ x \ K i \ cmod (\ (k n) x - g x) < e" + show ?thesis + by (rule *(1)[OF *(2,3)], drule *(4)) auto + qed (use comK in simp_all) + then show ?thesis + by auto + qed + have "\k g. strict_mono (k::nat\nat) \ (\e > 0. \N. \n\N. \x \ K i. norm((\ \ r \ k) n x - g x) < e)" + for i r + apply (rule *) + using rng_f by auto + then have **: "\i r. \k. strict_mono (k::nat\nat) \ (\g. \e>0. \N. \n\N. \x \ K i. norm((\ \ (r \ k)) n x - g x) < e)" + by (force simp: o_assoc) + obtain k :: "nat \ nat" where "strict_mono k" + and "\i. \g. \e>0. \N. \n\N. \x\K i. cmod ((\ \ (id \ k)) n x - g x) < e" + (* TODO: clean up this mess *) + apply (rule subsequence_diagonalization_lemma [OF **, of id id]) + apply (erule ex_forward all_forward imp_forward)+ + apply force + apply (erule exE) + apply (rename_tac i r k1 k2 N g e Na) + apply (rule_tac x="max N Na" in exI) + apply fastforce+ + done + then have lt_e: "\i. \g. \e>0. \N. \n\N. \x\K i. cmod ((\ \ k) n x - g x) < e" + by simp + have "\l. \e>0. \N. \n\N. norm(\ (k n) z - l) < e" if "z \ S" for z + proof - + obtain G where G: "\i e. e > 0 \ \M. \n\M. \x\K i. cmod ((\ \ k) n x - G i x) < e" + using lt_e by metis + obtain N where "\n. n \ N \ z \ K n" + using subK [of "{z}"] that \z \ S\ by auto + moreover have "\e. e > 0 \ \M. \n\M. \x\K N. cmod ((\ \ k) n x - G N x) < e" + using G by auto + ultimately show ?thesis + by (metis comp_apply order_refl) + qed + then obtain g where g: "\z e. \z \ S; e > 0\ \ \N. \n\N. norm(\ (k n) z - g z) < e" + by metis + show ?thesis + proof + show g_lim: "\x. x \ S \ (\n. \ (k n) x) \ g x" + by (simp add: lim_sequentially g dist_norm) + have dg_le_e: "\N. \n\N. \x\T. cmod (\ (k n) x - g x) < e" + if T: "compact T" "T \ S" and "0 < e" for T e + proof - + obtain N where N: "\n. n \ N \ T \ K n" + using subK [OF T] by blast + obtain h where h: "\e. e>0 \ \M. \n\M. \x\K N. cmod ((\ \ k) n x - h x) < e" + using lt_e by blast + have geq: "g w = h w" if "w \ T" for w + apply (rule LIMSEQ_unique [of "\n. \(k n) w"]) + using \T \ S\ g_lim that apply blast + using h N that by (force simp: lim_sequentially dist_norm) + show ?thesis + using T h N \0 < e\ by (fastforce simp add: geq) + qed + then show "\K. \compact K; K \ S\ + \ uniform_limit K (\ \ k) g sequentially" + by (simp add: uniform_limit_iff dist_norm eventually_sequentially) + show "g holomorphic_on S" + proof (rule holomorphic_uniform_sequence [OF \open S\ \]) + show "\n. (\ \ k) n \ \" + by (simp add: range_subsetD rng_f) + show "\d>0. cball z d \ S \ uniform_limit (cball z d) (\n. (\ \ k) n) g sequentially" + if "z \ S" for z + proof - + obtain d where d: "d>0" "cball z d \ S" + using \open S\ \z \ S\ open_contains_cball by blast + then have "uniform_limit (cball z d) (\ \ k) g sequentially" + using dg_le_e compact_cball by (auto simp: uniform_limit_iff eventually_sequentially dist_norm) + with d show ?thesis by blast + qed + qed + qed (auto simp: \strict_mono k\) +qed + + + +subsection\Some simple but useful cases of Hurwitz's theorem\ + +proposition Hurwitz_no_zeros: + assumes S: "open S" "connected S" + and holf: "\n::nat. \ n holomorphic_on S" + and holg: "g holomorphic_on S" + and ul_g: "\K. \compact K; K \ S\ \ uniform_limit K \ g sequentially" + and nonconst: "\ g constant_on S" + and nz: "\n z. z \ S \ \ n z \ 0" + and "z0 \ S" + shows "g z0 \ 0" +proof + assume g0: "g z0 = 0" + obtain h r m + where "0 < m" "0 < r" and subS: "ball z0 r \ S" + and holh: "h holomorphic_on ball z0 r" + and geq: "\w. w \ ball z0 r \ g w = (w - z0)^m * h w" + and hnz: "\w. w \ ball z0 r \ h w \ 0" + by (blast intro: holomorphic_factor_zero_nonconstant [OF holg S \z0 \ S\ g0 nonconst]) + then have holf0: "\ n holomorphic_on ball z0 r" for n + by (meson holf holomorphic_on_subset) + have *: "((\z. deriv (\ n) z / \ n z) has_contour_integral 0) (circlepath z0 (r/2))" for n + proof (rule Cauchy_theorem_disc_simple [of _ z0 r]) + show "(\z. deriv (\ n) z / \ n z) holomorphic_on ball z0 r" + apply (intro holomorphic_intros holomorphic_deriv holf holf0 open_ball nz) + using \ball z0 r \ S\ by blast + qed (use \0 < r\ in auto) + have hol_dg: "deriv g holomorphic_on S" + by (simp add: \open S\ holg holomorphic_deriv) + have "continuous_on (sphere z0 (r/2)) (deriv g)" + apply (intro holomorphic_on_imp_continuous_on holomorphic_on_subset [OF hol_dg]) + using \0 < r\ subS by auto + then have "compact (deriv g ` (sphere z0 (r/2)))" + by (rule compact_continuous_image [OF _ compact_sphere]) + then have bo_dg: "bounded (deriv g ` (sphere z0 (r/2)))" + using compact_imp_bounded by blast + have "continuous_on (sphere z0 (r/2)) (cmod \ g)" + apply (intro continuous_intros holomorphic_on_imp_continuous_on holomorphic_on_subset [OF holg]) + using \0 < r\ subS by auto + then have "compact ((cmod \ g) ` sphere z0 (r/2))" + by (rule compact_continuous_image [OF _ compact_sphere]) + moreover have "(cmod \ g) ` sphere z0 (r/2) \ {}" + using \0 < r\ by auto + ultimately obtain b where b: "b \ (cmod \ g) ` sphere z0 (r/2)" + "\t. t \ (cmod \ g) ` sphere z0 (r/2) \ b \ t" + using compact_attains_inf [of "(norm \ g) ` (sphere z0 (r/2))"] by blast + have "(\n. contour_integral (circlepath z0 (r/2)) (\z. deriv (\ n) z / \ n z)) \ + contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z)" + proof (rule contour_integral_uniform_limit_circlepath) + show "\\<^sub>F n in sequentially. (\z. deriv (\ n) z / \ n z) contour_integrable_on circlepath z0 (r/2)" + using * contour_integrable_on_def eventually_sequentiallyI by meson + show "uniform_limit (sphere z0 (r/2)) (\n z. deriv (\ n) z / \ n z) (\z. deriv g z / g z) sequentially" + proof (rule uniform_lim_divide [OF _ _ bo_dg]) + show "uniform_limit (sphere z0 (r/2)) (\a. deriv (\ a)) (deriv g) sequentially" + proof (rule uniform_limitI) + fix e::real + assume "0 < e" + have *: "dist (deriv (\ n) w) (deriv g w) < e" + if e8: "\x. dist z0 x \ 3 * r / 4 \ dist (\ n x) (g x) * 8 < r * e" + and w: "dist w z0 = r/2" for n w + proof - + have "ball w (r/4) \ ball z0 r" "cball w (r/4) \ ball z0 r" + using \0 < r\ by (simp_all add: ball_subset_ball_iff cball_subset_ball_iff w) + with subS have wr4_sub: "ball w (r/4) \ S" "cball w (r/4) \ S" by force+ + moreover + have "(\z. \ n z - g z) holomorphic_on S" + by (intro holomorphic_intros holf holg) + ultimately have hol: "(\z. \ n z - g z) holomorphic_on ball w (r/4)" + and cont: "continuous_on (cball w (r / 4)) (\z. \ n z - g z)" + using holomorphic_on_subset by (blast intro: holomorphic_on_imp_continuous_on)+ + have "w \ S" + using \0 < r\ wr4_sub by auto + have "\y. dist w y < r / 4 \ dist z0 y \ 3 * r / 4" + apply (rule dist_triangle_le [where z=w]) + using w by (simp add: dist_commute) + with e8 have in_ball: "\y. y \ ball w (r/4) \ \ n y - g y \ ball 0 (r/4 * e/2)" + by (simp add: dist_norm [symmetric]) + have "\ n field_differentiable at w" + by (metis holomorphic_on_imp_differentiable_at \w \ S\ holf \open S\) + moreover + have "g field_differentiable at w" + using \w \ S\ \open S\ holg holomorphic_on_imp_differentiable_at by auto + moreover + have "cmod (deriv (\w. \ n w - g w) w) * 2 \ e" + apply (rule Cauchy_higher_deriv_bound [OF hol cont in_ball, of 1, simplified]) + using \r > 0\ by auto + ultimately have "dist (deriv (\ n) w) (deriv g w) \ e/2" + by (simp add: dist_norm) + then show ?thesis + using \e > 0\ by auto + qed + have "cball z0 (3 * r / 4) \ ball z0 r" + by (simp add: cball_subset_ball_iff \0 < r\) + with subS have "uniform_limit (cball z0 (3 * r/4)) \ g sequentially" + by (force intro: ul_g) + then have "\\<^sub>F n in sequentially. \x\cball z0 (3 * r / 4). dist (\ n x) (g x) < r / 4 * e / 2" + using \0 < e\ \0 < r\ by (force simp: intro!: uniform_limitD) + then show "\\<^sub>F n in sequentially. \x \ sphere z0 (r/2). dist (deriv (\ n) x) (deriv g x) < e" + apply (simp add: eventually_sequentially) + apply (elim ex_forward all_forward imp_forward asm_rl) + using * apply (force simp: dist_commute) + done + qed + show "uniform_limit (sphere z0 (r/2)) \ g sequentially" + proof (rule uniform_limitI) + fix e::real + assume "0 < e" + have "sphere z0 (r/2) \ ball z0 r" + using \0 < r\ by auto + with subS have "uniform_limit (sphere z0 (r/2)) \ g sequentially" + by (force intro: ul_g) + then show "\\<^sub>F n in sequentially. \x \ sphere z0 (r/2). dist (\ n x) (g x) < e" + apply (rule uniform_limitD) + using \0 < e\ by force + qed + show "b > 0" "\x. x \ sphere z0 (r/2) \ b \ cmod (g x)" + using b \0 < r\ by (fastforce simp: geq hnz)+ + qed + qed (use \0 < r\ in auto) + then have "(\n. 0) \ contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z)" + by (simp add: contour_integral_unique [OF *]) + then have "contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z) = 0" + by (simp add: LIMSEQ_const_iff) + moreover + have "contour_integral (circlepath z0 (r/2)) (\z. deriv g z / g z) = + contour_integral (circlepath z0 (r/2)) (\z. m / (z - z0) + deriv h z / h z)" + proof (rule contour_integral_eq, use \0 < r\ in simp) + fix w + assume w: "dist z0 w * 2 = r" + then have w_inb: "w \ ball z0 r" + using \0 < r\ by auto + have h_der: "(h has_field_derivative deriv h w) (at w)" + using holh holomorphic_derivI w_inb by blast + have "deriv g w = ((of_nat m * h w + deriv h w * (w - z0)) * (w - z0) ^ m) / (w - z0)" + if "r = dist z0 w * 2" "w \ z0" + proof - + have "((\w. (w - z0) ^ m * h w) has_field_derivative + (m * h w + deriv h w * (w - z0)) * (w - z0) ^ m / (w - z0)) (at w)" + apply (rule derivative_eq_intros h_der refl)+ + using that \m > 0\ \0 < r\ apply (simp add: divide_simps distrib_right) + apply (metis Suc_pred mult.commute power_Suc) + done + then show ?thesis + apply (rule DERIV_imp_deriv [OF has_field_derivative_transform_within_open [where S = "ball z0 r"]]) + using that \m > 0\ \0 < r\ + apply (simp_all add: hnz geq) + done + qed + with \0 < r\ \0 < m\ w w_inb show "deriv g w / g w = of_nat m / (w - z0) + deriv h w / h w" + by (auto simp: geq field_split_simps hnz) + qed + moreover + have "contour_integral (circlepath z0 (r/2)) (\z. m / (z - z0) + deriv h z / h z) = + 2 * of_real pi * \ * m + 0" + proof (rule contour_integral_unique [OF has_contour_integral_add]) + show "((\x. m / (x - z0)) has_contour_integral 2 * of_real pi * \ * m) (circlepath z0 (r/2))" + by (force simp: \0 < r\ intro: Cauchy_integral_circlepath_simple) + show "((\x. deriv h x / h x) has_contour_integral 0) (circlepath z0 (r/2))" + apply (rule Cauchy_theorem_disc_simple [of _ z0 r]) + using hnz holh holomorphic_deriv holomorphic_on_divide \0 < r\ + apply force+ + done + qed + ultimately show False using \0 < m\ by auto +qed + +corollary Hurwitz_injective: + assumes S: "open S" "connected S" + and holf: "\n::nat. \ n holomorphic_on S" + and holg: "g holomorphic_on S" + and ul_g: "\K. \compact K; K \ S\ \ uniform_limit K \ g sequentially" + and nonconst: "\ g constant_on S" + and inj: "\n. inj_on (\ n) S" + shows "inj_on g S" +proof - + have False if z12: "z1 \ S" "z2 \ S" "z1 \ z2" "g z2 = g z1" for z1 z2 + proof - + obtain z0 where "z0 \ S" and z0: "g z0 \ g z2" + using constant_on_def nonconst by blast + have "(\z. g z - g z1) holomorphic_on S" + by (intro holomorphic_intros holg) + then obtain r where "0 < r" "ball z2 r \ S" "\z. dist z2 z < r \ z \ z2 \ g z \ g z1" + apply (rule isolated_zeros [of "\z. g z - g z1" S z2 z0]) + using S \z0 \ S\ z0 z12 by auto + have "g z2 - g z1 \ 0" + proof (rule Hurwitz_no_zeros [of "S - {z1}" "\n z. \ n z - \ n z1" "\z. g z - g z1"]) + show "open (S - {z1})" + by (simp add: S open_delete) + show "connected (S - {z1})" + by (simp add: connected_open_delete [OF S]) + show "\n. (\z. \ n z - \ n z1) holomorphic_on S - {z1}" + by (intro holomorphic_intros holomorphic_on_subset [OF holf]) blast + show "(\z. g z - g z1) holomorphic_on S - {z1}" + by (intro holomorphic_intros holomorphic_on_subset [OF holg]) blast + show "uniform_limit K (\n z. \ n z - \ n z1) (\z. g z - g z1) sequentially" + if "compact K" "K \ S - {z1}" for K + proof (rule uniform_limitI) + fix e::real + assume "e > 0" + have "uniform_limit K \ g sequentially" + using that ul_g by fastforce + then have K: "\\<^sub>F n in sequentially. \x \ K. dist (\ n x) (g x) < e/2" + using \0 < e\ by (force simp: intro!: uniform_limitD) + have "uniform_limit {z1} \ g sequentially" + by (simp add: ul_g z12) + then have "\\<^sub>F n in sequentially. \x \ {z1}. dist (\ n x) (g x) < e/2" + using \0 < e\ by (force simp: intro!: uniform_limitD) + then have z1: "\\<^sub>F n in sequentially. dist (\ n z1) (g z1) < e/2" + by simp + have "\\<^sub>F n in sequentially. \x\K. dist (\ n x - \ n z1) (g x - g z1) < e/2 + e/2" + apply (rule eventually_mono [OF eventually_conj [OF K z1]]) + apply (simp add: dist_norm algebra_simps del: divide_const_simps) + by (metis add.commute dist_commute dist_norm dist_triangle_add_half) + have "\\<^sub>F n in sequentially. \x\K. dist (\ n x - \ n z1) (g x - g z1) < e/2 + e/2" + using eventually_conj [OF K z1] + apply (rule eventually_mono) + by (metis (no_types, hide_lams) diff_add_eq diff_diff_eq2 dist_commute dist_norm dist_triangle_add_half field_sum_of_halves) + then + show "\\<^sub>F n in sequentially. \x\K. dist (\ n x - \ n z1) (g x - g z1) < e" + by simp + qed + show "\ (\z. g z - g z1) constant_on S - {z1}" + unfolding constant_on_def + by (metis Diff_iff \z0 \ S\ empty_iff insert_iff right_minus_eq z0 z12) + show "\n z. z \ S - {z1} \ \ n z - \ n z1 \ 0" + by (metis DiffD1 DiffD2 eq_iff_diff_eq_0 inj inj_onD insertI1 \z1 \ S\) + show "z2 \ S - {z1}" + using \z2 \ S\ \z1 \ z2\ by auto + qed + with z12 show False by auto + qed + then show ?thesis by (auto simp: inj_on_def) +qed + + + +subsection\The Great Picard theorem\ + +lemma GPicard1: + assumes S: "open S" "connected S" and "w \ S" "0 < r" "Y \ X" + and holX: "\h. h \ X \ h holomorphic_on S" + and X01: "\h z. \h \ X; z \ S\ \ h z \ 0 \ h z \ 1" + and r: "\h. h \ Y \ norm(h w) \ r" + obtains B Z where "0 < B" "open Z" "w \ Z" "Z \ S" "\h z. \h \ Y; z \ Z\ \ norm(h z) \ B" +proof - + obtain e where "e > 0" and e: "cball w e \ S" + using assms open_contains_cball_eq by blast + show ?thesis + proof + show "0 < exp(pi * exp(pi * (2 + 2 * r + 12)))" + by simp + show "ball w (e / 2) \ S" + using e ball_divide_subset_numeral ball_subset_cball by blast + show "cmod (h z) \ exp (pi * exp (pi * (2 + 2 * r + 12)))" + if "h \ Y" "z \ ball w (e / 2)" for h z + proof - + have "h \ X" + using \Y \ X\ \h \ Y\ by blast + with holX have "h holomorphic_on S" + by auto + then have "h holomorphic_on cball w e" + by (metis e holomorphic_on_subset) + then have hol_h_o: "(h \ (\z. (w + of_real e * z))) holomorphic_on cball 0 1" + apply (intro holomorphic_intros holomorphic_on_compose) + apply (erule holomorphic_on_subset) + using that \e > 0\ by (auto simp: dist_norm norm_mult) + have norm_le_r: "cmod ((h \ (\z. w + complex_of_real e * z)) 0) \ r" + by (auto simp: r \h \ Y\) + have le12: "norm (of_real(inverse e) * (z - w)) \ 1/2" + using that \e > 0\ by (simp add: inverse_eq_divide dist_norm norm_minus_commute norm_divide) + have non01: "\z::complex. cmod z \ 1 \ h (w + e * z) \ 0 \ h (w + e * z) \ 1" + apply (rule X01 [OF \h \ X\]) + apply (rule subsetD [OF e]) + using \0 < e\ by (auto simp: dist_norm norm_mult) + have "cmod (h z) \ cmod (h (w + of_real e * (inverse e * (z - w))))" + using \0 < e\ by (simp add: field_split_simps) + also have "... \ exp (pi * exp (pi * (14 + 2 * r)))" + using r [OF \h \ Y\] Schottky [OF hol_h_o norm_le_r _ _ _ le12] non01 by auto + finally + show ?thesis by simp + qed + qed (use \e > 0\ in auto) +qed + +lemma GPicard2: + assumes "S \ T" "connected T" "S \ {}" "open S" "\x. \x islimpt S; x \ T\ \ x \ S" + shows "S = T" + by (metis assms open_subset connected_clopen closedin_limpt) + + +lemma GPicard3: + assumes S: "open S" "connected S" "w \ S" and "Y \ X" + and holX: "\h. h \ X \ h holomorphic_on S" + and X01: "\h z. \h \ X; z \ S\ \ h z \ 0 \ h z \ 1" + and no_hw_le1: "\h. h \ Y \ norm(h w) \ 1" + and "compact K" "K \ S" + obtains B where "\h z. \h \ Y; z \ K\ \ norm(h z) \ B" +proof - + define U where "U \ {z \ S. \B Z. 0 < B \ open Z \ z \ Z \ Z \ S \ + (\h z'. h \ Y \ z' \ Z \ norm(h z') \ B)}" + then have "U \ S" by blast + have "U = S" + proof (rule GPicard2 [OF \U \ S\ \connected S\]) + show "U \ {}" + proof - + obtain B Z where "0 < B" "open Z" "w \ Z" "Z \ S" + and "\h z. \h \ Y; z \ Z\ \ norm(h z) \ B" + apply (rule GPicard1 [OF S zero_less_one \Y \ X\ holX]) + using no_hw_le1 X01 by force+ + then show ?thesis + unfolding U_def using \w \ S\ by blast + qed + show "open U" + unfolding open_subopen [of U] by (auto simp: U_def) + fix v + assume v: "v islimpt U" "v \ S" + have "\ (\r>0. \h\Y. r < cmod (h v))" + proof + assume "\r>0. \h\Y. r < cmod (h v)" + then have "\n. \h\Y. Suc n < cmod (h v)" + by simp + then obtain \ where FY: "\n. \ n \ Y" and ltF: "\n. Suc n < cmod (\ n v)" + by metis + define \ where "\ \ \n z. inverse(\ n z)" + have hol\: "\ n holomorphic_on S" for n + apply (simp add: \_def) + using FY X01 \Y \ X\ holX apply (blast intro: holomorphic_on_inverse) + done + have \not0: "\ n z \ 0" and \not1: "\ n z \ 1" if "z \ S" for n z + using FY X01 \Y \ X\ that by (force simp: \_def)+ + have \_le1: "cmod (\ n v) \ 1" for n + using less_le_trans linear ltF + by (fastforce simp add: \_def norm_inverse inverse_le_1_iff) + define W where "W \ {h. h holomorphic_on S \ (\z \ S. h z \ 0 \ h z \ 1)}" + obtain B Z where "0 < B" "open Z" "v \ Z" "Z \ S" + and B: "\h z. \h \ range \; z \ Z\ \ norm(h z) \ B" + apply (rule GPicard1 [OF \open S\ \connected S\ \v \ S\ zero_less_one, of "range \" W]) + using hol\ \not0 \not1 \_le1 by (force simp: W_def)+ + then obtain e where "e > 0" and e: "ball v e \ Z" + by (meson open_contains_ball) + obtain h j where holh: "h holomorphic_on ball v e" and "strict_mono j" + and lim: "\x. x \ ball v e \ (\n. \ (j n) x) \ h x" + and ulim: "\K. \compact K; K \ ball v e\ + \ uniform_limit K (\ \ j) h sequentially" + proof (rule Montel) + show "\h. h \ range \ \ h holomorphic_on ball v e" + by (metis \Z \ S\ e hol\ holomorphic_on_subset imageE) + show "\K. \compact K; K \ ball v e\ \ \B. \h\range \. \z\K. cmod (h z) \ B" + using B e by blast + qed auto + have "h v = 0" + proof (rule LIMSEQ_unique) + show "(\n. \ (j n) v) \ h v" + using \e > 0\ lim by simp + have lt_Fj: "real x \ cmod (\ (j x) v)" for x + by (metis of_nat_Suc ltF \strict_mono j\ add.commute less_eq_real_def less_le_trans nat_le_real_less seq_suble) + show "(\n. \ (j n) v) \ 0" + proof (rule Lim_null_comparison [OF eventually_sequentiallyI lim_inverse_n]) + show "cmod (\ (j x) v) \ inverse (real x)" if "1 \ x" for x + using that by (simp add: \_def norm_inverse_le_norm [OF lt_Fj]) + qed + qed + have "h v \ 0" + proof (rule Hurwitz_no_zeros [of "ball v e" "\ \ j" h]) + show "\n. (\ \ j) n holomorphic_on ball v e" + using \Z \ S\ e hol\ by force + show "\n z. z \ ball v e \ (\ \ j) n z \ 0" + using \not0 \Z \ S\ e by fastforce + show "\ h constant_on ball v e" + proof (clarsimp simp: constant_on_def) + fix c + have False if "\z. dist v z < e \ h z = c" + proof - + have "h v = c" + by (simp add: \0 < e\ that) + obtain y where "y \ U" "y \ v" and y: "dist y v < e" + using v \e > 0\ by (auto simp: islimpt_approachable) + then obtain C T where "y \ S" "C > 0" "open T" "y \ T" "T \ S" + and "\h z'. \h \ Y; z' \ T\ \ cmod (h z') \ C" + using \y \ U\ by (auto simp: U_def) + then have le_C: "\n. cmod (\ n y) \ C" + using FY by blast + have "\\<^sub>F n in sequentially. dist (\ (j n) y) (h y) < inverse C" + using uniform_limitD [OF ulim [of "{y}"], of "inverse C"] \C > 0\ y + by (simp add: dist_commute) + then obtain n where "dist (\ (j n) y) (h y) < inverse C" + by (meson eventually_at_top_linorder order_refl) + moreover + have "h y = h v" + by (metis \h v = c\ dist_commute that y) + ultimately have "norm (\ (j n) y) < inverse C" + by (simp add: \h v = 0\) + then have "C < norm (\ (j n) y)" + apply (simp add: \_def) + by (metis FY X01 \0 < C\ \y \ S\ \Y \ X\ inverse_less_iff_less norm_inverse subsetD zero_less_norm_iff) + show False + using \C < cmod (\ (j n) y)\ le_C not_less by blast + qed + then show "\x\ball v e. h x \ c" by force + qed + show "h holomorphic_on ball v e" + by (simp add: holh) + show "\K. \compact K; K \ ball v e\ \ uniform_limit K (\ \ j) h sequentially" + by (simp add: ulim) + qed (use \e > 0\ in auto) + with \h v = 0\ show False by blast + qed + then show "v \ U" + apply (clarsimp simp add: U_def v) + apply (rule GPicard1[OF \open S\ \connected S\ \v \ S\ _ \Y \ X\ holX]) + using X01 no_hw_le1 apply (meson | force simp: not_less)+ + done + qed + have "\x. x \ K \ x \ U" + using \U = S\ \K \ S\ by blast + then have "\x. x \ K \ (\B Z. 0 < B \ open Z \ x \ Z \ + (\h z'. h \ Y \ z' \ Z \ norm(h z') \ B))" + unfolding U_def by blast + then obtain F Z where F: "\x. x \ K \ open (Z x) \ x \ Z x \ + (\h z'. h \ Y \ z' \ Z x \ norm(h z') \ F x)" + by metis + then obtain L where "L \ K" "finite L" and L: "K \ (\c \ L. Z c)" + by (auto intro: compactE_image [OF \compact K\, of K Z]) + then have *: "\x h z'. \x \ L; h \ Y \ z' \ Z x\ \ cmod (h z') \ F x" + using F by blast + have "\B. \h z. h \ Y \ z \ K \ norm(h z) \ B" + proof (cases "L = {}") + case True with L show ?thesis by simp + next + case False + with \finite L\ show ?thesis + apply (rule_tac x = "Max (F ` L)" in exI) + apply (simp add: linorder_class.Max_ge_iff) + using * F by (metis L UN_E subsetD) + qed + with that show ?thesis by metis +qed + + +lemma GPicard4: + assumes "0 < k" and holf: "f holomorphic_on (ball 0 k - {0})" + and AE: "\e. \0 < e; e < k\ \ \d. 0 < d \ d < e \ (\z \ sphere 0 d. norm(f z) \ B)" + obtains \ where "0 < \" "\ < k" "\z. z \ ball 0 \ - {0} \ norm(f z) \ B" +proof - + obtain \ where "0 < \" "\ < k/2" and \: "\z. norm z = \ \ norm(f z) \ B" + using AE [of "k/2"] \0 < k\ by auto + show ?thesis + proof + show "\ < k" + using \0 < k\ \\ < k/2\ by auto + show "cmod (f \) \ B" if \: "\ \ ball 0 \ - {0}" for \ + proof - + obtain d where "0 < d" "d < norm \" and d: "\z. norm z = d \ norm(f z) \ B" + using AE [of "norm \"] \\ < k\ \ by auto + have [simp]: "closure (cball 0 \ - ball 0 d) = cball 0 \ - ball 0 d" + by (blast intro!: closure_closed) + have [simp]: "interior (cball 0 \ - ball 0 d) = ball 0 \ - cball (0::complex) d" + using \0 < \\ \0 < d\ by (simp add: interior_diff) + have *: "norm(f w) \ B" if "w \ cball 0 \ - ball 0 d" for w + proof (rule maximum_modulus_frontier [of f "cball 0 \ - ball 0 d"]) + show "f holomorphic_on interior (cball 0 \ - ball 0 d)" + apply (rule holomorphic_on_subset [OF holf]) + using \\ < k\ \0 < d\ that by auto + show "continuous_on (closure (cball 0 \ - ball 0 d)) f" + apply (rule holomorphic_on_imp_continuous_on) + apply (rule holomorphic_on_subset [OF holf]) + using \0 < d\ \\ < k\ by auto + show "\z. z \ frontier (cball 0 \ - ball 0 d) \ cmod (f z) \ B" + apply (simp add: frontier_def) + using \ d less_eq_real_def by blast + qed (use that in auto) + show ?thesis + using * \d < cmod \\ that by auto + qed + qed (use \0 < \\ in auto) +qed + + +lemma GPicard5: + assumes holf: "f holomorphic_on (ball 0 1 - {0})" + and f01: "\z. z \ ball 0 1 - {0} \ f z \ 0 \ f z \ 1" + obtains e B where "0 < e" "e < 1" "0 < B" + "(\z \ ball 0 e - {0}. norm(f z) \ B) \ + (\z \ ball 0 e - {0}. norm(f z) \ B)" +proof - + have [simp]: "1 + of_nat n \ (0::complex)" for n + using of_nat_eq_0_iff by fastforce + have [simp]: "cmod (1 + of_nat n) = 1 + of_nat n" for n + by (metis norm_of_nat of_nat_Suc) + have *: "(\x::complex. x / of_nat (Suc n)) ` (ball 0 1 - {0}) \ ball 0 1 - {0}" for n + by (auto simp: norm_divide field_split_simps split: if_split_asm) + define h where "h \ \n z::complex. f (z / (Suc n))" + have holh: "(h n) holomorphic_on ball 0 1 - {0}" for n + unfolding h_def + proof (rule holomorphic_on_compose_gen [unfolded o_def, OF _ holf *]) + show "(\x. x / of_nat (Suc n)) holomorphic_on ball 0 1 - {0}" + by (intro holomorphic_intros) auto + qed + have h01: "\n z. z \ ball 0 1 - {0} \ h n z \ 0 \ h n z \ 1" + unfolding h_def + apply (rule f01) + using * by force + obtain w where w: "w \ ball 0 1 - {0::complex}" + by (rule_tac w = "1/2" in that) auto + consider "infinite {n. norm(h n w) \ 1}" | "infinite {n. 1 \ norm(h n w)}" + by (metis (mono_tags, lifting) infinite_nat_iff_unbounded_le le_cases mem_Collect_eq) + then show ?thesis + proof cases + case 1 + with infinite_enumerate obtain r :: "nat \ nat" + where "strict_mono r" and r: "\n. r n \ {n. norm(h n w) \ 1}" + by blast + obtain B where B: "\j z. \norm z = 1/2; j \ range (h \ r)\ \ norm(j z) \ B" + proof (rule GPicard3 [OF _ _ w, where K = "sphere 0 (1/2)"]) + show "range (h \ r) \ + {g. g holomorphic_on ball 0 1 - {0} \ (\z\ball 0 1 - {0}. g z \ 0 \ g z \ 1)}" + apply clarsimp + apply (intro conjI holomorphic_intros holomorphic_on_compose holh) + using h01 apply auto + done + show "connected (ball 0 1 - {0::complex})" + by (simp add: connected_open_delete) + qed (use r in auto) + have normf_le_B: "cmod(f z) \ B" if "norm z = 1 / (2 * (1 + of_nat (r n)))" for z n + proof - + have *: "\w. norm w = 1/2 \ cmod((f (w / (1 + of_nat (r n))))) \ B" + using B by (auto simp: h_def o_def) + have half: "norm (z * (1 + of_nat (r n))) = 1/2" + by (simp add: norm_mult divide_simps that) + show ?thesis + using * [OF half] by simp + qed + obtain \ where "0 < \" "\ < 1" "\z. z \ ball 0 \ - {0} \ cmod(f z) \ B" + proof (rule GPicard4 [OF zero_less_one holf, of B]) + fix e::real + assume "0 < e" "e < 1" + obtain n where "(1/e - 2) / 2 < real n" + using reals_Archimedean2 by blast + also have "... \ r n" + using \strict_mono r\ by (simp add: seq_suble) + finally have "(1/e - 2) / 2 < real (r n)" . + with \0 < e\ have e: "e > 1 / (2 + 2 * real (r n))" + by (simp add: field_simps) + show "\d>0. d < e \ (\z\sphere 0 d. cmod (f z) \ B)" + apply (rule_tac x="1 / (2 * (1 + of_nat (r n)))" in exI) + using normf_le_B by (simp add: e) + qed blast + then have \: "cmod (f z) \ \B\ + 1" if "cmod z < \" "z \ 0" for z + using that by fastforce + have "0 < \B\ + 1" + by simp + then show ?thesis + apply (rule that [OF \0 < \\ \\ < 1\]) + using \ by auto + next + case 2 + with infinite_enumerate obtain r :: "nat \ nat" + where "strict_mono r" and r: "\n. r n \ {n. norm(h n w) \ 1}" + by blast + obtain B where B: "\j z. \norm z = 1/2; j \ range (\n. inverse \ h (r n))\ \ norm(j z) \ B" + proof (rule GPicard3 [OF _ _ w, where K = "sphere 0 (1/2)"]) + show "range (\n. inverse \ h (r n)) \ + {g. g holomorphic_on ball 0 1 - {0} \ (\z\ball 0 1 - {0}. g z \ 0 \ g z \ 1)}" + apply clarsimp + apply (intro conjI holomorphic_intros holomorphic_on_compose_gen [unfolded o_def, OF _ holh] holomorphic_on_compose) + using h01 apply auto + done + show "connected (ball 0 1 - {0::complex})" + by (simp add: connected_open_delete) + show "\j. j \ range (\n. inverse \ h (r n)) \ cmod (j w) \ 1" + using r norm_inverse_le_norm by fastforce + qed (use r in auto) + have norm_if_le_B: "cmod(inverse (f z)) \ B" if "norm z = 1 / (2 * (1 + of_nat (r n)))" for z n + proof - + have *: "inverse (cmod((f (z / (1 + of_nat (r n)))))) \ B" if "norm z = 1/2" for z + using B [OF that] by (force simp: norm_inverse h_def) + have half: "norm (z * (1 + of_nat (r n))) = 1/2" + by (simp add: norm_mult divide_simps that) + show ?thesis + using * [OF half] by (simp add: norm_inverse) + qed + have hol_if: "(inverse \ f) holomorphic_on (ball 0 1 - {0})" + by (metis (no_types, lifting) holf comp_apply f01 holomorphic_on_inverse holomorphic_transform) + obtain \ where "0 < \" "\ < 1" and leB: "\z. z \ ball 0 \ - {0} \ cmod((inverse \ f) z) \ B" + proof (rule GPicard4 [OF zero_less_one hol_if, of B]) + fix e::real + assume "0 < e" "e < 1" + obtain n where "(1/e - 2) / 2 < real n" + using reals_Archimedean2 by blast + also have "... \ r n" + using \strict_mono r\ by (simp add: seq_suble) + finally have "(1/e - 2) / 2 < real (r n)" . + with \0 < e\ have e: "e > 1 / (2 + 2 * real (r n))" + by (simp add: field_simps) + show "\d>0. d < e \ (\z\sphere 0 d. cmod ((inverse \ f) z) \ B)" + apply (rule_tac x="1 / (2 * (1 + of_nat (r n)))" in exI) + using norm_if_le_B by (simp add: e) + qed blast + have \: "cmod (f z) \ inverse B" and "B > 0" if "cmod z < \" "z \ 0" for z + proof - + have "inverse (cmod (f z)) \ B" + using leB that by (simp add: norm_inverse) + moreover + have "f z \ 0" + using \\ < 1\ f01 that by auto + ultimately show "cmod (f z) \ inverse B" + by (simp add: norm_inverse inverse_le_imp_le) + show "B > 0" + using \f z \ 0\ \inverse (cmod (f z)) \ B\ not_le order.trans by fastforce + qed + then have "B > 0" + by (metis \0 < \\ dense leI order.asym vector_choose_size) + then have "inverse B > 0" + by (simp add: field_split_simps) + then show ?thesis + apply (rule that [OF \0 < \\ \\ < 1\]) + using \ by auto + qed +qed + + +lemma GPicard6: + assumes "open M" "z \ M" "a \ 0" and holf: "f holomorphic_on (M - {z})" + and f0a: "\w. w \ M - {z} \ f w \ 0 \ f w \ a" + obtains r where "0 < r" "ball z r \ M" + "bounded(f ` (ball z r - {z})) \ + bounded((inverse \ f) ` (ball z r - {z}))" +proof - + obtain r where "0 < r" and r: "ball z r \ M" + using assms openE by blast + let ?g = "\w. f (z + of_real r * w) / a" + obtain e B where "0 < e" "e < 1" "0 < B" + and B: "(\z \ ball 0 e - {0}. norm(?g z) \ B) \ (\z \ ball 0 e - {0}. norm(?g z) \ B)" + proof (rule GPicard5) + show "?g holomorphic_on ball 0 1 - {0}" + apply (intro holomorphic_intros holomorphic_on_compose_gen [unfolded o_def, OF _ holf]) + using \0 < r\ \a \ 0\ r + by (auto simp: dist_norm norm_mult subset_eq) + show "\w. w \ ball 0 1 - {0} \ f (z + of_real r * w) / a \ 0 \ f (z + of_real r * w) / a \ 1" + apply (simp add: field_split_simps \a \ 0\) + apply (rule f0a) + using \0 < r\ r by (auto simp: dist_norm norm_mult subset_eq) + qed + show ?thesis + proof + show "0 < e*r" + by (simp add: \0 < e\ \0 < r\) + have "ball z (e * r) \ ball z r" + by (simp add: \0 < r\ \e < 1\ order.strict_implies_order subset_ball) + then show "ball z (e * r) \ M" + using r by blast + consider "\z. z \ ball 0 e - {0} \ norm(?g z) \ B" | "\z. z \ ball 0 e - {0} \ norm(?g z) \ B" + using B by blast + then show "bounded (f ` (ball z (e * r) - {z})) \ + bounded ((inverse \ f) ` (ball z (e * r) - {z}))" + proof cases + case 1 + have "\dist z w < e * r; w \ z\ \ cmod (f w) \ B * norm a" for w + using \a \ 0\ \0 < r\ 1 [of "(w - z) / r"] + by (simp add: norm_divide dist_norm field_split_simps) + then show ?thesis + by (force simp: intro!: boundedI) + next + case 2 + have "\dist z w < e * r; w \ z\ \ cmod (f w) \ B * norm a" for w + using \a \ 0\ \0 < r\ 2 [of "(w - z) / r"] + by (simp add: norm_divide dist_norm field_split_simps) + then have "\dist z w < e * r; w \ z\ \ inverse (cmod (f w)) \ inverse (B * norm a)" for w + by (metis \0 < B\ \a \ 0\ mult_pos_pos norm_inverse norm_inverse_le_norm zero_less_norm_iff) + then show ?thesis + by (force simp: norm_inverse intro!: boundedI) + qed + qed +qed + + +theorem great_Picard: + assumes "open M" "z \ M" "a \ b" and holf: "f holomorphic_on (M - {z})" + and fab: "\w. w \ M - {z} \ f w \ a \ f w \ b" + obtains l where "(f \ l) (at z) \ ((inverse \ f) \ l) (at z)" +proof - + obtain r where "0 < r" and zrM: "ball z r \ M" + and r: "bounded((\z. f z - a) ` (ball z r - {z})) \ + bounded((inverse \ (\z. f z - a)) ` (ball z r - {z}))" + proof (rule GPicard6 [OF \open M\ \z \ M\]) + show "b - a \ 0" + using assms by auto + show "(\z. f z - a) holomorphic_on M - {z}" + by (intro holomorphic_intros holf) + qed (use fab in auto) + have holfb: "f holomorphic_on ball z r - {z}" + apply (rule holomorphic_on_subset [OF holf]) + using zrM by auto + have holfb_i: "(\z. inverse(f z - a)) holomorphic_on ball z r - {z}" + apply (intro holomorphic_intros holfb) + using fab zrM by fastforce + show ?thesis + using r + proof + assume "bounded ((\z. f z - a) ` (ball z r - {z}))" + then obtain B where B: "\w. w \ (\z. f z - a) ` (ball z r - {z}) \ norm w \ B" + by (force simp: bounded_iff) + have "\\<^sub>F w in at z. cmod (f w - a) \ B" + apply (simp add: eventually_at) + apply (rule_tac x=r in exI) + using \0 < r\ by (auto simp: dist_commute intro!: B) + then have "\B. \\<^sub>F w in at z. cmod (f w) \ B" + apply (rule_tac x="B + norm a" in exI) + apply (erule eventually_mono) + by (metis add.commute add_le_cancel_right norm_triangle_sub order.trans) + then obtain g where holg: "g holomorphic_on ball z r" and gf: "\w. w \ ball z r - {z} \ g w = f w" + using \0 < r\ holomorphic_on_extend_bounded [OF holfb] by auto + then have "g \z\ g z" + apply (simp add: continuous_at [symmetric]) + using \0 < r\ centre_in_ball field_differentiable_imp_continuous_at holomorphic_on_imp_differentiable_at by blast + then have "(f \ g z) (at z)" + apply (rule Lim_transform_within_open [of g "g z" z UNIV "ball z r"]) + using \0 < r\ by (auto simp: gf) + then show ?thesis + using that by blast + next + assume "bounded((inverse \ (\z. f z - a)) ` (ball z r - {z}))" + then obtain B where B: "\w. w \ (inverse \ (\z. f z - a)) ` (ball z r - {z}) \ norm w \ B" + by (force simp: bounded_iff) + have "\\<^sub>F w in at z. cmod (inverse (f w - a)) \ B" + apply (simp add: eventually_at) + apply (rule_tac x=r in exI) + using \0 < r\ by (auto simp: dist_commute intro!: B) + then have "\B. \\<^sub>F z in at z. cmod (inverse (f z - a)) \ B" + by blast + then obtain g where holg: "g holomorphic_on ball z r" and gf: "\w. w \ ball z r - {z} \ g w = inverse (f w - a)" + using \0 < r\ holomorphic_on_extend_bounded [OF holfb_i] by auto + then have gz: "g \z\ g z" + apply (simp add: continuous_at [symmetric]) + using \0 < r\ centre_in_ball field_differentiable_imp_continuous_at holomorphic_on_imp_differentiable_at by blast + have gnz: "\w. w \ ball z r - {z} \ g w \ 0" + using gf fab zrM by fastforce + show ?thesis + proof (cases "g z = 0") + case True + have *: "\g \ 0; inverse g = f - a\ \ g / (1 + a * g) = inverse f" for f g::complex + by (auto simp: field_simps) + have "(inverse \ f) \z\ 0" + proof (rule Lim_transform_within_open [of "\w. g w / (1 + a * g w)" _ _ UNIV "ball z r"]) + show "(\w. g w / (1 + a * g w)) \z\ 0" + using True by (auto simp: intro!: tendsto_eq_intros gz) + show "\x. \x \ ball z r; x \ z\ \ g x / (1 + a * g x) = (inverse \ f) x" + using * gf gnz by simp + qed (use \0 < r\ in auto) + with that show ?thesis by blast + next + case False + show ?thesis + proof (cases "1 + a * g z = 0") + case True + have "(f \ 0) (at z)" + proof (rule Lim_transform_within_open [of "\w. (1 + a * g w) / g w" _ _ _ "ball z r"]) + show "(\w. (1 + a * g w) / g w) \z\ 0" + apply (rule tendsto_eq_intros refl gz \g z \ 0\)+ + by (simp add: True) + show "\x. \x \ ball z r; x \ z\ \ (1 + a * g x) / g x = f x" + using fab fab zrM by (fastforce simp add: gf field_split_simps) + qed (use \0 < r\ in auto) + then show ?thesis + using that by blast + next + case False + have *: "\g \ 0; inverse g = f - a\ \ g / (1 + a * g) = inverse f" for f g::complex + by (auto simp: field_simps) + have "(inverse \ f) \z\ g z / (1 + a * g z)" + proof (rule Lim_transform_within_open [of "\w. g w / (1 + a * g w)" _ _ UNIV "ball z r"]) + show "(\w. g w / (1 + a * g w)) \z\ g z / (1 + a * g z)" + using False by (auto simp: False intro!: tendsto_eq_intros gz) + show "\x. \x \ ball z r; x \ z\ \ g x / (1 + a * g x) = (inverse \ f) x" + using * gf gnz by simp + qed (use \0 < r\ in auto) + with that show ?thesis by blast + qed + qed + qed +qed + + +corollary great_Picard_alt: + assumes M: "open M" "z \ M" and holf: "f holomorphic_on (M - {z})" + and non: "\l. \ (f \ l) (at z)" "\l. \ ((inverse \ f) \ l) (at z)" + obtains a where "- {a} \ f ` (M - {z})" + apply (simp add: subset_iff image_iff) + by (metis great_Picard [OF M _ holf] non) + + +corollary great_Picard_infinite: + assumes M: "open M" "z \ M" and holf: "f holomorphic_on (M - {z})" + and non: "\l. \ (f \ l) (at z)" "\l. \ ((inverse \ f) \ l) (at z)" + obtains a where "\w. w \ a \ infinite {x. x \ M - {z} \ f x = w}" +proof - + have False if "a \ b" and ab: "finite {x. x \ M - {z} \ f x = a}" "finite {x. x \ M - {z} \ f x = b}" for a b + proof - + have finab: "finite {x. x \ M - {z} \ f x \ {a,b}}" + using finite_UnI [OF ab] unfolding mem_Collect_eq insert_iff empty_iff + by (simp add: conj_disj_distribL) + obtain r where "0 < r" and zrM: "ball z r \ M" and r: "\x. \x \ M - {z}; f x \ {a,b}\ \ x \ ball z r" + proof - + obtain e where "e > 0" and e: "ball z e \ M" + using assms openE by blast + show ?thesis + proof (cases "{x \ M - {z}. f x \ {a, b}} = {}") + case True + then show ?thesis + apply (rule_tac r=e in that) + using e \e > 0\ by auto + next + case False + let ?r = "min e (Min (dist z ` {x \ M - {z}. f x \ {a,b}}))" + show ?thesis + proof + show "0 < ?r" + using min_less_iff_conj Min_gr_iff finab False \0 < e\ by auto + have "ball z ?r \ ball z e" + by (simp add: subset_ball) + with e show "ball z ?r \ M" by blast + show "\x. \x \ M - {z}; f x \ {a, b}\ \ x \ ball z ?r" + using min_less_iff_conj Min_gr_iff finab False \0 < e\ by auto + qed + qed + qed + have holfb: "f holomorphic_on (ball z r - {z})" + apply (rule holomorphic_on_subset [OF holf]) + using zrM by auto + show ?thesis + apply (rule great_Picard [OF open_ball _ \a \ b\ holfb]) + using non \0 < r\ r zrM by auto + qed + with that show thesis + by meson +qed + +theorem Casorati_Weierstrass: + assumes "open M" "z \ M" "f holomorphic_on (M - {z})" + and "\l. \ (f \ l) (at z)" "\l. \ ((inverse \ f) \ l) (at z)" + shows "closure(f ` (M - {z})) = UNIV" +proof - + obtain a where a: "- {a} \ f ` (M - {z})" + using great_Picard_alt [OF assms] . + have "UNIV = closure(- {a})" + by (simp add: closure_interior) + also have "... \ closure(f ` (M - {z}))" + by (simp add: a closure_mono) + finally show ?thesis + by blast +qed + +end diff --git a/src/HOL/Complex_Analysis/Riemann_Mapping.thy b/src/HOL/Complex_Analysis/Riemann_Mapping.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/Riemann_Mapping.thy @@ -0,0 +1,1489 @@ +(* Title: HOL/Analysis/Riemann_Mapping.thy + Authors: LC Paulson, based on material from HOL Light +*) + +section \Moebius functions, Equivalents of Simply Connected Sets, Riemann Mapping Theorem\ + +theory Riemann_Mapping +imports Great_Picard +begin + +subsection\Moebius functions are biholomorphisms of the unit disc\ + +definition\<^marker>\tag important\ Moebius_function :: "[real,complex,complex] \ complex" where + "Moebius_function \ \t w z. exp(\ * of_real t) * (z - w) / (1 - cnj w * z)" + +lemma Moebius_function_simple: + "Moebius_function 0 w z = (z - w) / (1 - cnj w * z)" + by (simp add: Moebius_function_def) + +lemma Moebius_function_eq_zero: + "Moebius_function t w w = 0" + by (simp add: Moebius_function_def) + +lemma Moebius_function_of_zero: + "Moebius_function t w 0 = - exp(\ * of_real t) * w" + by (simp add: Moebius_function_def) + +lemma Moebius_function_norm_lt_1: + assumes w1: "norm w < 1" and z1: "norm z < 1" + shows "norm (Moebius_function t w z) < 1" +proof - + have "1 - cnj w * z \ 0" + by (metis complex_cnj_cnj complex_mod_sqrt_Re_mult_cnj mult.commute mult_less_cancel_right1 norm_ge_zero norm_mult norm_one order.asym right_minus_eq w1 z1) + then have VV: "1 - w * cnj z \ 0" + by (metis complex_cnj_cnj complex_cnj_mult complex_cnj_one right_minus_eq) + then have "1 - norm (Moebius_function t w z) ^ 2 = + ((1 - norm w ^ 2) / (norm (1 - cnj w * z) ^ 2)) * (1 - norm z ^ 2)" + apply (cases w) + apply (cases z) + apply (simp add: Moebius_function_def divide_simps norm_divide norm_mult) + apply (simp add: complex_norm complex_diff complex_mult one_complex.code complex_cnj) + apply (auto simp: algebra_simps power2_eq_square) + done + then have "1 - (cmod (Moebius_function t w z))\<^sup>2 = (1 - cmod (w * w)) / (cmod (1 - cnj w * z))\<^sup>2 * (1 - cmod (z * z))" + by (simp add: norm_mult power2_eq_square) + moreover have "0 < 1 - cmod (z * z)" + by (metis (no_types) z1 diff_gt_0_iff_gt mult.left_neutral norm_mult_less) + ultimately have "0 < 1 - norm (Moebius_function t w z) ^ 2" + using \1 - cnj w * z \ 0\ w1 norm_mult_less by fastforce + then show ?thesis + using linorder_not_less by fastforce +qed + +lemma Moebius_function_holomorphic: + assumes "norm w < 1" + shows "Moebius_function t w holomorphic_on ball 0 1" +proof - + have *: "1 - z * w \ 0" if "norm z < 1" for z + proof - + have "norm (1::complex) \ norm (z * w)" + using assms that norm_mult_less by fastforce + then show ?thesis by auto + qed + show ?thesis + apply (simp add: Moebius_function_def) + apply (intro holomorphic_intros) + using assms * + by (metis complex_cnj_cnj complex_cnj_mult complex_cnj_one complex_mod_cnj mem_ball_0 mult.commute right_minus_eq) +qed + +lemma Moebius_function_compose: + assumes meq: "-w1 = w2" and "norm w1 < 1" "norm z < 1" + shows "Moebius_function 0 w1 (Moebius_function 0 w2 z) = z" +proof - + have "norm w2 < 1" + using assms by auto + then have "-w1 = z" if "cnj w2 * z = 1" + by (metis assms(3) complex_mod_cnj less_irrefl mult.right_neutral norm_mult_less norm_one that) + moreover have "z=0" if "1 - cnj w2 * z = cnj w1 * (z - w2)" + proof - + have "w2 * cnj w2 = 1" + using that meq by (auto simp: algebra_simps) + then show "z = 0" + by (metis (no_types) \cmod w2 < 1\ complex_mod_cnj less_irrefl mult.right_neutral norm_mult_less norm_one) + qed + moreover have "z - w2 - w1 * (1 - cnj w2 * z) = z * (1 - cnj w2 * z - cnj w1 * (z - w2))" + using meq by (fastforce simp: algebra_simps) + ultimately + show ?thesis + by (simp add: Moebius_function_def divide_simps norm_divide norm_mult) +qed + +lemma ball_biholomorphism_exists: + assumes "a \ ball 0 1" + obtains f g where "f a = 0" + "f holomorphic_on ball 0 1" "f ` ball 0 1 \ ball 0 1" + "g holomorphic_on ball 0 1" "g ` ball 0 1 \ ball 0 1" + "\z. z \ ball 0 1 \ f (g z) = z" + "\z. z \ ball 0 1 \ g (f z) = z" +proof + show "Moebius_function 0 a holomorphic_on ball 0 1" "Moebius_function 0 (-a) holomorphic_on ball 0 1" + using Moebius_function_holomorphic assms mem_ball_0 by auto + show "Moebius_function 0 a a = 0" + by (simp add: Moebius_function_eq_zero) + show "Moebius_function 0 a ` ball 0 1 \ ball 0 1" + "Moebius_function 0 (- a) ` ball 0 1 \ ball 0 1" + using Moebius_function_norm_lt_1 assms by auto + show "Moebius_function 0 a (Moebius_function 0 (- a) z) = z" + "Moebius_function 0 (- a) (Moebius_function 0 a z) = z" if "z \ ball 0 1" for z + using Moebius_function_compose assms that by auto +qed + + +subsection\A big chain of equivalents of simple connectedness for an open set\ + +lemma biholomorphic_to_disc_aux: + assumes "open S" "connected S" "0 \ S" and S01: "S \ ball 0 1" + and prev: "\f. \f holomorphic_on S; \z. z \ S \ f z \ 0; inj_on f S\ + \ \g. g holomorphic_on S \ (\z \ S. f z = (g z)\<^sup>2)" + shows "\f g. f holomorphic_on S \ g holomorphic_on ball 0 1 \ + (\z \ S. f z \ ball 0 1 \ g(f z) = z) \ + (\z \ ball 0 1. g z \ S \ f(g z) = z)" +proof - + define F where "F \ {h. h holomorphic_on S \ h ` S \ ball 0 1 \ h 0 = 0 \ inj_on h S}" + have idF: "id \ F" + using S01 by (auto simp: F_def) + then have "F \ {}" + by blast + have imF_ne: "((\h. norm(deriv h 0)) ` F) \ {}" + using idF by auto + have holF: "\h. h \ F \ h holomorphic_on S" + by (auto simp: F_def) + obtain f where "f \ F" and normf: "\h. h \ F \ norm(deriv h 0) \ norm(deriv f 0)" + proof - + obtain r where "r > 0" and r: "ball 0 r \ S" + using \open S\ \0 \ S\ openE by auto + have bdd: "bdd_above ((\h. norm(deriv h 0)) ` F)" + proof (intro bdd_aboveI exI ballI, clarify) + show "norm (deriv f 0) \ 1 / r" if "f \ F" for f + proof - + have r01: "(*) (complex_of_real r) ` ball 0 1 \ S" + using that \r > 0\ by (auto simp: norm_mult r [THEN subsetD]) + then have "f holomorphic_on (*) (complex_of_real r) ` ball 0 1" + using holomorphic_on_subset [OF holF] by (simp add: that) + then have holf: "f \ (\z. (r * z)) holomorphic_on (ball 0 1)" + by (intro holomorphic_intros holomorphic_on_compose) + have f0: "(f \ (*) (complex_of_real r)) 0 = 0" + using F_def that by auto + have "f ` S \ ball 0 1" + using F_def that by blast + with r01 have fr1: "\z. norm z < 1 \ norm ((f \ (*)(of_real r))z) < 1" + by force + have *: "((\w. f (r * w)) has_field_derivative deriv f (r * z) * r) (at z)" + if "z \ ball 0 1" for z::complex + proof (rule DERIV_chain' [where g=f]) + show "(f has_field_derivative deriv f (of_real r * z)) (at (of_real r * z))" + apply (rule holomorphic_derivI [OF holF \open S\]) + apply (rule \f \ F\) + by (meson imageI r01 subset_iff that) + qed simp + have df0: "((\w. f (r * w)) has_field_derivative deriv f 0 * r) (at 0)" + using * [of 0] by simp + have deq: "deriv (\x. f (complex_of_real r * x)) 0 = deriv f 0 * complex_of_real r" + using DERIV_imp_deriv df0 by blast + have "norm (deriv (f \ (*) (complex_of_real r)) 0) \ 1" + by (auto intro: Schwarz_Lemma [OF holf f0 fr1, of 0]) + with \r > 0\ show ?thesis + by (simp add: deq norm_mult divide_simps o_def) + qed + qed + define l where "l \ SUP h\F. norm (deriv h 0)" + have eql: "norm (deriv f 0) = l" if le: "l \ norm (deriv f 0)" and "f \ F" for f + apply (rule order_antisym [OF _ le]) + using \f \ F\ bdd cSUP_upper by (fastforce simp: l_def) + obtain \ where \in: "\n. \ n \ F" and \lim: "(\n. norm (deriv (\ n) 0)) \ l" + proof - + have "\f. f \ F \ \norm (deriv f 0) - l\ < 1 / (Suc n)" for n + proof - + obtain f where "f \ F" and f: "l < norm (deriv f 0) + 1/(Suc n)" + using cSup_least [OF imF_ne, of "l - 1/(Suc n)"] by (fastforce simp add: l_def) + then have "\norm (deriv f 0) - l\ < 1 / (Suc n)" + by (fastforce simp add: abs_if not_less eql) + with \f \ F\ show ?thesis + by blast + qed + then obtain \ where fF: "\n. (\ n) \ F" + and fless: "\n. \norm (deriv (\ n) 0) - l\ < 1 / (Suc n)" + by metis + have "(\n. norm (deriv (\ n) 0)) \ l" + proof (rule metric_LIMSEQ_I) + fix e::real + assume "e > 0" + then obtain N::nat where N: "e > 1/(Suc N)" + using nat_approx_posE by blast + show "\N. \n\N. dist (norm (deriv (\ n) 0)) l < e" + proof (intro exI allI impI) + fix n assume "N \ n" + have "dist (norm (deriv (\ n) 0)) l < 1 / (Suc n)" + using fless by (simp add: dist_norm) + also have "... < e" + using N \N \ n\ inverse_of_nat_le le_less_trans by blast + finally show "dist (norm (deriv (\ n) 0)) l < e" . + qed + qed + with fF show ?thesis + using that by blast + qed + have "\K. \compact K; K \ S\ \ \B. \h\F. \z\K. norm (h z) \ B" + by (rule_tac x=1 in exI) (force simp: F_def) + moreover have "range \ \ F" + using \\n. \ n \ F\ by blast + ultimately obtain f and r :: "nat \ nat" + where holf: "f holomorphic_on S" and r: "strict_mono r" + and limf: "\x. x \ S \ (\n. \ (r n) x) \ f x" + and ulimf: "\K. \compact K; K \ S\ \ uniform_limit K (\ \ r) f sequentially" + using Montel [of S F \, OF \open S\ holF] by auto+ + have der: "\n x. x \ S \ ((\ \ r) n has_field_derivative ((\n. deriv (\ n)) \ r) n x) (at x)" + using \\n. \ n \ F\ \open S\ holF holomorphic_derivI by fastforce + have ulim: "\x. x \ S \ \d>0. cball x d \ S \ uniform_limit (cball x d) (\ \ r) f sequentially" + by (meson ulimf \open S\ compact_cball open_contains_cball) + obtain f' :: "complex\complex" where f': "(f has_field_derivative f' 0) (at 0)" + and tof'0: "(\n. ((\n. deriv (\ n)) \ r) n 0) \ f' 0" + using has_complex_derivative_uniform_sequence [OF \open S\ der ulim] \0 \ S\ by metis + then have derf0: "deriv f 0 = f' 0" + by (simp add: DERIV_imp_deriv) + have "f field_differentiable (at 0)" + using field_differentiable_def f' by blast + have "(\x. (norm (deriv (\ (r x)) 0))) \ norm (deriv f 0)" + using isCont_tendsto_compose [OF continuous_norm [OF continuous_ident] tof'0] derf0 by auto + with LIMSEQ_subseq_LIMSEQ [OF \lim r] have no_df0: "norm(deriv f 0) = l" + by (force simp: o_def intro: tendsto_unique) + have nonconstf: "\ f constant_on S" + proof - + have False if "\x. x \ S \ f x = c" for c + proof - + have "deriv f 0 = 0" + by (metis that \open S\ \0 \ S\ DERIV_imp_deriv [OF has_field_derivative_transform_within_open [OF DERIV_const]]) + with no_df0 have "l = 0" + by auto + with eql [OF _ idF] show False by auto + qed + then show ?thesis + by (meson constant_on_def) + qed + show ?thesis + proof + show "f \ F" + unfolding F_def + proof (intro CollectI conjI holf) + have "norm(f z) \ 1" if "z \ S" for z + proof (intro Lim_norm_ubound [OF _ limf] always_eventually allI that) + fix n + have "\ (r n) \ F" + by (simp add: \in) + then show "norm (\ (r n) z) \ 1" + using that by (auto simp: F_def) + qed simp + then have fless1: "norm(f z) < 1" if "z \ S" for z + using maximum_modulus_principle [OF holf \open S\ \connected S\ \open S\] nonconstf that + by fastforce + then show "f ` S \ ball 0 1" + by auto + have "(\n. \ (r n) 0) \ 0" + using \in by (auto simp: F_def) + then show "f 0 = 0" + using tendsto_unique [OF _ limf ] \0 \ S\ trivial_limit_sequentially by blast + show "inj_on f S" + proof (rule Hurwitz_injective [OF \open S\ \connected S\ _ holf]) + show "\n. (\ \ r) n holomorphic_on S" + by (simp add: \in holF) + show "\K. \compact K; K \ S\ \ uniform_limit K(\ \ r) f sequentially" + by (metis ulimf) + show "\ f constant_on S" + using nonconstf by auto + show "\n. inj_on ((\ \ r) n) S" + using \in by (auto simp: F_def) + qed + qed + show "\h. h \ F \ norm (deriv h 0) \ norm (deriv f 0)" + by (metis eql le_cases no_df0) + qed + qed + have holf: "f holomorphic_on S" and injf: "inj_on f S" and f01: "f ` S \ ball 0 1" + using \f \ F\ by (auto simp: F_def) + obtain g where holg: "g holomorphic_on (f ` S)" + and derg: "\z. z \ S \ deriv f z * deriv g (f z) = 1" + and gf: "\z. z \ S \ g(f z) = z" + using holomorphic_has_inverse [OF holf \open S\ injf] by metis + have "ball 0 1 \ f ` S" + proof + fix a::complex + assume a: "a \ ball 0 1" + have False if "\x. x \ S \ f x \ a" + proof - + obtain h k where "h a = 0" + and holh: "h holomorphic_on ball 0 1" and h01: "h ` ball 0 1 \ ball 0 1" + and holk: "k holomorphic_on ball 0 1" and k01: "k ` ball 0 1 \ ball 0 1" + and hk: "\z. z \ ball 0 1 \ h (k z) = z" + and kh: "\z. z \ ball 0 1 \ k (h z) = z" + using ball_biholomorphism_exists [OF a] by blast + have nf1: "\z. z \ S \ norm(f z) < 1" + using \f \ F\ by (auto simp: F_def) + have 1: "h \ f holomorphic_on S" + using F_def \f \ F\ holh holomorphic_on_compose holomorphic_on_subset by blast + have 2: "\z. z \ S \ (h \ f) z \ 0" + by (metis \h a = 0\ a comp_eq_dest_lhs nf1 kh mem_ball_0 that) + have 3: "inj_on (h \ f) S" + by (metis (no_types, lifting) F_def \f \ F\ comp_inj_on inj_on_inverseI injf kh mem_Collect_eq subset_inj_on) + obtain \ where hol\: "\ holomorphic_on ((h \ f) ` S)" + and \2: "\z. z \ S \ \(h (f z)) ^ 2 = h(f z)" + proof (rule exE [OF prev [OF 1 2 3]], safe) + fix \ + assume hol\: "\ holomorphic_on S" and \2: "(\z\S. (h \ f) z = (\ z)\<^sup>2)" + show thesis + proof + show "(\ \ g \ k) holomorphic_on (h \ f) ` S" + proof (intro holomorphic_on_compose) + show "k holomorphic_on (h \ f) ` S" + apply (rule holomorphic_on_subset [OF holk]) + using f01 h01 by force + show "g holomorphic_on k ` (h \ f) ` S" + apply (rule holomorphic_on_subset [OF holg]) + by (auto simp: kh nf1) + show "\ holomorphic_on g ` k ` (h \ f) ` S" + apply (rule holomorphic_on_subset [OF hol\]) + by (auto simp: gf kh nf1) + qed + show "((\ \ g \ k) (h (f z)))\<^sup>2 = h (f z)" if "z \ S" for z + proof - + have "f z \ ball 0 1" + by (simp add: nf1 that) + then have "(\ (g (k (h (f z)))))\<^sup>2 = (\ (g (f z)))\<^sup>2" + by (metis kh) + also have "... = h (f z)" + using \2 gf that by auto + finally show ?thesis + by (simp add: o_def) + qed + qed + qed + have norm\1: "norm(\ (h (f z))) < 1" if "z \ S" for z + proof - + have "norm (\ (h (f z)) ^ 2) < 1" + by (metis (no_types) that DIM_complex \2 h01 image_subset_iff mem_ball_0 nf1) + then show ?thesis + by (metis le_less_trans mult_less_cancel_left2 norm_ge_zero norm_power not_le power2_eq_square) + qed + then have \01: "\ (h (f 0)) \ ball 0 1" + by (simp add: \0 \ S\) + obtain p q where p0: "p (\ (h (f 0))) = 0" + and holp: "p holomorphic_on ball 0 1" and p01: "p ` ball 0 1 \ ball 0 1" + and holq: "q holomorphic_on ball 0 1" and q01: "q ` ball 0 1 \ ball 0 1" + and pq: "\z. z \ ball 0 1 \ p (q z) = z" + and qp: "\z. z \ ball 0 1 \ q (p z) = z" + using ball_biholomorphism_exists [OF \01] by metis + have "p \ \ \ h \ f \ F" + unfolding F_def + proof (intro CollectI conjI holf) + show "p \ \ \ h \ f holomorphic_on S" + proof (intro holomorphic_on_compose holf) + show "h holomorphic_on f ` S" + apply (rule holomorphic_on_subset [OF holh]) + using f01 by force + show "\ holomorphic_on h ` f ` S" + apply (rule holomorphic_on_subset [OF hol\]) + by auto + show "p holomorphic_on \ ` h ` f ` S" + apply (rule holomorphic_on_subset [OF holp]) + by (auto simp: norm\1) + qed + show "(p \ \ \ h \ f) ` S \ ball 0 1" + apply clarsimp + by (meson norm\1 p01 image_subset_iff mem_ball_0) + show "(p \ \ \ h \ f) 0 = 0" + by (simp add: \p (\ (h (f 0))) = 0\) + show "inj_on (p \ \ \ h \ f) S" + unfolding inj_on_def o_def + by (metis \2 dist_0_norm gf kh mem_ball nf1 norm\1 qp) + qed + then have le_norm_df0: "norm (deriv (p \ \ \ h \ f) 0) \ norm (deriv f 0)" + by (rule normf) + have 1: "k \ power2 \ q holomorphic_on ball 0 1" + proof (intro holomorphic_on_compose holq) + show "power2 holomorphic_on q ` ball 0 1" + using holomorphic_on_subset holomorphic_on_power + by (blast intro: holomorphic_on_ident) + show "k holomorphic_on power2 ` q ` ball 0 1" + apply (rule holomorphic_on_subset [OF holk]) + using q01 by (auto simp: norm_power abs_square_less_1) + qed + have 2: "(k \ power2 \ q) 0 = 0" + using p0 F_def \f \ F\ \01 \2 \0 \ S\ kh qp by force + have 3: "norm ((k \ power2 \ q) z) < 1" if "norm z < 1" for z + proof - + have "norm ((power2 \ q) z) < 1" + using that q01 by (force simp: norm_power abs_square_less_1) + with k01 show ?thesis + by fastforce + qed + have False if c: "\z. norm z < 1 \ (k \ power2 \ q) z = c * z" and "norm c = 1" for c + proof - + have "c \ 0" using that by auto + have "norm (p(1/2)) < 1" "norm (p(-1/2)) < 1" + using p01 by force+ + then have "(k \ power2 \ q) (p(1/2)) = c * p(1/2)" "(k \ power2 \ q) (p(-1/2)) = c * p(-1/2)" + using c by force+ + then have "p (1/2) = p (- (1/2))" + by (auto simp: \c \ 0\ qp o_def) + then have "q (p (1/2)) = q (p (- (1/2)))" + by simp + then have "1/2 = - (1/2::complex)" + by (auto simp: qp) + then show False + by simp + qed + moreover + have False if "norm (deriv (k \ power2 \ q) 0) \ 1" "norm (deriv (k \ power2 \ q) 0) \ 1" + and le: "\\. norm \ < 1 \ norm ((k \ power2 \ q) \) \ norm \" + proof - + have "norm (deriv (k \ power2 \ q) 0) < 1" + using that by simp + moreover have eq: "deriv f 0 = deriv (k \ (\z. z ^ 2) \ q) 0 * deriv (p \ \ \ h \ f) 0" + proof (intro DERIV_imp_deriv has_field_derivative_transform_within_open [OF DERIV_chain]) + show "(k \ power2 \ q has_field_derivative deriv (k \ power2 \ q) 0) (at ((p \ \ \ h \ f) 0))" + using "1" holomorphic_derivI p0 by auto + show "(p \ \ \ h \ f has_field_derivative deriv (p \ \ \ h \ f) 0) (at 0)" + using \p \ \ \ h \ f \ F\ \open S\ \0 \ S\ holF holomorphic_derivI by blast + show "\x. x \ S \ (k \ power2 \ q \ (p \ \ \ h \ f)) x = f x" + using \2 f01 kh norm\1 qp by auto + qed (use assms in simp_all) + ultimately have "cmod (deriv (p \ \ \ h \ f) 0) \ 0" + using le_norm_df0 + by (metis linorder_not_le mult.commute mult_less_cancel_left2 norm_mult) + moreover have "1 \ norm (deriv f 0)" + using normf [of id] by (simp add: idF) + ultimately show False + by (simp add: eq) + qed + ultimately show ?thesis + using Schwarz_Lemma [OF 1 2 3] norm_one by blast + qed + then show "a \ f ` S" + by blast + qed + then have "f ` S = ball 0 1" + using F_def \f \ F\ by blast + then show ?thesis + apply (rule_tac x=f in exI) + apply (rule_tac x=g in exI) + using holf holg derg gf by safe force+ +qed + + +locale SC_Chain = + fixes S :: "complex set" + assumes openS: "open S" +begin + +lemma winding_number_zero: + assumes "simply_connected S" + shows "connected S \ + (\\ z. path \ \ path_image \ \ S \ + pathfinish \ = pathstart \ \ z \ S \ winding_number \ z = 0)" + using assms + by (auto simp: simply_connected_imp_connected simply_connected_imp_winding_number_zero) + +lemma contour_integral_zero: + assumes "valid_path g" "path_image g \ S" "pathfinish g = pathstart g" "f holomorphic_on S" + "\\ z. \path \; path_image \ \ S; pathfinish \ = pathstart \; z \ S\ \ winding_number \ z = 0" + shows "(f has_contour_integral 0) g" + using assms by (meson Cauchy_theorem_global openS valid_path_imp_path) + +lemma global_primitive: + assumes "connected S" and holf: "f holomorphic_on S" + and prev: "\\ f. \valid_path \; path_image \ \ S; pathfinish \ = pathstart \; f holomorphic_on S\ \ (f has_contour_integral 0) \" + shows "\h. \z \ S. (h has_field_derivative f z) (at z)" +proof (cases "S = {}") +case True then show ?thesis + by simp +next + case False + then obtain a where "a \ S" + by blast + show ?thesis + proof (intro exI ballI) + fix x assume "x \ S" + then obtain d where "d > 0" and d: "cball x d \ S" + using openS open_contains_cball_eq by blast + let ?g = "\z. (SOME g. polynomial_function g \ path_image g \ S \ pathstart g = a \ pathfinish g = z)" + show "((\z. contour_integral (?g z) f) has_field_derivative f x) + (at x)" + proof (simp add: has_field_derivative_def has_derivative_at2 bounded_linear_mult_right, rule Lim_transform) + show "(\y. inverse(norm(y - x)) *\<^sub>R (contour_integral(linepath x y) f - f x * (y - x))) \x\ 0" + proof (clarsimp simp add: Lim_at) + fix e::real assume "e > 0" + moreover have "continuous (at x) f" + using openS \x \ S\ holf continuous_on_eq_continuous_at holomorphic_on_imp_continuous_on by auto + ultimately obtain d1 where "d1 > 0" + and d1: "\x'. dist x' x < d1 \ dist (f x') (f x) < e/2" + unfolding continuous_at_eps_delta + by (metis less_divide_eq_numeral1(1) mult_zero_left) + obtain d2 where "d2 > 0" and d2: "ball x d2 \ S" + using openS \x \ S\ open_contains_ball_eq by blast + have "inverse (norm (y - x)) * norm (contour_integral (linepath x y) f - f x * (y - x)) < e" + if "0 < d1" "0 < d2" "y \ x" "dist y x < d1" "dist y x < d2" for y + proof - + have "f contour_integrable_on linepath x y" + proof (rule contour_integrable_continuous_linepath [OF continuous_on_subset]) + show "continuous_on S f" + by (simp add: holf holomorphic_on_imp_continuous_on) + have "closed_segment x y \ ball x d2" + by (meson dist_commute_lessI dist_in_closed_segment le_less_trans mem_ball subsetI that(5)) + with d2 show "closed_segment x y \ S" + by blast + qed + then obtain z where z: "(f has_contour_integral z) (linepath x y)" + by (force simp: contour_integrable_on_def) + have con: "((\w. f x) has_contour_integral f x * (y - x)) (linepath x y)" + using has_contour_integral_const_linepath [of "f x" y x] by metis + have "norm (z - f x * (y - x)) \ (e/2) * norm (y - x)" + proof (rule has_contour_integral_bound_linepath) + show "((\w. f w - f x) has_contour_integral z - f x * (y - x)) (linepath x y)" + by (rule has_contour_integral_diff [OF z con]) + show "\w. w \ closed_segment x y \ norm (f w - f x) \ e/2" + by (metis d1 dist_norm less_le_trans not_less not_less_iff_gr_or_eq segment_bound1 that(4)) + qed (use \e > 0\ in auto) + with \e > 0\ have "inverse (norm (y - x)) * norm (z - f x * (y - x)) \ e/2" + by (simp add: field_split_simps) + also have "... < e" + using \e > 0\ by simp + finally show ?thesis + by (simp add: contour_integral_unique [OF z]) + qed + with \d1 > 0\ \d2 > 0\ + show "\d>0. \z. z \ x \ dist z x < d \ + inverse (norm (z - x)) * norm (contour_integral (linepath x z) f - f x * (z - x)) < e" + by (rule_tac x="min d1 d2" in exI) auto + qed + next + have *: "(1 / norm (y - x)) *\<^sub>R (contour_integral (?g y) f - + (contour_integral (?g x) f + f x * (y - x))) = + (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R norm (y - x)" + if "0 < d" "y \ x" and yx: "dist y x < d" for y + proof - + have "y \ S" + by (metis subsetD d dist_commute less_eq_real_def mem_cball yx) + have gxy: "polynomial_function (?g x) \ path_image (?g x) \ S \ pathstart (?g x) = a \ pathfinish (?g x) = x" + "polynomial_function (?g y) \ path_image (?g y) \ S \ pathstart (?g y) = a \ pathfinish (?g y) = y" + using someI_ex [OF connected_open_polynomial_connected [OF openS \connected S\ \a \ S\]] \x \ S\ \y \ S\ + by meson+ + then have vp: "valid_path (?g x)" "valid_path (?g y)" + by (simp_all add: valid_path_polynomial_function) + have f0: "(f has_contour_integral 0) ((?g x) +++ linepath x y +++ reversepath (?g y))" + proof (rule prev) + show "valid_path ((?g x) +++ linepath x y +++ reversepath (?g y))" + using gxy vp by (auto simp: valid_path_join) + have "closed_segment x y \ cball x d" + using yx by (auto simp: dist_commute dest!: dist_in_closed_segment) + then have "closed_segment x y \ S" + using d by blast + then show "path_image ((?g x) +++ linepath x y +++ reversepath (?g y)) \ S" + using gxy by (auto simp: path_image_join) + qed (use gxy holf in auto) + then have fintxy: "f contour_integrable_on linepath x y" + by (metis (no_types, lifting) contour_integrable_joinD1 contour_integrable_joinD2 gxy(2) has_contour_integral_integrable pathfinish_linepath pathstart_reversepath valid_path_imp_reverse valid_path_join valid_path_linepath vp(2)) + have fintgx: "f contour_integrable_on (?g x)" "f contour_integrable_on (?g y)" + using openS contour_integrable_holomorphic_simple gxy holf vp by blast+ + show ?thesis + apply (clarsimp simp add: divide_simps) + using contour_integral_unique [OF f0] + apply (simp add: fintxy gxy contour_integrable_reversepath contour_integral_reversepath fintgx vp) + apply (simp add: algebra_simps) + done + qed + show "(\z. (1 / norm (z - x)) *\<^sub>R + (contour_integral (?g z) f - (contour_integral (?g x) f + f x * (z - x))) - + (contour_integral (linepath x z) f - f x * (z - x)) /\<^sub>R norm (z - x)) + \x\ 0" + apply (rule tendsto_eventually) + apply (simp add: eventually_at) + apply (rule_tac x=d in exI) + using \d > 0\ * by simp + qed + qed +qed + +lemma holomorphic_log: + assumes "connected S" and holf: "f holomorphic_on S" and nz: "\z. z \ S \ f z \ 0" + and prev: "\f. f holomorphic_on S \ \h. \z \ S. (h has_field_derivative f z) (at z)" + shows "\g. g holomorphic_on S \ (\z \ S. f z = exp(g z))" +proof - + have "(\z. deriv f z / f z) holomorphic_on S" + by (simp add: openS holf holomorphic_deriv holomorphic_on_divide nz) + then obtain g where g: "\z. z \ S \ (g has_field_derivative deriv f z / f z) (at z)" + using prev [of "\z. deriv f z / f z"] by metis + have hfd: "\x. x \ S \ ((\z. exp (g z) / f z) has_field_derivative 0) (at x)" + apply (rule derivative_eq_intros g| simp)+ + apply (subst DERIV_deriv_iff_field_differentiable) + using openS holf holomorphic_on_imp_differentiable_at nz apply auto + done + obtain c where c: "\x. x \ S \ exp (g x) / f x = c" + proof (rule DERIV_zero_connected_constant[OF \connected S\ openS finite.emptyI]) + show "continuous_on S (\z. exp (g z) / f z)" + by (metis (full_types) openS g continuous_on_divide continuous_on_exp holf holomorphic_on_imp_continuous_on holomorphic_on_open nz) + then show "\x\S - {}. ((\z. exp (g z) / f z) has_field_derivative 0) (at x)" + using hfd by (blast intro: DERIV_zero_connected_constant [OF \connected S\ openS finite.emptyI, of "\z. exp(g z) / f z"]) + qed auto + show ?thesis + proof (intro exI ballI conjI) + show "(\z. Ln(inverse c) + g z) holomorphic_on S" + apply (intro holomorphic_intros) + using openS g holomorphic_on_open by blast + fix z :: complex + assume "z \ S" + then have "exp (g z) / c = f z" + by (metis c divide_divide_eq_right exp_not_eq_zero nonzero_mult_div_cancel_left) + moreover have "1 / c \ 0" + using \z \ S\ c nz by fastforce + ultimately show "f z = exp (Ln (inverse c) + g z)" + by (simp add: exp_add inverse_eq_divide) + qed +qed + +lemma holomorphic_sqrt: + assumes holf: "f holomorphic_on S" and nz: "\z. z \ S \ f z \ 0" + and prev: "\f. \f holomorphic_on S; \z \ S. f z \ 0\ \ \g. g holomorphic_on S \ (\z \ S. f z = exp(g z))" + shows "\g. g holomorphic_on S \ (\z \ S. f z = (g z)\<^sup>2)" +proof - + obtain g where holg: "g holomorphic_on S" and g: "\z. z \ S \ f z = exp (g z)" + using prev [of f] holf nz by metis + show ?thesis + proof (intro exI ballI conjI) + show "(\z. exp(g z/2)) holomorphic_on S" + by (intro holomorphic_intros) (auto simp: holg) + show "\z. z \ S \ f z = (exp (g z/2))\<^sup>2" + by (metis (no_types) g exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) + qed +qed + +lemma biholomorphic_to_disc: + assumes "connected S" and S: "S \ {}" "S \ UNIV" + and prev: "\f. \f holomorphic_on S; \z \ S. f z \ 0\ \ \g. g holomorphic_on S \ (\z \ S. f z = (g z)\<^sup>2)" + shows "\f g. f holomorphic_on S \ g holomorphic_on ball 0 1 \ + (\z \ S. f z \ ball 0 1 \ g(f z) = z) \ + (\z \ ball 0 1. g z \ S \ f(g z) = z)" +proof - + obtain a b where "a \ S" "b \ S" + using S by blast + then obtain \ where "\ > 0" and \: "ball a \ \ S" + using openS openE by blast + obtain g where holg: "g holomorphic_on S" and eqg: "\z. z \ S \ z - b = (g z)\<^sup>2" + proof (rule exE [OF prev [of "\z. z - b"]]) + show "(\z. z - b) holomorphic_on S" + by (intro holomorphic_intros) + qed (use \b \ S\ in auto) + have "\ g constant_on S" + proof - + have "(a + \/2) \ ball a \" "a + (\/2) \ a" + using \\ > 0\ by (simp_all add: dist_norm) + then show ?thesis + unfolding constant_on_def + using eqg [of a] eqg [of "a + \/2"] \a \ S\ \ + by (metis diff_add_cancel subset_eq) + qed + then have "open (g ` ball a \)" + using open_mapping_thm [of g S "ball a \", OF holg openS \connected S\] \ by blast + then obtain r where "r > 0" and r: "ball (g a) r \ (g ` ball a \)" + by (metis \0 < \\ centre_in_ball imageI openE) + have g_not_r: "g z \ ball (-(g a)) r" if "z \ S" for z + proof + assume "g z \ ball (-(g a)) r" + then have "- g z \ ball (g a) r" + by (metis add.inverse_inverse dist_minus mem_ball) + with r have "- g z \ (g ` ball a \)" + by blast + then obtain w where w: "- g z = g w" "dist a w < \" + by auto + then have "w \ ball a \" + by simp + then have "w \ S" + using \ by blast + then have "w = z" + by (metis diff_add_cancel eqg power_minus_Bit0 that w(1)) + then have "g z = 0" + using \- g z = g w\ by auto + with eqg [OF that] have "z = b" + by auto + with that \b \ S\ show False + by simp + qed + then have nz: "\z. z \ S \ g z + g a \ 0" + by (metis \0 < r\ add.commute add_diff_cancel_left' centre_in_ball diff_0) + let ?f = "\z. (r/3) / (g z + g a) - (r/3) / (g a + g a)" + obtain h where holh: "h holomorphic_on S" and "h a = 0" and h01: "h ` S \ ball 0 1" and "inj_on h S" + proof + show "?f holomorphic_on S" + by (intro holomorphic_intros holg nz) + have 3: "\norm x \ 1/3; norm y \ 1/3\ \ norm(x - y) < 1" for x y::complex + using norm_triangle_ineq4 [of x y] by simp + have "norm ((r/3) / (g z + g a) - (r/3) / (g a + g a)) < 1" if "z \ S" for z + apply (rule 3) + unfolding norm_divide + using \r > 0\ g_not_r [OF \z \ S\] g_not_r [OF \a \ S\] + by (simp_all add: field_split_simps dist_commute dist_norm) + then show "?f ` S \ ball 0 1" + by auto + show "inj_on ?f S" + using \r > 0\ eqg apply (clarsimp simp: inj_on_def) + by (metis diff_add_cancel) + qed auto + obtain k where holk: "k holomorphic_on (h ` S)" + and derk: "\z. z \ S \ deriv h z * deriv k (h z) = 1" + and kh: "\z. z \ S \ k(h z) = z" + using holomorphic_has_inverse [OF holh openS \inj_on h S\] by metis + + have 1: "open (h ` S)" + by (simp add: \inj_on h S\ holh openS open_mapping_thm3) + have 2: "connected (h ` S)" + by (simp add: connected_continuous_image \connected S\ holh holomorphic_on_imp_continuous_on) + have 3: "0 \ h ` S" + using \a \ S\ \h a = 0\ by auto + have 4: "\g. g holomorphic_on h ` S \ (\z\h ` S. f z = (g z)\<^sup>2)" + if holf: "f holomorphic_on h ` S" and nz: "\z. z \ h ` S \ f z \ 0" "inj_on f (h ` S)" for f + proof - + obtain g where holg: "g holomorphic_on S" and eqg: "\z. z \ S \ (f \ h) z = (g z)\<^sup>2" + proof - + have "f \ h holomorphic_on S" + by (simp add: holh holomorphic_on_compose holf) + moreover have "\z\S. (f \ h) z \ 0" + by (simp add: nz) + ultimately show thesis + using prev that by blast + qed + show ?thesis + proof (intro exI conjI) + show "g \ k holomorphic_on h ` S" + proof - + have "k ` h ` S \ S" + by (simp add: \\z. z \ S \ k (h z) = z\ image_subset_iff) + then show ?thesis + by (meson holg holk holomorphic_on_compose holomorphic_on_subset) + qed + show "\z\h ` S. f z = ((g \ k) z)\<^sup>2" + using eqg kh by auto + qed + qed + obtain f g where f: "f holomorphic_on h ` S" and g: "g holomorphic_on ball 0 1" + and gf: "\z\h ` S. f z \ ball 0 1 \ g (f z) = z" and fg:"\z\ball 0 1. g z \ h ` S \ f (g z) = z" + using biholomorphic_to_disc_aux [OF 1 2 3 h01 4] by blast + show ?thesis + proof (intro exI conjI) + show "f \ h holomorphic_on S" + by (simp add: f holh holomorphic_on_compose) + show "k \ g holomorphic_on ball 0 1" + by (metis holomorphic_on_subset image_subset_iff fg holk g holomorphic_on_compose) + qed (use fg gf kh in auto) +qed + +lemma homeomorphic_to_disc: + assumes S: "S \ {}" + and prev: "S = UNIV \ + (\f g. f holomorphic_on S \ g holomorphic_on ball 0 1 \ + (\z \ S. f z \ ball 0 1 \ g(f z) = z) \ + (\z \ ball 0 1. g z \ S \ f(g z) = z))" (is "_ \ ?P") + shows "S homeomorphic ball (0::complex) 1" + using prev +proof + assume "S = UNIV" then show ?thesis + using homeomorphic_ball01_UNIV homeomorphic_sym by blast +next + assume ?P + then show ?thesis + unfolding homeomorphic_minimal + using holomorphic_on_imp_continuous_on by blast +qed + +lemma homeomorphic_to_disc_imp_simply_connected: + assumes "S = {} \ S homeomorphic ball (0::complex) 1" + shows "simply_connected S" + using assms homeomorphic_simply_connected_eq convex_imp_simply_connected by auto + +end + +proposition + assumes "open S" + shows simply_connected_eq_winding_number_zero: + "simply_connected S \ + connected S \ + (\g z. path g \ path_image g \ S \ + pathfinish g = pathstart g \ (z \ S) + \ winding_number g z = 0)" (is "?wn0") + and simply_connected_eq_contour_integral_zero: + "simply_connected S \ + connected S \ + (\g f. valid_path g \ path_image g \ S \ + pathfinish g = pathstart g \ f holomorphic_on S + \ (f has_contour_integral 0) g)" (is "?ci0") + and simply_connected_eq_global_primitive: + "simply_connected S \ + connected S \ + (\f. f holomorphic_on S \ + (\h. \z. z \ S \ (h has_field_derivative f z) (at z)))" (is "?gp") + and simply_connected_eq_holomorphic_log: + "simply_connected S \ + connected S \ + (\f. f holomorphic_on S \ (\z \ S. f z \ 0) + \ (\g. g holomorphic_on S \ (\z \ S. f z = exp(g z))))" (is "?log") + and simply_connected_eq_holomorphic_sqrt: + "simply_connected S \ + connected S \ + (\f. f holomorphic_on S \ (\z \ S. f z \ 0) + \ (\g. g holomorphic_on S \ (\z \ S. f z = (g z)\<^sup>2)))" (is "?sqrt") + and simply_connected_eq_biholomorphic_to_disc: + "simply_connected S \ + S = {} \ S = UNIV \ + (\f g. f holomorphic_on S \ g holomorphic_on ball 0 1 \ + (\z \ S. f z \ ball 0 1 \ g(f z) = z) \ + (\z \ ball 0 1. g z \ S \ f(g z) = z))" (is "?bih") + and simply_connected_eq_homeomorphic_to_disc: + "simply_connected S \ S = {} \ S homeomorphic ball (0::complex) 1" (is "?disc") +proof - + interpret SC_Chain + using assms by (simp add: SC_Chain_def) + have "?wn0 \ ?ci0 \ ?gp \ ?log \ ?sqrt \ ?bih \ ?disc" +proof - + have *: "\\ \ \; \ \ \; \ \ \; \ \ \; \ \ \; \ \ \; \ \ \; \ \ \\ + \ (\ \ \) \ (\ \ \) \ (\ \ \) \ (\ \ \) \ + (\ \ \) \ (\ \ \) \ (\ \ \)" for \ \ \ \ \ \ \ \ + by blast + show ?thesis + apply (rule *) + using winding_number_zero apply metis + using contour_integral_zero apply metis + using global_primitive apply metis + using holomorphic_log apply metis + using holomorphic_sqrt apply simp + using biholomorphic_to_disc apply blast + using homeomorphic_to_disc apply blast + using homeomorphic_to_disc_imp_simply_connected apply blast + done +qed + then show ?wn0 ?ci0 ?gp ?log ?sqrt ?bih ?disc + by safe +qed + +corollary contractible_eq_simply_connected_2d: + fixes S :: "complex set" + shows "open S \ (contractible S \ simply_connected S)" + apply safe + apply (simp add: contractible_imp_simply_connected) + using convex_imp_contractible homeomorphic_contractible_eq simply_connected_eq_homeomorphic_to_disc by auto + +subsection\A further chain of equivalences about components of the complement of a simply connected set\ + +text\(following 1.35 in Burckel'S book)\ + +context SC_Chain +begin + +lemma frontier_properties: + assumes "simply_connected S" + shows "if bounded S then connected(frontier S) + else \C \ components(frontier S). \ bounded C" +proof - + have "S = {} \ S homeomorphic ball (0::complex) 1" + using simply_connected_eq_homeomorphic_to_disc assms openS by blast + then show ?thesis + proof + assume "S = {}" + then have "bounded S" + by simp + with \S = {}\ show ?thesis + by simp + next + assume S01: "S homeomorphic ball (0::complex) 1" + then obtain g f + where gim: "g ` S = ball 0 1" and fg: "\x. x \ S \ f(g x) = x" + and fim: "f ` ball 0 1 = S" and gf: "\y. cmod y < 1 \ g(f y) = y" + and contg: "continuous_on S g" and contf: "continuous_on (ball 0 1) f" + by (fastforce simp: homeomorphism_def homeomorphic_def) + define D where "D \ \n. ball (0::complex) (1 - 1/(of_nat n + 2))" + define A where "A \ \n. {z::complex. 1 - 1/(of_nat n + 2) < norm z \ norm z < 1}" + define X where "X \ \n::nat. closure(f ` A n)" + have D01: "D n \ ball 0 1" for n + by (simp add: D_def ball_subset_ball_iff) + have A01: "A n \ ball 0 1" for n + by (auto simp: A_def) + have cloX: "closed(X n)" for n + by (simp add: X_def) + have Xsubclo: "X n \ closure S" for n + unfolding X_def by (metis A01 closure_mono fim image_mono) + have connX: "connected(X n)" for n + unfolding X_def + apply (rule connected_imp_connected_closure) + apply (rule connected_continuous_image) + apply (simp add: continuous_on_subset [OF contf A01]) + using connected_annulus [of _ "0::complex"] by (simp add: A_def) + have nestX: "X n \ X m" if "m \ n" for m n + proof - + have "1 - 1 / (real m + 2) \ 1 - 1 / (real n + 2)" + using that by (auto simp: field_simps) + then show ?thesis + by (auto simp: X_def A_def intro!: closure_mono) + qed + have "closure S - S \ (\n. X n)" + proof + fix x + assume "x \ closure S - S" + then have "x \ closure S" "x \ S" by auto + show "x \ (\n. X n)" + proof + fix n + have "ball 0 1 = closure (D n) \ A n" + by (auto simp: D_def A_def le_less_trans) + with fim have Seq: "S = f ` (closure (D n)) \ f ` (A n)" + by (simp add: image_Un) + have "continuous_on (closure (D n)) f" + by (simp add: D_def cball_subset_ball_iff continuous_on_subset [OF contf]) + moreover have "compact (closure (D n))" + by (simp add: D_def) + ultimately have clo_fim: "closed (f ` closure (D n))" + using compact_continuous_image compact_imp_closed by blast + have *: "(f ` cball 0 (1 - 1 / (real n + 2))) \ S" + by (force simp: D_def Seq) + show "x \ X n" + using \x \ closure S\ unfolding X_def Seq + using \x \ S\ * D_def clo_fim by auto + qed + qed + moreover have "(\n. X n) \ closure S - S" + proof - + have "(\n. X n) \ closure S" + proof - + have "(\n. X n) \ X 0" + by blast + also have "... \ closure S" + apply (simp add: X_def fim [symmetric]) + apply (rule closure_mono) + by (auto simp: A_def) + finally show "(\n. X n) \ closure S" . + qed + moreover have "(\n. X n) \ S \ {}" + proof (clarify, clarsimp simp: X_def fim [symmetric]) + fix x assume x [rule_format]: "\n. f x \ closure (f ` A n)" and "cmod x < 1" + then obtain n where n: "1 / (1 - norm x) < of_nat n" + using reals_Archimedean2 by blast + with \cmod x < 1\ gr0I have XX: "1 / of_nat n < 1 - norm x" and "n > 0" + by (fastforce simp: field_split_simps algebra_simps)+ + have "f x \ f ` (D n)" + using n \cmod x < 1\ by (auto simp: field_split_simps algebra_simps D_def) + moreover have " f ` D n \ closure (f ` A n) = {}" + proof - + have op_fDn: "open(f ` (D n))" + proof (rule invariance_of_domain) + show "continuous_on (D n) f" + by (rule continuous_on_subset [OF contf D01]) + show "open (D n)" + by (simp add: D_def) + show "inj_on f (D n)" + unfolding inj_on_def using D01 by (metis gf mem_ball_0 subsetCE) + qed + have injf: "inj_on f (ball 0 1)" + by (metis mem_ball_0 inj_on_def gf) + have "D n \ A n \ ball 0 1" + using D01 A01 by simp + moreover have "D n \ A n = {}" + by (auto simp: D_def A_def) + ultimately have "f ` D n \ f ` A n = {}" + by (metis A01 D01 image_is_empty inj_on_image_Int injf) + then show ?thesis + by (simp add: open_Int_closure_eq_empty [OF op_fDn]) + qed + ultimately show False + using x [of n] by blast + qed + ultimately + show "(\n. X n) \ closure S - S" + using closure_subset disjoint_iff_not_equal by blast + qed + ultimately have "closure S - S = (\n. X n)" by blast + then have frontierS: "frontier S = (\n. X n)" + by (simp add: frontier_def openS interior_open) + show ?thesis + proof (cases "bounded S") + case True + have bouX: "bounded (X n)" for n + apply (simp add: X_def) + apply (rule bounded_closure) + by (metis A01 fim image_mono bounded_subset [OF True]) + have compaX: "compact (X n)" for n + apply (simp add: compact_eq_bounded_closed bouX) + apply (auto simp: X_def) + done + have "connected (\n. X n)" + by (metis nestX compaX connX connected_nest) + then show ?thesis + by (simp add: True \frontier S = (\n. X n)\) + next + case False + have unboundedX: "\ bounded(X n)" for n + proof + assume bXn: "bounded(X n)" + have "continuous_on (cball 0 (1 - 1 / (2 + real n))) f" + by (simp add: cball_subset_ball_iff continuous_on_subset [OF contf]) + then have "bounded (f ` cball 0 (1 - 1 / (2 + real n)))" + by (simp add: compact_imp_bounded [OF compact_continuous_image]) + moreover have "bounded (f ` A n)" + by (auto simp: X_def closure_subset image_subset_iff bounded_subset [OF bXn]) + ultimately have "bounded (f ` (cball 0 (1 - 1/(2 + real n)) \ A n))" + by (simp add: image_Un) + then have "bounded (f ` ball 0 1)" + apply (rule bounded_subset) + apply (auto simp: A_def algebra_simps) + done + then show False + using False by (simp add: fim [symmetric]) + qed + have clo_INTX: "closed(\(range X))" + by (metis cloX closed_INT) + then have lcX: "locally compact (\(range X))" + by (metis closed_imp_locally_compact) + have False if C: "C \ components (frontier S)" and boC: "bounded C" for C + proof - + have "closed C" + by (metis C closed_components frontier_closed) + then have "compact C" + by (metis boC compact_eq_bounded_closed) + have Cco: "C \ components (\(range X))" + by (metis frontierS C) + obtain K where "C \ K" "compact K" + and Ksub: "K \ \(range X)" and clo: "closed(\(range X) - K)" + proof (cases "{k. C \ k \ compact k \ openin (top_of_set (\(range X))) k} = {}") + case True + then show ?thesis + using Sura_Bura [OF lcX Cco \compact C\] boC + by (simp add: True) + next + case False + then obtain L where "compact L" "C \ L" and K: "openin (top_of_set (\x. X x)) L" + by blast + show ?thesis + proof + show "L \ \(range X)" + by (metis K openin_imp_subset) + show "closed (\(range X) - L)" + by (metis closedin_diff closedin_self closedin_closed_trans [OF _ clo_INTX] K) + qed (use \compact L\ \C \ L\ in auto) + qed + obtain U V where "open U" and "compact (closure U)" and "open V" "K \ U" + and V: "\(range X) - K \ V" and "U \ V = {}" + using separation_normal_compact [OF \compact K\ clo] by blast + then have "U \ (\ (range X) - K) = {}" + by blast + have "(closure U - U) \ (\n. X n \ closure U) \ {}" + proof (rule compact_imp_fip) + show "compact (closure U - U)" + by (metis \compact (closure U)\ \open U\ compact_diff) + show "\T. T \ range (\n. X n \ closure U) \ closed T" + by clarify (metis cloX closed_Int closed_closure) + show "(closure U - U) \ \\ \ {}" + if "finite \" and \: "\ \ range (\n. X n \ closure U)" for \ + proof + assume empty: "(closure U - U) \ \\ = {}" + obtain J where "finite J" and J: "\ = (\n. X n \ closure U) ` J" + using finite_subset_image [OF \finite \\ \] by auto + show False + proof (cases "J = {}") + case True + with J empty have "closed U" + by (simp add: closure_subset_eq) + have "C \ {}" + using C in_components_nonempty by blast + then have "U \ {}" + using \K \ U\ \C \ K\ by blast + moreover have "U \ UNIV" + using \compact (closure U)\ by auto + ultimately show False + using \open U\ \closed U\ clopen by blast + next + case False + define j where "j \ Max J" + have "j \ J" + by (simp add: False \finite J\ j_def) + have jmax: "\m. m \ J \ m \ j" + by (simp add: j_def \finite J\) + have "\ ((\n. X n \ closure U) ` J) = X j \ closure U" + using False jmax nestX \j \ J\ by auto + then have "X j \ closure U = X j \ U" + apply safe + using DiffI J empty apply auto[1] + using closure_subset by blast + then have "openin (top_of_set (X j)) (X j \ closure U)" + by (simp add: openin_open_Int \open U\) + moreover have "closedin (top_of_set (X j)) (X j \ closure U)" + by (simp add: closedin_closed_Int) + moreover have "X j \ closure U \ X j" + by (metis unboundedX \compact (closure U)\ bounded_subset compact_eq_bounded_closed inf.order_iff) + moreover have "X j \ closure U \ {}" + proof - + have "C \ {}" + using C in_components_nonempty by blast + moreover have "C \ X j \ closure U" + using \C \ K\ \K \ U\ Ksub closure_subset by blast + ultimately show ?thesis by blast + qed + ultimately show False + using connX [of j] by (force simp: connected_clopen) + qed + qed + qed + moreover have "(\n. X n \ closure U) = (\n. X n) \ closure U" + by blast + moreover have "x \ U" if "\n. x \ X n" "x \ closure U" for x + proof - + have "x \ V" + using \U \ V = {}\ \open V\ closure_iff_nhds_not_empty that(2) by blast + then show ?thesis + by (metis (no_types) Diff_iff INT_I V \K \ U\ contra_subsetD that(1)) + qed + ultimately show False + by (auto simp: open_Int_closure_eq_empty [OF \open V\, of U]) + qed + then show ?thesis + by (auto simp: False) + qed + qed +qed + + +lemma unbounded_complement_components: + assumes C: "C \ components (- S)" and S: "connected S" + and prev: "if bounded S then connected(frontier S) + else \C \ components(frontier S). \ bounded C" + shows "\ bounded C" +proof (cases "bounded S") + case True + with prev have "S \ UNIV" and confr: "connected(frontier S)" + by auto + obtain w where C_ccsw: "C = connected_component_set (- S) w" and "w \ S" + using C by (auto simp: components_def) + show ?thesis + proof (cases "S = {}") + case True with C show ?thesis by auto + next + case False + show ?thesis + proof + assume "bounded C" + then have "outside C \ {}" + using outside_bounded_nonempty by metis + then obtain z where z: "\ bounded (connected_component_set (- C) z)" and "z \ C" + by (auto simp: outside_def) + have clo_ccs: "closed (connected_component_set (- S) x)" for x + by (simp add: closed_Compl closed_connected_component openS) + have "connected_component_set (- S) w = connected_component_set (- S) z" + proof (rule joinable_connected_component_eq [OF confr]) + show "frontier S \ - S" + using openS by (auto simp: frontier_def interior_open) + have False if "connected_component_set (- S) w \ frontier (- S) = {}" + proof - + have "C \ frontier S = {}" + using that by (simp add: C_ccsw) + then show False + by (metis C_ccsw ComplI Compl_eq_Compl_iff Diff_subset False \w \ S\ clo_ccs closure_closed compl_bot_eq connected_component_eq_UNIV connected_component_eq_empty empty_subsetI frontier_complement frontier_def frontier_not_empty frontier_of_connected_component_subset le_inf_iff subset_antisym) + qed + then show "connected_component_set (- S) w \ frontier S \ {}" + by auto + have *: "\frontier C \ C; frontier C \ F; frontier C \ {}\ \ C \ F \ {}" for C F::"complex set" + by blast + have "connected_component_set (- S) z \ frontier (- S) \ {}" + proof (rule *) + show "frontier (connected_component_set (- S) z) \ connected_component_set (- S) z" + by (auto simp: closed_Compl closed_connected_component frontier_def openS) + show "frontier (connected_component_set (- S) z) \ frontier (- S)" + using frontier_of_connected_component_subset by fastforce + have "\ bounded (-S)" + by (simp add: True cobounded_imp_unbounded) + then have "connected_component_set (- S) z \ {}" + apply (simp only: connected_component_eq_empty) + using confr openS \bounded C\ \w \ S\ + apply (simp add: frontier_def interior_open C_ccsw) + by (metis ComplI Compl_eq_Diff_UNIV connected_UNIV closed_closure closure_subset connected_component_eq_self + connected_diff_open_from_closed subset_UNIV) + then show "frontier (connected_component_set (- S) z) \ {}" + apply (simp add: frontier_eq_empty connected_component_eq_UNIV) + apply (metis False compl_top_eq double_compl) + done + qed + then show "connected_component_set (- S) z \ frontier S \ {}" + by auto + qed + then show False + by (metis C_ccsw Compl_iff \w \ S\ \z \ C\ connected_component_eq_empty connected_component_idemp) + qed + qed +next + case False + obtain w where C_ccsw: "C = connected_component_set (- S) w" and "w \ S" + using C by (auto simp: components_def) + have "frontier (connected_component_set (- S) w) \ connected_component_set (- S) w" + by (simp add: closed_Compl closed_connected_component frontier_subset_eq openS) + moreover have "frontier (connected_component_set (- S) w) \ frontier S" + using frontier_complement frontier_of_connected_component_subset by blast + moreover have "frontier (connected_component_set (- S) w) \ {}" + by (metis C C_ccsw False bounded_empty compl_top_eq connected_component_eq_UNIV double_compl frontier_not_empty in_components_nonempty) + ultimately obtain z where zin: "z \ frontier S" and z: "z \ connected_component_set (- S) w" + by blast + have *: "connected_component_set (frontier S) z \ components(frontier S)" + by (simp add: \z \ frontier S\ componentsI) + with prev False have "\ bounded (connected_component_set (frontier S) z)" + by simp + moreover have "connected_component (- S) w = connected_component (- S) z" + using connected_component_eq [OF z] by force + ultimately show ?thesis + by (metis C_ccsw * zin bounded_subset closed_Compl closure_closed connected_component_maximal + connected_component_refl connected_connected_component frontier_closures in_components_subset le_inf_iff mem_Collect_eq openS) +qed + +lemma empty_inside: + assumes "connected S" "\C. C \ components (- S) \ \ bounded C" + shows "inside S = {}" + using assms by (auto simp: components_def inside_def) + +lemma empty_inside_imp_simply_connected: + "\connected S; inside S = {}\ \ simply_connected S" + by (metis ComplI inside_Un_outside openS outside_mono simply_connected_eq_winding_number_zero subsetCE sup_bot.left_neutral winding_number_zero_in_outside) + +end + +proposition + fixes S :: "complex set" + assumes "open S" + shows simply_connected_eq_frontier_properties: + "simply_connected S \ + connected S \ + (if bounded S then connected(frontier S) + else (\C \ components(frontier S). \bounded C))" (is "?fp") + and simply_connected_eq_unbounded_complement_components: + "simply_connected S \ + connected S \ (\C \ components(- S). \bounded C)" (is "?ucc") + and simply_connected_eq_empty_inside: + "simply_connected S \ + connected S \ inside S = {}" (is "?ei") +proof - + interpret SC_Chain + using assms by (simp add: SC_Chain_def) + have "?fp \ ?ucc \ ?ei" +proof - + have *: "\\ \ \; \ \ \; \ \ \; \ \ \\ + \ (\ \ \) \ (\ \ \) \ (\ \ \)" for \ \ \ \ + by blast + show ?thesis + apply (rule *) + using frontier_properties simply_connected_imp_connected apply blast +apply clarify + using unbounded_complement_components simply_connected_imp_connected apply blast + using empty_inside apply blast + using empty_inside_imp_simply_connected apply blast + done +qed + then show ?fp ?ucc ?ei + by safe +qed + +lemma simply_connected_iff_simple: + fixes S :: "complex set" + assumes "open S" "bounded S" + shows "simply_connected S \ connected S \ connected(- S)" + apply (simp add: simply_connected_eq_unbounded_complement_components assms, safe) + apply (metis DIM_complex assms(2) cobounded_has_bounded_component double_compl order_refl) + by (meson assms inside_bounded_complement_connected_empty simply_connected_eq_empty_inside simply_connected_eq_unbounded_complement_components) + +subsection\Further equivalences based on continuous logs and sqrts\ + +context SC_Chain +begin + +lemma continuous_log: + fixes f :: "complex\complex" + assumes S: "simply_connected S" + and contf: "continuous_on S f" and nz: "\z. z \ S \ f z \ 0" + shows "\g. continuous_on S g \ (\z \ S. f z = exp(g z))" +proof - + consider "S = {}" | "S homeomorphic ball (0::complex) 1" + using simply_connected_eq_homeomorphic_to_disc [OF openS] S by metis + then show ?thesis + proof cases + case 1 then show ?thesis + by simp + next + case 2 + then obtain h k :: "complex\complex" + where kh: "\x. x \ S \ k(h x) = x" and him: "h ` S = ball 0 1" + and conth: "continuous_on S h" + and hk: "\y. y \ ball 0 1 \ h(k y) = y" and kim: "k ` ball 0 1 = S" + and contk: "continuous_on (ball 0 1) k" + unfolding homeomorphism_def homeomorphic_def by metis + obtain g where contg: "continuous_on (ball 0 1) g" + and expg: "\z. z \ ball 0 1 \ (f \ k) z = exp (g z)" + proof (rule continuous_logarithm_on_ball) + show "continuous_on (ball 0 1) (f \ k)" + apply (rule continuous_on_compose [OF contk]) + using kim continuous_on_subset [OF contf] + by blast + show "\z. z \ ball 0 1 \ (f \ k) z \ 0" + using kim nz by auto + qed auto + then show ?thesis + by (metis comp_apply conth continuous_on_compose him imageI kh) + qed +qed + +lemma continuous_sqrt: + fixes f :: "complex\complex" + assumes contf: "continuous_on S f" and nz: "\z. z \ S \ f z \ 0" + and prev: "\f::complex\complex. + \continuous_on S f; \z. z \ S \ f z \ 0\ + \ \g. continuous_on S g \ (\z \ S. f z = exp(g z))" + shows "\g. continuous_on S g \ (\z \ S. f z = (g z)\<^sup>2)" +proof - + obtain g where contg: "continuous_on S g" and geq: "\z. z \ S \ f z = exp(g z)" + using contf nz prev by metis + show ?thesis +proof (intro exI ballI conjI) + show "continuous_on S (\z. exp(g z/2))" + by (intro continuous_intros) (auto simp: contg) + show "\z. z \ S \ f z = (exp (g z/2))\<^sup>2" + by (metis (no_types, lifting) divide_inverse exp_double geq mult.left_commute mult.right_neutral right_inverse zero_neq_numeral) + qed +qed + +lemma continuous_sqrt_imp_simply_connected: + assumes "connected S" + and prev: "\f::complex\complex. \continuous_on S f; \z \ S. f z \ 0\ + \ \g. continuous_on S g \ (\z \ S. f z = (g z)\<^sup>2)" + shows "simply_connected S" +proof (clarsimp simp add: simply_connected_eq_holomorphic_sqrt [OF openS] \connected S\) + fix f + assume "f holomorphic_on S" and nz: "\z\S. f z \ 0" + then obtain g where contg: "continuous_on S g" and geq: "\z. z \ S \ f z = (g z)\<^sup>2" + by (metis holomorphic_on_imp_continuous_on prev) + show "\g. g holomorphic_on S \ (\z\S. f z = (g z)\<^sup>2)" + proof (intro exI ballI conjI) + show "g holomorphic_on S" + proof (clarsimp simp add: holomorphic_on_open [OF openS]) + fix z + assume "z \ S" + with nz geq have "g z \ 0" + by auto + obtain \ where "0 < \" "\w. \w \ S; dist w z < \\ \ dist (g w) (g z) < cmod (g z)" + using contg [unfolded continuous_on_iff] by (metis \g z \ 0\ \z \ S\ zero_less_norm_iff) + then have \: "\w. \w \ S; w \ ball z \\ \ g w + g z \ 0" + apply (clarsimp simp: dist_norm) + by (metis \g z \ 0\ add_diff_cancel_left' diff_0_right norm_eq_zero norm_increases_online norm_minus_commute norm_not_less_zero not_less_iff_gr_or_eq) + have *: "(\x. (f x - f z) / (x - z) / (g x + g z)) \z\ deriv f z / (g z + g z)" + apply (intro tendsto_intros) + using SC_Chain.openS SC_Chain_axioms \f holomorphic_on S\ \z \ S\ has_field_derivativeD holomorphic_derivI apply fastforce + using \z \ S\ contg continuous_on_eq_continuous_at isCont_def openS apply blast + by (simp add: \g z \ 0\) + then have "(g has_field_derivative deriv f z / (g z + g z)) (at z)" + unfolding has_field_derivative_iff + proof (rule Lim_transform_within_open) + show "open (ball z \ \ S)" + by (simp add: openS open_Int) + show "z \ ball z \ \ S" + using \z \ S\ \0 < \\ by simp + show "\x. \x \ ball z \ \ S; x \ z\ + \ (f x - f z) / (x - z) / (g x + g z) = (g x - g z) / (x - z)" + using \ + apply (simp add: geq \z \ S\ divide_simps) + apply (auto simp: algebra_simps power2_eq_square) + done + qed + then show "\f'. (g has_field_derivative f') (at z)" .. + qed + qed (use geq in auto) +qed + +end + +proposition + fixes S :: "complex set" + assumes "open S" + shows simply_connected_eq_continuous_log: + "simply_connected S \ + connected S \ + (\f::complex\complex. continuous_on S f \ (\z \ S. f z \ 0) + \ (\g. continuous_on S g \ (\z \ S. f z = exp (g z))))" (is "?log") + and simply_connected_eq_continuous_sqrt: + "simply_connected S \ + connected S \ + (\f::complex\complex. continuous_on S f \ (\z \ S. f z \ 0) + \ (\g. continuous_on S g \ (\z \ S. f z = (g z)\<^sup>2)))" (is "?sqrt") +proof - + interpret SC_Chain + using assms by (simp add: SC_Chain_def) + have "?log \ ?sqrt" +proof - + have *: "\\ \ \; \ \ \; \ \ \\ + \ (\ \ \) \ (\ \ \)" for \ \ \ + by blast + show ?thesis + apply (rule *) + apply (simp add: local.continuous_log winding_number_zero) + apply (simp add: continuous_sqrt) + apply (simp add: continuous_sqrt_imp_simply_connected) + done +qed + then show ?log ?sqrt + by safe +qed + + +subsection\<^marker>\tag unimportant\ \More Borsukian results\ + +lemma Borsukian_componentwise_eq: + fixes S :: "'a::euclidean_space set" + assumes S: "locally connected S \ compact S" + shows "Borsukian S \ (\C \ components S. Borsukian C)" +proof - + have *: "ANR(-{0::complex})" + by (simp add: ANR_delete open_Compl open_imp_ANR) + show ?thesis + using cohomotopically_trivial_on_components [OF assms *] by (auto simp: Borsukian_alt) +qed + +lemma Borsukian_componentwise: + fixes S :: "'a::euclidean_space set" + assumes "locally connected S \ compact S" "\C. C \ components S \ Borsukian C" + shows "Borsukian S" + by (metis Borsukian_componentwise_eq assms) + +lemma simply_connected_eq_Borsukian: + fixes S :: "complex set" + shows "open S \ (simply_connected S \ connected S \ Borsukian S)" + by (auto simp: simply_connected_eq_continuous_log Borsukian_continuous_logarithm) + +lemma Borsukian_eq_simply_connected: + fixes S :: "complex set" + shows "open S \ Borsukian S \ (\C \ components S. simply_connected C)" +apply (auto simp: Borsukian_componentwise_eq open_imp_locally_connected) + using in_components_connected open_components simply_connected_eq_Borsukian apply blast + using open_components simply_connected_eq_Borsukian by blast + +lemma Borsukian_separation_open_closed: + fixes S :: "complex set" + assumes S: "open S \ closed S" and "bounded S" + shows "Borsukian S \ connected(- S)" + using S +proof + assume "open S" + show ?thesis + unfolding Borsukian_eq_simply_connected [OF \open S\] + by (meson \open S\ \bounded S\ bounded_subset in_components_connected in_components_subset nonseparation_by_component_eq open_components simply_connected_iff_simple) +next + assume "closed S" + with \bounded S\ show ?thesis + by (simp add: Borsukian_separation_compact compact_eq_bounded_closed) +qed + + +subsection\Finally, the Riemann Mapping Theorem\ + +theorem Riemann_mapping_theorem: + "open S \ simply_connected S \ + S = {} \ S = UNIV \ + (\f g. f holomorphic_on S \ g holomorphic_on ball 0 1 \ + (\z \ S. f z \ ball 0 1 \ g(f z) = z) \ + (\z \ ball 0 1. g z \ S \ f(g z) = z))" + (is "_ = ?rhs") +proof - + have "simply_connected S \ ?rhs" if "open S" + by (simp add: simply_connected_eq_biholomorphic_to_disc that) + moreover have "open S" if "?rhs" + proof - + { fix f g + assume g: "g holomorphic_on ball 0 1" "\z\ball 0 1. g z \ S \ f (g z) = z" + and "\z\S. cmod (f z) < 1 \ g (f z) = z" + then have "S = g ` (ball 0 1)" + by (force simp:) + then have "open S" + by (metis open_ball g inj_on_def open_mapping_thm3) + } + with that show "open S" by auto + qed + ultimately show ?thesis by metis +qed + +end diff --git a/src/HOL/Complex_Analysis/Winding_Numbers.thy b/src/HOL/Complex_Analysis/Winding_Numbers.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/Winding_Numbers.thy @@ -0,0 +1,1153 @@ +section \Winding Numbers\ + +text\By John Harrison et al. Ported from HOL Light by L C Paulson (2017)\ + +theory Winding_Numbers +imports + 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\ + +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 + diff --git a/src/HOL/Complex_Analysis/document/root.bib b/src/HOL/Complex_Analysis/document/root.bib new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/document/root.bib @@ -0,0 +1,3 @@ + + +@misc{dummy} diff --git a/src/HOL/Complex_Analysis/document/root.tex b/src/HOL/Complex_Analysis/document/root.tex new file mode 100644 --- /dev/null +++ b/src/HOL/Complex_Analysis/document/root.tex @@ -0,0 +1,43 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{graphicx} +\usepackage{isabelle} +\usepackage{isabellesym} +\usepackage{latexsym} +\usepackage{textcomp} +\usepackage{amsmath} +\usepackage{amssymb} +\usepackage[only,bigsqcap]{stmaryrd} +\usepackage{pdfsetup} + +\usepackage{tocloft} + +\urlstyle{rm} +\isabellestyle{literalunderscore} +\pagestyle{myheadings} + +\raggedbottom + +\begin{document} + +\title{Complex Analysis} +\maketitle + +\tableofcontents + +\begin{center} + \includegraphics[height=\textheight]{session_graph} +\end{center} + +\newpage + +\renewcommand{\setisabellecontext}[1]{\markright{\href{#1.html}{#1.thy}}} + +\parindent 0pt\parskip 0.5ex +\input{session} + +\pagestyle{headings} +\bibliographystyle{abbrv} +\bibliography{root} +\nocite{dummy} + +\end{document} diff --git a/src/HOL/ROOT b/src/HOL/ROOT --- a/src/HOL/ROOT +++ b/src/HOL/ROOT @@ -1,1116 +1,1125 @@ chapter HOL session HOL (main) = Pure + description " Classical Higher-order Logic. " options [strict_facts] sessions Tools theories Main (global) Complex_Main (global) document_files "root.bib" "root.tex" session "HOL-Proofs" (timing) in Proofs = Pure + description " HOL-Main with explicit proof terms. " options [quick_and_dirty = false, record_proofs = 2, parallel_limit = 500] sessions "HOL-Library" theories "HOL-Library.Realizers" session "HOL-Library" (main timing) in Library = HOL + description " Classical Higher-order Logic -- batteries included. " theories Library (*conflicting type class instantiations and dependent applications*) Finite_Lattice List_Lexorder Prefix_Order Product_Lexorder Product_Order Subseq_Order (*conflicting syntax*) Datatype_Records (*data refinements and dependent applications*) AList_Mapping Code_Binary_Nat Code_Prolog Code_Real_Approx_By_Float Code_Target_Numeral DAList DAList_Multiset RBT_Mapping RBT_Set (*printing modifications*) OptionalSugar (*prototypic tools*) Predicate_Compile_Quickcheck (*legacy tools*) Old_Datatype Old_Recdef Realizers Refute document_files "root.bib" "root.tex" session "HOL-Analysis" (main timing) in Analysis = HOL + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] sessions "HOL-Library" "HOL-Computational_Algebra" theories Analysis document_files "root.tex" "root.bib" +session "HOL-Complex_Analysis" (main timing) in Complex_Analysis = "HOL-Analysis" + + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", + document_variants = "document:manual=-proof,-ML,-unimportant"] + theories + Complex_Analysis + document_files + "root.tex" + "root.bib" + session "HOL-Analysis-ex" in "Analysis/ex" = "HOL-Analysis" + theories Approximations Metric_Arith_Examples session "HOL-Homology" (timing) in Homology = "HOL-Analysis" + options [document_tags = "theorem%important,corollary%important,proposition%important,class%important,instantiation%important,subsubsection%unimportant,%unimportant", document_variants = "document:manual=-proof,-ML,-unimportant"] sessions "HOL-Algebra" theories Homology document_files "root.tex" session "HOL-Computational_Algebra" (main timing) in "Computational_Algebra" = "HOL-Library" + theories Computational_Algebra (*conflicting type class instantiations and dependent applications*) Field_as_Ring session "HOL-Real_Asymp" in Real_Asymp = HOL + sessions "HOL-Decision_Procs" theories Real_Asymp Real_Asymp_Approx Real_Asymp_Examples session "HOL-Real_Asymp-Manual" in "Real_Asymp/Manual" = "HOL-Real_Asymp" + theories Real_Asymp_Doc document_files (in "~~/src/Doc") "iman.sty" "extra.sty" "isar.sty" document_files "root.tex" "style.sty" session "HOL-Hahn_Banach" in Hahn_Banach = "HOL-Library" + description " Author: Gertrud Bauer, TU Munich The Hahn-Banach theorem for real vector spaces. This is the proof of the Hahn-Banach theorem for real vectorspaces, following H. Heuser, Funktionalanalysis, p. 228 -232. The Hahn-Banach theorem is one of the fundamental theorems of functional analysis. It is a conclusion of Zorn's lemma. Two different formaulations of the theorem are presented, one for general real vectorspaces and its application to normed vectorspaces. The theorem says, that every continous linearform, defined on arbitrary subspaces (not only one-dimensional subspaces), can be extended to a continous linearform on the whole vectorspace. " sessions "HOL-Analysis" theories Hahn_Banach document_files "root.bib" "root.tex" session "HOL-Induct" in Induct = "HOL-Library" + description " Examples of (Co)Inductive Definitions. Comb proves the Church-Rosser theorem for combinators (see http://www.cl.cam.ac.uk/ftp/papers/reports/TR396-lcp-generic-automatic-proof-tools.ps.gz). Mutil is the famous Mutilated Chess Board problem (see http://www.cl.cam.ac.uk/ftp/papers/reports/TR394-lcp-mutilated-chess-board.dvi.gz). PropLog proves the completeness of a formalization of propositional logic (see http://www.cl.cam.ac.uk/Research/Reports/TR312-lcp-set-II.ps.gz). Exp demonstrates the use of iterated inductive definitions to reason about mutually recursive relations. " theories [quick_and_dirty] Common_Patterns theories Nested_Datatype QuoDataType QuoNestedDataType Term SList ABexp Infinitely_Branching_Tree Ordinals Sigma_Algebra Comb PropLog Com document_files "root.tex" session "HOL-IMP" (timing) in IMP = "HOL-Library" + options [document_variants = document] theories BExp ASM Finite_Reachable Denotational Compiler2 Poly_Types Sec_Typing Sec_TypingT Def_Init_Big Def_Init_Small Fold Live Live_True Hoare_Examples Hoare_Sound_Complete VCG Hoare_Total VCG_Total_EX VCG_Total_EX2 Collecting1 Collecting_Examples Abs_Int_Tests Abs_Int1_parity Abs_Int1_const Abs_Int3 Procs_Dyn_Vars_Dyn Procs_Stat_Vars_Dyn Procs_Stat_Vars_Stat C_like OO document_files "root.bib" "root.tex" session "HOL-IMPP" in IMPP = HOL + description \ Author: David von Oheimb Copyright 1999 TUM IMPP -- An imperative language with procedures. This is an extension of IMP with local variables and mutually recursive procedures. For documentation see "Hoare Logic for Mutual Recursion and Local Variables" (https://isabelle.in.tum.de/Bali/papers/FSTTCS99.html). \ theories EvenOdd session "HOL-Data_Structures" (timing) in Data_Structures = HOL + options [document_variants = document] sessions "HOL-Number_Theory" theories [document = false] Less_False theories Sorting Balance Tree_Map AVL_Map RBT_Map Tree23_Map Tree234_Map Brother12_Map AA_Map Set2_Join_RBT Array_Braun Trie_Fun Trie_Map Tries_Binary Leftist_Heap Binomial_Heap document_files "root.tex" "root.bib" session "HOL-Import" in Import = HOL + theories HOL_Light_Maps theories [condition = HOL_LIGHT_BUNDLE] HOL_Light_Import session "HOL-Number_Theory" (main timing) in Number_Theory = "HOL-Computational_Algebra" + description " Fundamental Theorem of Arithmetic, Chinese Remainder Theorem, Fermat/Euler Theorem, Wilson's Theorem, some lemmas for Quadratic Reciprocity. " sessions "HOL-Algebra" theories Number_Theory document_files "root.tex" session "HOL-Hoare" in Hoare = HOL + description " Verification of imperative programs (verification conditions are generated automatically from pre/post conditions and loop invariants). " theories Hoare document_files "root.bib" "root.tex" session "HOL-Hoare_Parallel" (timing) in Hoare_Parallel = HOL + description " Verification of shared-variable imperative programs a la Owicki-Gries. (verification conditions are generated automatically). " theories Hoare_Parallel document_files "root.bib" "root.tex" session "HOL-Codegenerator_Test" in Codegenerator_Test = "HOL-Number_Theory" + sessions "HOL-Data_Structures" "HOL-ex" theories Generate Generate_Binary_Nat Generate_Target_Nat Generate_Efficient_Datastructures Code_Lazy_Test Code_Test_PolyML Code_Test_Scala theories [condition = ISABELLE_GHC] Code_Test_GHC theories [condition = ISABELLE_MLTON] Code_Test_MLton theories [condition = ISABELLE_OCAMLFIND] Code_Test_OCaml theories [condition = ISABELLE_SMLNJ] Code_Test_SMLNJ session "HOL-Metis_Examples" (timing) in Metis_Examples = "HOL-Library" + description " Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Jasmin Blanchette, TU Muenchen Testing Metis and Sledgehammer. " sessions "HOL-Decision_Procs" theories Abstraction Big_O Binary_Tree Clausification Message Proxies Tarski Trans_Closure Sets session "HOL-Nitpick_Examples" in Nitpick_Examples = "HOL-Library" + description " Author: Jasmin Blanchette, TU Muenchen Copyright 2009 " theories [quick_and_dirty] Nitpick_Examples session "HOL-Algebra" (main timing) in Algebra = "HOL-Computational_Algebra" + description " Author: Clemens Ballarin, started 24 September 1999, and many others The Isabelle Algebraic Library. " sessions "HOL-Cardinals" theories (* Orders and Lattices *) Galois_Connection (* Knaster-Tarski theorem and Galois connections *) (* Groups *) FiniteProduct (* Product operator for commutative groups *) Sylow (* Sylow's theorem *) Bij (* Automorphism Groups *) Multiplicative_Group Zassenhaus (* The Zassenhaus lemma *) (* Rings *) Divisibility (* Rings *) IntRing (* Ideals and residue classes *) UnivPoly (* Polynomials *) (* Main theory *) Algebra document_files "root.bib" "root.tex" session "HOL-Auth" (timing) in Auth = "HOL-Library" + description " A new approach to verifying authentication protocols. " directories "Smartcard" "Guard" theories Auth_Shared Auth_Public "Smartcard/Auth_Smartcard" "Guard/Auth_Guard_Shared" "Guard/Auth_Guard_Public" document_files "root.tex" session "HOL-UNITY" (timing) in UNITY = "HOL-Auth" + description " Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge Verifying security protocols using Chandy and Misra's UNITY formalism. " directories "Simple" "Comp" theories (*Basic meta-theory*) UNITY_Main (*Simple examples: no composition*) "Simple/Deadlock" "Simple/Common" "Simple/Network" "Simple/Token" "Simple/Channel" "Simple/Lift" "Simple/Mutex" "Simple/Reach" "Simple/Reachability" (*Verifying security protocols using UNITY*) "Simple/NSP_Bad" (*Example of composition*) "Comp/Handshake" (*Universal properties examples*) "Comp/Counter" "Comp/Counterc" "Comp/Priority" "Comp/TimerArray" "Comp/Progress" "Comp/Alloc" "Comp/AllocImpl" "Comp/Client" (*obsolete*) ELT document_files "root.tex" session "HOL-Unix" in Unix = "HOL-Library" + options [print_mode = "no_brackets,no_type_brackets"] theories Unix document_files "root.bib" "root.tex" session "HOL-ZF" in ZF = "HOL-Library" + theories MainZF Games document_files "root.tex" session "HOL-Imperative_HOL" (timing) in Imperative_HOL = "HOL-Library" + options [print_mode = "iff,no_brackets"] directories "ex" theories Imperative_HOL_ex document_files "root.bib" "root.tex" session "HOL-Decision_Procs" (timing) in Decision_Procs = "HOL-Algebra" + description " Various decision procedures, typically involving reflection. " directories "ex" theories Decision_Procs session "HOL-Proofs-ex" in "Proofs/ex" = "HOL-Proofs" + sessions "HOL-Isar_Examples" theories Hilbert_Classical Proof_Terms XML_Data session "HOL-Proofs-Extraction" (timing) in "Proofs/Extraction" = "HOL-Proofs" + description " Examples for program extraction in Higher-Order Logic. " options [quick_and_dirty = false] sessions "HOL-Computational_Algebra" theories Greatest_Common_Divisor Warshall Higman_Extraction Pigeonhole Euclid document_files "root.bib" "root.tex" session "HOL-Proofs-Lambda" (timing) in "Proofs/Lambda" = "HOL-Proofs" + description \ Lambda Calculus in de Bruijn's Notation. This session defines lambda-calculus terms with de Bruijn indixes and proves confluence of beta, eta and beta+eta. The paper "More Church-Rosser Proofs (in Isabelle/HOL)" describes the whole theory (see http://www.in.tum.de/~nipkow/pubs/jar2001.html). \ options [print_mode = "no_brackets", quick_and_dirty = false] sessions "HOL-Library" theories Eta StrongNorm Standardization WeakNorm document_files "root.bib" "root.tex" session "HOL-Prolog" in Prolog = HOL + description " Author: David von Oheimb (based on a lecture on Lambda Prolog by Nadathur) A bare-bones implementation of Lambda-Prolog. This is a simple exploratory implementation of Lambda-Prolog in HOL, including some minimal examples (in Test.thy) and a more typical example of a little functional language and its type system. " theories Test Type session "HOL-MicroJava" (timing) in MicroJava = "HOL-Library" + description " Formalization of a fragment of Java, together with a corresponding virtual machine and a specification of its bytecode verifier and a lightweight bytecode verifier, including proofs of type-safety. " sessions "HOL-Eisbach" directories "BV" "Comp" "DFA" "J" "JVM" theories MicroJava document_files "introduction.tex" "root.bib" "root.tex" session "HOL-NanoJava" in NanoJava = HOL + description " Hoare Logic for a tiny fragment of Java. " theories Example document_files "root.bib" "root.tex" session "HOL-Bali" (timing) in Bali = "HOL-Library" + theories AxExample AxSound AxCompl Trans TypeSafe document_files "root.tex" session "HOL-IOA" in IOA = HOL + description \ Author: Tobias Nipkow and Konrad Slind and Olaf Müller Copyright 1994--1996 TU Muenchen The meta-theory of I/O-Automata in HOL. This formalization has been significantly changed and extended, see HOLCF/IOA. There are also the proofs of two communication protocols which formerly have been here. @inproceedings{Nipkow-Slind-IOA, author={Tobias Nipkow and Konrad Slind}, title={{I/O} Automata in {Isabelle/HOL}}, booktitle={Proc.\ TYPES Workshop 1994}, publisher=Springer, series=LNCS, note={To appear}} ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/ioa.ps.gz and @inproceedings{Mueller-Nipkow, author={Olaf M\"uller and Tobias Nipkow}, title={Combining Model Checking and Deduction for {I/O}-Automata}, booktitle={Proc.\ TACAS Workshop}, organization={Aarhus University, BRICS report}, year=1995} ftp://ftp.informatik.tu-muenchen.de/local/lehrstuhl/nipkow/tacas.dvi.gz \ theories Solve session "HOL-Lattice" in Lattice = HOL + description " Author: Markus Wenzel, TU Muenchen Basic theory of lattices and orders. " theories CompleteLattice document_files "root.tex" session "HOL-ex" (timing) in ex = "HOL-Number_Theory" + description " Miscellaneous examples for Higher-Order Logic. " theories Adhoc_Overloading_Examples Antiquote Argo_Examples Arith_Examples Ballot BinEx Birthday_Paradox Bit_Lists Bubblesort CTL Cartouche_Examples Case_Product Chinese Classical Code_Binary_Nat_examples Code_Lazy_Demo Code_Timing Coercion_Examples Coherent Commands Computations Conditional_Parametricity_Examples Cubic_Quartic Datatype_Record_Examples Dedekind_Real Erdoes_Szekeres Eval_Examples Executable_Relation Execute_Choice Functions Function_Growth Gauge_Integration Groebner_Examples Guess HarmonicSeries Hebrew Hex_Bin_Examples IArray_Examples Iff_Oracle Induction_Schema Intuitionistic Join_Theory Lagrange List_to_Set_Comprehension_Examples LocaleTest2 "ML" MergeSort MonoidGroup Multiquote NatSum Normalization_by_Evaluation PER Parallel_Example Peano_Axioms Perm_Fragments PresburgerEx Primrec Pythagoras Quicksort Radix_Sort Records Reflection_Examples Refute_Examples Residue_Ring Rewrite_Examples SOS SOS_Cert Seq Serbian Set_Comprehension_Pointfree_Examples Set_Theory Simproc_Tests Simps_Case_Conv_Examples Sketch_and_Explore Sorting_Algorithms_Examples Sqrt Sqrt_Script Sudoku Sum_of_Powers Tarski Termination ThreeDivides Transfer_Debug Transfer_Int_Nat Transitive_Closure_Table_Ex Tree23 Triangular_Numbers Unification While_Combinator_Example Word veriT_Preprocessing theories [skip_proofs = false] SAT_Examples Meson_Test session "HOL-Isar_Examples" in Isar_Examples = "HOL-Computational_Algebra" + description " Miscellaneous Isabelle/Isar examples. " options [quick_and_dirty] theories Knaster_Tarski Peirce Drinker Cantor Structured_Statements Basic_Logic Expr_Compiler Fibonacci Group Group_Context Group_Notepad Hoare_Ex Mutilated_Checkerboard Puzzle Summation First_Order_Logic Higher_Order_Logic document_files "root.bib" "root.tex" session "HOL-Eisbach" in Eisbach = HOL + description \ The Eisbach proof method language and "match" method. \ sessions FOL "HOL-Analysis" theories Eisbach Tests Examples Examples_FOL Example_Metric session "HOL-SET_Protocol" (timing) in SET_Protocol = "HOL-Library" + description " Verification of the SET Protocol. " theories SET_Protocol document_files "root.tex" session "HOL-Matrix_LP" in Matrix_LP = "HOL-Library" + description " Two-dimensional matrices and linear programming. " directories "Compute_Oracle" theories Cplex document_files "root.tex" session "HOL-TLA" in TLA = HOL + description " Lamport's Temporal Logic of Actions. " theories TLA session "HOL-TLA-Inc" in "TLA/Inc" = "HOL-TLA" + theories Inc session "HOL-TLA-Buffer" in "TLA/Buffer" = "HOL-TLA" + theories DBuffer session "HOL-TLA-Memory" in "TLA/Memory" = "HOL-TLA" + theories MemoryImplementation session "HOL-TPTP" in TPTP = "HOL-Library" + description " Author: Jasmin Blanchette, TU Muenchen Author: Nik Sultana, University of Cambridge Copyright 2011 TPTP-related extensions. " theories ATP_Theory_Export MaSh_Eval TPTP_Interpret THF_Arith TPTP_Proof_Reconstruction theories ATP_Problem_Import session "HOL-Probability" (main timing) in "Probability" = "HOL-Analysis" + theories Probability document_files "root.tex" session "HOL-Probability-ex" (timing) in "Probability/ex" = "HOL-Probability" + theories Dining_Cryptographers Koepf_Duermuth_Countermeasure Measure_Not_CCC session "HOL-Nominal" in Nominal = "HOL-Library" + theories Nominal session "HOL-Nominal-Examples" (timing) in "Nominal/Examples" = "HOL-Nominal" + theories Class3 CK_Machine Compile Contexts Crary CR_Takahashi CR Fsub Height Lambda_mu Lam_Funs LocalWeakening Pattern SN SOS Standardization Support Type_Preservation Weakening W theories [quick_and_dirty] VC_Condition session "HOL-Cardinals" (timing) in Cardinals = HOL + description " Ordinals and Cardinals, Full Theories. " theories Cardinals Bounded_Set document_files "intro.tex" "root.tex" "root.bib" session "HOL-Datatype_Examples" (timing) in Datatype_Examples = "HOL-Library" + description " (Co)datatype Examples. " directories "Derivation_Trees" theories Compat Lambda_Term Process TreeFsetI "Derivation_Trees/Gram_Lang" "Derivation_Trees/Parallel_Composition" Koenig Lift_BNF Milner_Tofte Stream_Processor Misc_Codatatype Misc_Datatype Misc_Primcorec Misc_Primrec session "HOL-Corec_Examples" (timing) in Corec_Examples = "HOL-Library" + description " Corecursion Examples. " directories "Tests" theories LFilter Paper_Examples Stream_Processor "Tests/Simple_Nesting" "Tests/Iterate_GPV" theories [quick_and_dirty] "Tests/GPV_Bare_Bones" "Tests/Merge_D" "Tests/Merge_Poly" "Tests/Misc_Mono" "Tests/Misc_Poly" "Tests/Small_Concrete" "Tests/Stream_Friends" "Tests/TLList_Friends" "Tests/Type_Class" session "HOL-Word" (main timing) in Word = HOL + sessions "HOL-Library" theories Word More_Word Word_Examples document_files "root.bib" "root.tex" session "HOL-Statespace" in Statespace = HOL + theories [skip_proofs = false] StateSpaceEx document_files "root.tex" session "HOL-Nonstandard_Analysis" (timing) in Nonstandard_Analysis = "HOL-Computational_Algebra" + description " Nonstandard analysis. " theories Nonstandard_Analysis document_files "root.tex" session "HOL-Nonstandard_Analysis-Examples" (timing) in "Nonstandard_Analysis/Examples" = "HOL-Nonstandard_Analysis" + theories NSPrimes session "HOL-Mirabelle" in Mirabelle = HOL + theories Mirabelle_Test session "HOL-Mirabelle-ex" in "Mirabelle/ex" = "HOL-Mirabelle" + options [timeout = 60] theories Ex session "HOL-Word-SMT_Examples" (timing) in SMT_Examples = "HOL-Word" + options [quick_and_dirty] theories Boogie SMT_Examples SMT_Word_Examples SMT_Tests session "HOL-SPARK" in "SPARK" = "HOL-Word" + theories SPARK session "HOL-SPARK-Examples" in "SPARK/Examples" = "HOL-SPARK" + directories "Gcd" "Liseq" "RIPEMD-160" "Sqrt" theories "Gcd/Greatest_Common_Divisor" "Liseq/Longest_Increasing_Subsequence" "RIPEMD-160/F" "RIPEMD-160/Hash" "RIPEMD-160/K_L" "RIPEMD-160/K_R" "RIPEMD-160/R_L" "RIPEMD-160/Round" "RIPEMD-160/R_R" "RIPEMD-160/S_L" "RIPEMD-160/S_R" "Sqrt/Sqrt" export_files (in ".") "*:**.prv" session "HOL-SPARK-Manual" in "SPARK/Manual" = "HOL-SPARK" + options [show_question_marks = false] sessions "HOL-SPARK-Examples" theories Example_Verification VC_Principles Reference Complex_Types document_files "complex_types.ads" "complex_types_app.adb" "complex_types_app.ads" "Gcd.adb" "Gcd.ads" "intro.tex" "loop_invariant.adb" "loop_invariant.ads" "root.bib" "root.tex" "Simple_Gcd.adb" "Simple_Gcd.ads" session "HOL-Mutabelle" in Mutabelle = "HOL-Library" + theories MutabelleExtra session "HOL-Quickcheck_Examples" (timing) in Quickcheck_Examples = "HOL-Library" + theories Quickcheck_Examples Quickcheck_Lattice_Examples Completeness Quickcheck_Interfaces Quickcheck_Nesting_Example theories [condition = ISABELLE_GHC] Hotel_Example Quickcheck_Narrowing_Examples session "HOL-Quotient_Examples" (timing) in Quotient_Examples = "HOL-Algebra" + description " Author: Cezary Kaliszyk and Christian Urban " theories DList Quotient_FSet Quotient_Int Quotient_Message Lift_FSet Lift_Set Lift_Fun Quotient_Rat Lift_DList Int_Pow Lifting_Code_Dt_Test session "HOL-Predicate_Compile_Examples" (timing) in Predicate_Compile_Examples = "HOL-Library" + theories Examples Predicate_Compile_Tests Predicate_Compile_Quickcheck_Examples Specialisation_Examples IMP_1 IMP_2 (* FIXME since 21-Jul-2011 Hotel_Example_Small_Generator *) IMP_3 IMP_4 theories [condition = ISABELLE_SWIPL] Code_Prolog_Examples Context_Free_Grammar_Example Hotel_Example_Prolog Lambda_Example List_Examples theories [condition = ISABELLE_SWIPL, quick_and_dirty] Reg_Exp_Example session "HOL-Types_To_Sets" in Types_To_Sets = HOL + description " Experimental extension of Higher-Order Logic to allow translation of types to sets. " directories "Examples" theories Types_To_Sets "Examples/Prerequisites" "Examples/Finite" "Examples/T2_Spaces" "Examples/Unoverload_Def" "Examples/Linear_Algebra_On" session HOLCF (main timing) in HOLCF = HOL + description " Author: Franz Regensburger Author: Brian Huffman HOLCF -- a semantic extension of HOL by the LCF logic. " sessions "HOL-Library" theories HOLCF (global) document_files "root.tex" session "HOLCF-Tutorial" in "HOLCF/Tutorial" = HOLCF + theories Domain_ex Fixrec_ex New_Domain document_files "root.tex" session "HOLCF-Library" in "HOLCF/Library" = HOLCF + theories HOLCF_Library HOL_Cpo session "HOLCF-IMP" in "HOLCF/IMP" = HOLCF + description " IMP -- A WHILE-language and its Semantics. This is the HOLCF-based denotational semantics of a simple WHILE-language. " sessions "HOL-IMP" theories HoareEx document_files "isaverbatimwrite.sty" "root.tex" "root.bib" session "HOLCF-ex" in "HOLCF/ex" = "HOLCF-Library" + description " Miscellaneous examples for HOLCF. " theories Dnat Dagstuhl Focus_ex Fix2 Hoare Concurrency_Monad Loop Powerdomain_ex Domain_Proofs Letrec Pattern_Match session "HOLCF-FOCUS" in "HOLCF/FOCUS" = "HOLCF-Library" + description \ FOCUS: a theory of stream-processing functions Isabelle/HOLCF. For introductions to FOCUS, see "The Design of Distributed Systems - An Introduction to FOCUS" http://www4.in.tum.de/publ/html.php?e=2 "Specification and Refinement of a Buffer of Length One" http://www4.in.tum.de/publ/html.php?e=15 "Specification and Development of Interactive Systems: Focus on Streams, Interfaces, and Refinement" http://www4.in.tum.de/publ/html.php?e=321 \ theories Fstreams FOCUS Buffer_adm session IOA (timing) in "HOLCF/IOA" = HOLCF + description " Author: Olaf Mueller Copyright 1997 TU München A formalization of I/O automata in HOLCF. The distribution contains simulation relations, temporal logic, and an abstraction theory. Everything is based upon a domain-theoretic model of finite and infinite sequences. " theories Abstraction session "IOA-ABP" in "HOLCF/IOA/ABP" = IOA + description " Author: Olaf Mueller The Alternating Bit Protocol performed in I/O-Automata. " theories Correctness Spec session "IOA-NTP" in "HOLCF/IOA/NTP" = IOA + description " Author: Tobias Nipkow & Konrad Slind A network transmission protocol, performed in the I/O automata formalization by Olaf Mueller. " theories Correctness session "IOA-Storage" in "HOLCF/IOA/Storage" = IOA + description " Author: Olaf Mueller Memory storage case study. " theories Correctness session "IOA-ex" in "HOLCF/IOA/ex" = IOA + description " Author: Olaf Mueller " theories TrivEx TrivEx2