diff --git a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy @@ -1,7843 +1,7843 @@ 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 continuous_on_const) lemma C1_differentiable_on_const [simp, derivative_intros]: "(\z. a) C1_differentiable_on S" by (auto simp: C1_differentiable_on_eq continuous_on_const) lemma C1_differentiable_on_add [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x + g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_minus [simp, derivative_intros]: "f C1_differentiable_on S \ (\x. - f x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_diff [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x - g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto intro: continuous_intros) lemma C1_differentiable_on_mult [simp, derivative_intros]: fixes f g :: "real \ 'a :: real_normed_algebra" shows "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x * g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (auto simp: continuous_on_add continuous_on_mult continuous_at_imp_continuous_on differentiable_imp_continuous_within) lemma C1_differentiable_on_scaleR [simp, derivative_intros]: "f C1_differentiable_on S \ g C1_differentiable_on S \ (\x. f x *\<^sub>R g x) C1_differentiable_on S" unfolding C1_differentiable_on_eq by (rule continuous_intros | simp add: continuous_at_imp_continuous_on differentiable_imp_continuous_within)+ definition\<^marker>\tag important\ piecewise_C1_differentiable_on (infixr "piecewise'_C1'_differentiable'_on" 50) where "f piecewise_C1_differentiable_on i \ continuous_on i f \ (\S. finite S \ (f C1_differentiable_on (i - S)))" lemma C1_differentiable_imp_piecewise: "f C1_differentiable_on S \ f piecewise_C1_differentiable_on S" by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_at_imp_continuous_on differentiable_imp_continuous_within) lemma piecewise_C1_imp_differentiable: "f piecewise_C1_differentiable_on i \ f piecewise_differentiable_on i" by (auto simp: piecewise_C1_differentiable_on_def piecewise_differentiable_on_def C1_differentiable_on_def differentiable_def has_vector_derivative_def intro: has_derivative_at_withinI) lemma piecewise_C1_differentiable_compose: assumes fg: "f piecewise_C1_differentiable_on S" "g piecewise_C1_differentiable_on (f ` S)" and fin: "\x. finite (S \ f-`{x})" shows "(g \ f) piecewise_C1_differentiable_on S" proof - have "continuous_on S (\x. g (f x))" by (metis continuous_on_compose2 fg order_refl piecewise_C1_differentiable_on_def) moreover have "\T. finite T \ g \ f C1_differentiable_on S - T" proof - obtain F where "finite F" and F: "f C1_differentiable_on S - F" and f: "f piecewise_C1_differentiable_on S" using fg by (auto simp: piecewise_C1_differentiable_on_def) obtain G where "finite G" and G: "g C1_differentiable_on f ` S - G" and g: "g piecewise_C1_differentiable_on f ` S" using fg by (auto simp: piecewise_C1_differentiable_on_def) show ?thesis proof (intro exI conjI) show "finite (F \ (\x\G. S \ f-`{x}))" using fin by (auto simp only: Int_Union \finite F\ \finite G\ finite_UN finite_imageI) show "g \ f C1_differentiable_on S - (F \ (\x\G. S \ f -` {x}))" apply (rule C1_differentiable_compose) apply (blast intro: C1_differentiable_on_subset [OF F]) apply (blast intro: C1_differentiable_on_subset [OF G]) by (simp add: C1_differentiable_on_subset G Diff_Int_distrib2 fin) qed qed ultimately show ?thesis by (simp add: piecewise_C1_differentiable_on_def) qed lemma piecewise_C1_differentiable_on_subset: "f piecewise_C1_differentiable_on S \ T \ S \ f piecewise_C1_differentiable_on T" by (auto simp: piecewise_C1_differentiable_on_def elim!: continuous_on_subset C1_differentiable_on_subset) lemma C1_differentiable_imp_continuous_on: "f C1_differentiable_on S \ continuous_on S f" unfolding C1_differentiable_on_eq continuous_on_eq_continuous_within using differentiable_at_withinI differentiable_imp_continuous_within by blast lemma C1_differentiable_on_empty [iff]: "f C1_differentiable_on {}" unfolding C1_differentiable_on_def by auto lemma piecewise_C1_differentiable_affine: fixes m::real assumes "f piecewise_C1_differentiable_on ((\x. m * x + c) ` S)" shows "(f \ (\x. m *\<^sub>R x + c)) piecewise_C1_differentiable_on S" proof (cases "m = 0") case True then show ?thesis unfolding o_def by (auto simp: piecewise_C1_differentiable_on_def continuous_on_const) next case False have *: "\x. finite (S \ {y. m * y + c = x})" using False not_finite_existsD by fastforce show ?thesis apply (rule piecewise_C1_differentiable_compose [OF C1_differentiable_imp_piecewise]) apply (rule * assms derivative_intros | simp add: False vimage_def)+ done qed lemma piecewise_C1_differentiable_cases: fixes c::real assumes "f piecewise_C1_differentiable_on {a..c}" "g piecewise_C1_differentiable_on {c..b}" "a \ c" "c \ b" "f c = g c" shows "(\x. if x \ c then f x else g x) piecewise_C1_differentiable_on {a..b}" proof - obtain S T where st: "f C1_differentiable_on ({a..c} - S)" "g C1_differentiable_on ({c..b} - T)" "finite S" "finite T" using assms by (force simp: piecewise_C1_differentiable_on_def) then have f_diff: "f differentiable_on {a..x. if x \ c then f x else g x)" using continuous_on_cases [OF closed_real_atLeastAtMost [of a c], OF closed_real_atLeastAtMost [of c b], of f g "\x. x\c"] assms by (force simp: ivl_disj_un_two_touch) { fix x assume x: "x \ {a..b} - insert c (S \ T)" have "(\x. if x \ c then f x else g x) differentiable at x" (is "?diff_fg") proof (cases x c rule: le_cases) case le show ?diff_fg apply (rule differentiable_transform_within [where f=f and d = "dist x c"]) using x dist_real_def le st by (auto simp: C1_differentiable_on_eq) next case ge show ?diff_fg apply (rule differentiable_transform_within [where f=g and d = "dist x c"]) using dist_nz x dist_real_def ge st x by (auto simp: C1_differentiable_on_eq) qed } then have "(\x \ {a..b} - insert c (S \ T). (\x. if x \ c then f x else g x) differentiable at x)" by auto moreover { assume fcon: "continuous_on ({a<..x. vector_derivative f (at x))" and gcon: "continuous_on ({c<..x. vector_derivative g (at x))" have "open ({a<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative f (at x)) (at x)" if "a < x" "x < c" "x \ S" for x proof - have f: "f differentiable at x" by (meson C1_differentiable_on_eq Diff_iff atLeastAtMost_iff less_eq_real_def st(1) that) show ?thesis using that apply (rule_tac f=f and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] f) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF fcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed moreover have "continuous_on ({c<..x. vector_derivative (\x. if x \ c then f x else g x) (at x))" proof - have "((\x. if x \ c then f x else g x) has_vector_derivative vector_derivative g (at x)) (at x)" if "c < x" "x < b" "x \ T" for x proof - have g: "g differentiable at x" by (metis C1_differentiable_on_eq DiffD1 DiffI atLeastAtMost_diff_ends greaterThanLessThan_iff st(2) that) show ?thesis using that apply (rule_tac f=g and d="dist x c" in has_vector_derivative_transform_within) apply (auto simp: dist_norm vector_derivative_works [symmetric] g) done qed then show ?thesis by (metis (no_types, lifting) continuous_on_eq [OF gcon] DiffE greaterThanLessThan_iff vector_derivative_at) qed ultimately have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" by (rule continuous_on_subset [OF continuous_on_open_Un], auto) } note * = this have "continuous_on ({a<.. T)) (\x. vector_derivative (\x. if x \ c then f x else g x) (at x))" using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset intro: *) ultimately have "\S. finite S \ ((\x. if x \ c then f x else g x) C1_differentiable_on {a..b} - S)" apply (rule_tac x="{a,b,c} \ S \ T" in exI) using st by (auto simp: C1_differentiable_on_eq elim!: continuous_on_subset) with cab show ?thesis by (simp add: piecewise_C1_differentiable_on_def) qed lemma piecewise_C1_differentiable_neg: "f piecewise_C1_differentiable_on S \ (\x. -(f x)) piecewise_C1_differentiable_on S" unfolding piecewise_C1_differentiable_on_def by (auto intro!: continuous_on_minus C1_differentiable_on_minus) lemma piecewise_C1_differentiable_add: assumes "f piecewise_C1_differentiable_on i" "g piecewise_C1_differentiable_on i" shows "(\x. f x + g x) piecewise_C1_differentiable_on i" proof - obtain S t where st: "finite S" "finite t" "f C1_differentiable_on (i-S)" "g C1_differentiable_on (i-t)" using assms by (auto simp: piecewise_C1_differentiable_on_def) then have "finite (S \ t) \ (\x. f x + g x) C1_differentiable_on i - (S \ t)" by (auto intro: C1_differentiable_on_add elim!: C1_differentiable_on_subset) moreover have "continuous_on i f" "continuous_on i g" using assms piecewise_C1_differentiable_on_def by auto ultimately show ?thesis by (auto simp: piecewise_C1_differentiable_on_def continuous_on_add) qed lemma piecewise_C1_differentiable_diff: "\f piecewise_C1_differentiable_on S; g piecewise_C1_differentiable_on S\ \ (\x. f x - g x) piecewise_C1_differentiable_on S" unfolding diff_conv_add_uminus by (metis piecewise_C1_differentiable_add piecewise_C1_differentiable_neg) lemma piecewise_C1_differentiable_D1: fixes g1 :: "real \ 'a::real_normed_field" assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" shows "g1 piecewise_C1_differentiable_on {0..1}" proof - obtain S where "finite S" and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have g1D: "g1 differentiable at x" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x proof (rule differentiable_transform_within) show "g1 +++ g2 \ (*) (inverse 2) differentiable at x" using that g12D apply (simp only: joinpaths_def) by (rule differentiable_chain_at derivative_intros | force)+ show "\x'. \dist x' x < dist (x/2) (1/2)\ \ (g1 +++ g2 \ (*) (inverse 2)) x' = g1 x'" using that by (auto simp: dist_real_def joinpaths_def) qed (use that in \auto simp: dist_real_def\) have [simp]: "vector_derivative (g1 \ (*) 2) (at (x/2)) = 2 *\<^sub>R vector_derivative g1 (at x)" if "x \ {0..1} - insert 1 ((*) 2 ` S)" for x apply (subst vector_derivative_chain_at) using that apply (rule derivative_eq_intros g1D | simp)+ done have "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" using co12 by (rule continuous_on_subset) force then have coDhalf: "continuous_on ({0..1/2} - insert (1/2) S) (\x. vector_derivative (g1 \ (*)2) (at x))" proof (rule continuous_on_eq [OF _ vector_derivative_at]) show "(g1 +++ g2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" if "x \ {0..1/2} - insert (1/2) S" for x proof (rule has_vector_derivative_transform_within) show "(g1 \ (*) 2 has_vector_derivative vector_derivative (g1 \ (*) 2) (at x)) (at x)" using that by (force intro: g1D differentiable_chain_at simp: vector_derivative_works [symmetric]) show "\x'. \dist x' x < dist x (1/2)\ \ (g1 \ (*) 2) x' = (g1 +++ g2) x'" using that by (auto simp: dist_norm joinpaths_def) qed (use that in \auto simp: dist_norm\) qed have "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) ((\x. 1/2 * vector_derivative (g1 \ (*)2) (at x)) \ (*)(1/2))" apply (rule continuous_intros)+ using coDhalf apply (simp add: scaleR_conv_of_real image_set_diff image_image) done then have con_g1: "continuous_on ({0..1} - insert 1 ((*) 2 ` S)) (\x. vector_derivative g1 (at x))" by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) have "continuous_on {0..1} g1" using continuous_on_joinpaths_D1 assms piecewise_C1_differentiable_on_def by blast with \finite S\ show ?thesis apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) apply (rule_tac x="insert 1 (((*)2)`S)" in exI) apply (simp add: g1D con_g1) done qed lemma piecewise_C1_differentiable_D2: fixes g2 :: "real \ 'a::real_normed_field" assumes "(g1 +++ g2) piecewise_C1_differentiable_on {0..1}" "pathfinish g1 = pathstart g2" shows "g2 piecewise_C1_differentiable_on {0..1}" proof - obtain S where "finite S" and co12: "continuous_on ({0..1} - S) (\x. vector_derivative (g1 +++ g2) (at x))" and g12D: "\x\{0..1} - S. g1 +++ g2 differentiable at x" using assms by (auto simp: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have g2D: "g2 differentiable at x" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x proof (rule differentiable_transform_within) show "g1 +++ g2 \ (\x. (x + 1) / 2) differentiable at x" using g12D that apply (simp only: joinpaths_def) apply (drule_tac x= "(x+1) / 2" in bspec, force simp: field_split_simps) apply (rule differentiable_chain_at derivative_intros | force)+ done show "\x'. dist x' x < dist ((x + 1) / 2) (1/2) \ (g1 +++ g2 \ (\x. (x + 1) / 2)) x' = g2 x'" using that by (auto simp: dist_real_def joinpaths_def field_simps) qed (use that in \auto simp: dist_norm\) have [simp]: "vector_derivative (g2 \ (\x. 2*x-1)) (at ((x+1)/2)) = 2 *\<^sub>R vector_derivative g2 (at x)" if "x \ {0..1} - insert 0 ((\x. 2*x-1) ` S)" for x using that by (auto simp: vector_derivative_chain_at field_split_simps g2D) have "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g1 +++ g2) (at x))" using co12 by (rule continuous_on_subset) force then have coDhalf: "continuous_on ({1/2..1} - insert (1/2) S) (\x. vector_derivative (g2 \ (\x. 2*x-1)) (at x))" proof (rule continuous_on_eq [OF _ vector_derivative_at]) show "(g1 +++ g2 has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) (at x)" if "x \ {1 / 2..1} - insert (1 / 2) S" for x proof (rule_tac f="g2 \ (\x. 2*x-1)" and d="dist (3/4) ((x+1)/2)" in has_vector_derivative_transform_within) show "(g2 \ (\x. 2 * x - 1) has_vector_derivative vector_derivative (g2 \ (\x. 2 * x - 1)) (at x)) (at x)" using that by (force intro: g2D differentiable_chain_at simp: vector_derivative_works [symmetric]) show "\x'. \dist x' x < dist (3 / 4) ((x + 1) / 2)\ \ (g2 \ (\x. 2 * x - 1)) x' = (g1 +++ g2) x'" using that by (auto simp: dist_norm joinpaths_def add_divide_distrib) qed (use that in \auto simp: dist_norm\) qed have [simp]: "((\x. (x+1) / 2) ` ({0..1} - insert 0 ((\x. 2 * x - 1) ` S))) = ({1/2..1} - insert (1/2) S)" apply (simp add: image_set_diff inj_on_def image_image) apply (auto simp: image_affinity_atLeastAtMost_div add_divide_distrib) done have "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) ((\x. 1/2 * vector_derivative (g2 \ (\x. 2*x-1)) (at x)) \ (\x. (x+1)/2))" by (rule continuous_intros | simp add: coDhalf)+ then have con_g2: "continuous_on ({0..1} - insert 0 ((\x. 2*x-1) ` S)) (\x. vector_derivative g2 (at x))" by (rule continuous_on_eq) (simp add: scaleR_conv_of_real) have "continuous_on {0..1} g2" using continuous_on_joinpaths_D2 assms piecewise_C1_differentiable_on_def by blast with \finite S\ show ?thesis apply (clarsimp simp add: piecewise_C1_differentiable_on_def C1_differentiable_on_eq) apply (rule_tac x="insert 0 ((\x. 2 * x - 1) ` S)" in exI) apply (simp add: g2D con_g2) done qed subsection \Valid paths, and their start and finish\ definition\<^marker>\tag important\ valid_path :: "(real \ 'a :: real_normed_vector) \ bool" where "valid_path f \ f piecewise_C1_differentiable_on {0..1::real}" definition closed_path :: "(real \ 'a :: real_normed_vector) \ bool" where "closed_path g \ g 0 = g 1" text\In particular, all results for paths apply\ lemma valid_path_imp_path: "valid_path g \ path g" by (simp add: path_def piecewise_C1_differentiable_on_def valid_path_def) lemma connected_valid_path_image: "valid_path g \ connected(path_image g)" by (metis connected_path_image valid_path_imp_path) lemma compact_valid_path_image: "valid_path g \ compact(path_image g)" by (metis compact_path_image valid_path_imp_path) lemma bounded_valid_path_image: "valid_path g \ bounded(path_image g)" by (metis bounded_path_image valid_path_imp_path) lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" by (metis closed_path_image valid_path_imp_path) lemma valid_path_compose: assumes "valid_path g" and der: "\x. x \ path_image g \ f field_differentiable (at x)" and con: "continuous_on (path_image g) (deriv f)" shows "valid_path (f \ g)" proof - obtain S where "finite S" and g_diff: "g C1_differentiable_on {0..1} - S" using \valid_path g\ unfolding valid_path_def piecewise_C1_differentiable_on_def by auto have "f \ g differentiable at t" when "t\{0..1} - S" for t proof (rule differentiable_chain_at) show "g differentiable at t" using \valid_path g\ by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - S\ that) next have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis then show "f differentiable at (g t)" using der[THEN field_differentiable_imp_differentiable] by auto qed moreover have "continuous_on ({0..1} - S) (\x. vector_derivative (f \ g) (at x))" proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], rule continuous_intros) show "continuous_on ({0..1} - S) (\x. vector_derivative g (at x))" using g_diff C1_differentiable_on_eq by auto next have "continuous_on {0..1} (\x. deriv f (g x))" using continuous_on_compose[OF _ con[unfolded path_image_def],unfolded comp_def] \valid_path g\ piecewise_C1_differentiable_on_def valid_path_def by blast then show "continuous_on ({0..1} - S) (\x. deriv f (g x))" using continuous_on_subset by blast next show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" when "t \ {0..1} - S" for t proof (rule vector_derivative_chain_at_general[symmetric]) show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) next have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis then show "f field_differentiable at (g t)" using der by auto qed qed ultimately have "f \ g C1_differentiable_on {0..1} - S" using C1_differentiable_on_eq by blast moreover have "path (f \ g)" apply (rule path_continuous_image[OF valid_path_imp_path[OF \valid_path g\]]) using der by (simp add: continuous_at_imp_continuous_on field_differentiable_imp_continuous_at) ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def using \finite S\ by auto qed lemma valid_path_uminus_comp[simp]: fixes g::"real \ 'a ::real_normed_field" shows "valid_path (uminus \ g) \ valid_path g" proof show "valid_path g \ valid_path (uminus \ g)" for g::"real \ 'a" by (auto intro!: valid_path_compose derivative_intros simp add: deriv_linear[of "-1",simplified]) then show "valid_path g" when "valid_path (uminus \ g)" by (metis fun.map_comp group_add_class.minus_comp_minus id_comp that) qed lemma valid_path_offset[simp]: shows "valid_path (\t. g t - z) \ valid_path g" proof show *: "valid_path (g::real\'a) \ valid_path (\t. g t - z)" for g z unfolding valid_path_def by (fastforce intro:derivative_intros C1_differentiable_imp_piecewise piecewise_C1_differentiable_diff) show "valid_path (\t. g t - z) \ valid_path g" using *[of "\t. g t - z" "-z",simplified] . qed subsection\Contour Integrals along a path\ text\This definition is for complex numbers only, and does not generalise to line integrals in a vector field\ text\piecewise differentiable function on [0,1]\ definition\<^marker>\tag important\ has_contour_integral :: "(complex \ complex) \ complex \ (real \ complex) \ bool" (infixr "has'_contour'_integral" 50) where "(f has_contour_integral i) g \ ((\x. f(g x) * vector_derivative g (at x within {0..1})) has_integral i) {0..1}" definition\<^marker>\tag important\ contour_integrable_on (infixr "contour'_integrable'_on" 50) where "f contour_integrable_on g \ \i. (f has_contour_integral i) g" definition\<^marker>\tag important\ contour_integral where "contour_integral g f \ SOME i. (f has_contour_integral i) g \ \ f contour_integrable_on g \ i=0" lemma not_integrable_contour_integral: "\ f contour_integrable_on g \ contour_integral g f = 0" unfolding contour_integrable_on_def contour_integral_def by blast lemma contour_integral_unique: "(f has_contour_integral i) g \ contour_integral g f = i" apply (simp add: contour_integral_def has_contour_integral_def contour_integrable_on_def) using has_integral_unique by blast lemma has_contour_integral_eqpath: "\(f has_contour_integral y) p; f contour_integrable_on \; contour_integral p f = contour_integral \ f\ \ (f has_contour_integral y) \" using contour_integrable_on_def contour_integral_unique by auto lemma has_contour_integral_integral: "f contour_integrable_on i \ (f has_contour_integral (contour_integral i f)) i" by (metis contour_integral_unique contour_integrable_on_def) lemma has_contour_integral_unique: "(f has_contour_integral i) g \ (f has_contour_integral j) g \ i = j" using has_integral_unique by (auto simp: has_contour_integral_def) lemma has_contour_integral_integrable: "(f has_contour_integral i) g \ f contour_integrable_on g" using contour_integrable_on_def by blast text\Show that we can forget about the localized derivative.\ lemma has_integral_localized_vector_derivative: "((\x. f (g x) * vector_derivative g (at x within {a..b})) has_integral i) {a..b} \ ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {a..b}" proof - have *: "{a..b} - {a,b} = interior {a..b}" by (simp add: atLeastAtMost_diff_ends) show ?thesis apply (rule has_integral_spike_eq [of "{a,b}"]) apply (auto simp: at_within_interior [of _ "{a..b}"]) done qed lemma integrable_on_localized_vector_derivative: "(\x. f (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b} \ (\x. f (g x) * vector_derivative g (at x)) integrable_on {a..b}" by (simp add: integrable_on_def has_integral_localized_vector_derivative) lemma has_contour_integral: "(f has_contour_integral i) g \ ((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" by (simp add: has_integral_localized_vector_derivative has_contour_integral_def) lemma contour_integrable_on: "f contour_integrable_on g \ (\t. f(g t) * vector_derivative g (at t)) integrable_on {0..1}" by (simp add: has_contour_integral integrable_on_def contour_integrable_on_def) subsection\<^marker>\tag unimportant\ \Reversing a path\ lemma valid_path_imp_reverse: assumes "valid_path g" shows "valid_path(reversepath g)" proof - obtain S where "finite S" and S: "g C1_differentiable_on ({0..1} - S)" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) then have "finite ((-) 1 ` S)" by auto moreover have "(reversepath g C1_differentiable_on ({0..1} - (-) 1 ` S))" unfolding reversepath_def apply (rule C1_differentiable_compose [of "\x::real. 1-x" _ g, unfolded o_def]) using S by (force simp: finite_vimageI inj_on_def C1_differentiable_on_eq continuous_on_const elim!: continuous_on_subset)+ ultimately show ?thesis using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def path_def [symmetric]) qed lemma valid_path_reversepath [simp]: "valid_path(reversepath g) \ valid_path g" using valid_path_imp_reverse by force lemma has_contour_integral_reversepath: assumes "valid_path g" and f: "(f has_contour_integral i) g" shows "(f has_contour_integral (-i)) (reversepath g)" proof - { fix S x assume xs: "g C1_differentiable_on ({0..1} - S)" "x \ (-) 1 ` S" "0 \ x" "x \ 1" have "vector_derivative (\x. g (1 - x)) (at x within {0..1}) = - vector_derivative g (at (1 - x) within {0..1})" proof - obtain f' where f': "(g has_vector_derivative f') (at (1 - x))" using xs by (force simp: has_vector_derivative_def C1_differentiable_on_def) have "(g \ (\x. 1 - x) has_vector_derivative -1 *\<^sub>R f') (at x)" by (intro vector_diff_chain_within has_vector_derivative_at_within [OF f'] derivative_eq_intros | simp)+ then have mf': "((\x. g (1 - x)) has_vector_derivative -f') (at x)" by (simp add: o_def) show ?thesis using xs by (auto simp: vector_derivative_at_within_ivl [OF mf'] vector_derivative_at_within_ivl [OF f']) qed } note * = this obtain S where S: "continuous_on {0..1} g" "finite S" "g C1_differentiable_on {0..1} - S" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (g (1 - x)) * vector_derivative g (at (1 - x) within {0..1}))) has_integral -i) {0..1}" using has_integral_affinity01 [where m= "-1" and c=1, OF f [unfolded has_contour_integral_def]] by (simp add: has_integral_neg) then show ?thesis using S apply (clarsimp simp: reversepath_def has_contour_integral_def) apply (rule_tac S = "(\x. 1 - x) ` S" in has_integral_spike_finite) apply (auto simp: *) done qed lemma contour_integrable_reversepath: "valid_path g \ f contour_integrable_on g \ f contour_integrable_on (reversepath g)" using has_contour_integral_reversepath contour_integrable_on_def by blast lemma contour_integrable_reversepath_eq: "valid_path g \ (f contour_integrable_on (reversepath g) \ f contour_integrable_on g)" using contour_integrable_reversepath valid_path_reversepath by fastforce lemma contour_integral_reversepath: assumes "valid_path g" shows "contour_integral (reversepath g) f = - (contour_integral g f)" proof (cases "f contour_integrable_on g") case True then show ?thesis by (simp add: assms contour_integral_unique has_contour_integral_integral has_contour_integral_reversepath) next case False then have "\ f contour_integrable_on (reversepath g)" by (simp add: assms contour_integrable_reversepath_eq) with False show ?thesis by (simp add: not_integrable_contour_integral) qed subsection\<^marker>\tag unimportant\ \Joining two paths together\ lemma valid_path_join: assumes "valid_path g1" "valid_path g2" "pathfinish g1 = pathstart g2" shows "valid_path(g1 +++ g2)" proof - have "g1 1 = g2 0" using assms by (auto simp: pathfinish_def pathstart_def) moreover have "(g1 \ (\x. 2*x)) piecewise_C1_differentiable_on {0..1/2}" apply (rule piecewise_C1_differentiable_compose) using assms apply (auto simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_joinpaths) apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) done moreover have "(g2 \ (\x. 2*x-1)) piecewise_C1_differentiable_on {1/2..1}" apply (rule piecewise_C1_differentiable_compose) using assms unfolding valid_path_def piecewise_C1_differentiable_on_def by (auto intro!: continuous_intros finite_vimageI [where h = "(\x. 2*x - 1)"] inj_onI simp: image_affinity_atLeastAtMost_diff continuous_on_joinpaths) ultimately show ?thesis apply (simp only: valid_path_def continuous_on_joinpaths joinpaths_def) apply (rule piecewise_C1_differentiable_cases) apply (auto simp: o_def) done qed lemma valid_path_join_D1: fixes g1 :: "real \ 'a::real_normed_field" shows "valid_path (g1 +++ g2) \ valid_path g1" unfolding valid_path_def by (rule piecewise_C1_differentiable_D1) lemma valid_path_join_D2: fixes g2 :: "real \ 'a::real_normed_field" shows "\valid_path (g1 +++ g2); pathfinish g1 = pathstart g2\ \ valid_path g2" unfolding valid_path_def by (rule piecewise_C1_differentiable_D2) lemma valid_path_join_eq [simp]: fixes g2 :: "real \ 'a::real_normed_field" shows "pathfinish g1 = pathstart g2 \ (valid_path(g1 +++ g2) \ valid_path g1 \ valid_path g2)" using valid_path_join_D1 valid_path_join_D2 valid_path_join by blast lemma has_contour_integral_join: assumes "(f has_contour_integral i1) g1" "(f has_contour_integral i2) g2" "valid_path g1" "valid_path g2" shows "(f has_contour_integral (i1 + i2)) (g1 +++ g2)" proof - obtain s1 s2 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" and s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have 1: "((\x. f (g1 x) * vector_derivative g1 (at x)) has_integral i1) {0..1}" and 2: "((\x. f (g2 x) * vector_derivative g2 (at x)) has_integral i2) {0..1}" using assms by (auto simp: has_contour_integral) have i1: "((\x. (2*f (g1 (2*x))) * vector_derivative g1 (at (2*x))) has_integral i1) {0..1/2}" and i2: "((\x. (2*f (g2 (2*x - 1))) * vector_derivative g2 (at (2*x - 1))) has_integral i2) {1/2..1}" using has_integral_affinity01 [OF 1, where m= 2 and c=0, THEN has_integral_cmul [where c=2]] has_integral_affinity01 [OF 2, where m= 2 and c="-1", THEN has_integral_cmul [where c=2]] by (simp_all only: image_affinity_atLeastAtMost_div_diff, simp_all add: scaleR_conv_of_real mult_ac) have g1: "\0 \ z; z*2 < 1; z*2 \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = 2 *\<^sub>R vector_derivative g1 (at (z*2))" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\z - 1/2\"]]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x" 2 _ g1, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s1 apply (auto simp: algebra_simps vector_derivative_works) done have g2: "\1 < z*2; z \ 1; z*2 - 1 \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at z) = 2 *\<^sub>R vector_derivative g2 (at (z*2 - 1))" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2 (2*x - 1))" and d = "\z - 1/2\"]]) apply (simp_all add: dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. 2*x - 1" 2 _ g2, simplified o_def]) apply (simp add: has_vector_derivative_def has_derivative_def bounded_linear_mult_left) using s2 apply (auto simp: algebra_simps vector_derivative_works) done have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i1) {0..1/2}" apply (rule has_integral_spike_finite [OF _ _ i1, of "insert (1/2) ((*)2 -` s1)"]) using s1 apply (force intro: finite_vimageI [where h = "(*)2"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g1) done moreover have "((\x. f ((g1 +++ g2) x) * vector_derivative (g1 +++ g2) (at x)) has_integral i2) {1/2..1}" apply (rule has_integral_spike_finite [OF _ _ i2, of "insert (1/2) ((\x. 2*x-1) -` s2)"]) using s2 apply (force intro: finite_vimageI [where h = "\x. 2*x-1"] inj_onI) apply (clarsimp simp add: joinpaths_def scaleR_conv_of_real mult_ac g2) done ultimately show ?thesis apply (simp add: has_contour_integral) apply (rule has_integral_combine [where c = "1/2"], auto) done qed lemma contour_integrable_joinI: assumes "f contour_integrable_on g1" "f contour_integrable_on g2" "valid_path g1" "valid_path g2" shows "f contour_integrable_on (g1 +++ g2)" using assms by (meson has_contour_integral_join contour_integrable_on_def) lemma contour_integrable_joinD1: assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g1" shows "f contour_integrable_on g1" proof - obtain s1 where s1: "finite s1" "\x\{0..1} - s1. g1 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. f ((g1 +++ g2) (x/2)) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" using assms apply (auto simp: contour_integrable_on) apply (drule integrable_on_subcbox [where a=0 and b="1/2"]) apply (auto intro: integrable_affinity [of _ 0 "1/2::real" "1/2" 0, simplified]) done then have *: "(\x. (f ((g1 +++ g2) (x/2))/2) * vector_derivative (g1 +++ g2) (at (x/2))) integrable_on {0..1}" by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g1: "\0 < z; z < 1; z \ s1\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2)) = 2 *\<^sub>R vector_derivative g1 (at z)" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g1(2*x))" and d = "\(z-1)/2\"]]) apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x*2" 2 _ g1, simplified o_def]) using s1 apply (auto simp: vector_derivative_works has_vector_derivative_def has_derivative_def bounded_linear_mult_left) done show ?thesis using s1 apply (auto simp: contour_integrable_on) apply (rule integrable_spike_finite [of "{0,1} \ s1", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g1) done qed lemma contour_integrable_joinD2: assumes "f contour_integrable_on (g1 +++ g2)" "valid_path g2" shows "f contour_integrable_on g2" proof - obtain s2 where s2: "finite s2" "\x\{0..1} - s2. g2 differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have "(\x. f ((g1 +++ g2) (x/2 + 1/2)) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" using assms apply (auto simp: contour_integrable_on) apply (drule integrable_on_subcbox [where a="1/2" and b=1], auto) apply (drule integrable_affinity [of _ "1/2::real" 1 "1/2" "1/2", simplified]) apply (simp add: image_affinity_atLeastAtMost_diff) done then have *: "(\x. (f ((g1 +++ g2) (x/2 + 1/2))/2) * vector_derivative (g1 +++ g2) (at (x/2 + 1/2))) integrable_on {0..1}" by (auto dest: integrable_cmul [where c="1/2"] simp: scaleR_conv_of_real) have g2: "\0 < z; z < 1; z \ s2\ \ vector_derivative (\x. if x*2 \ 1 then g1 (2*x) else g2 (2*x - 1)) (at (z/2+1/2)) = 2 *\<^sub>R vector_derivative g2 (at z)" for z apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g2(2*x-1))" and d = "\z/2\"]]) apply (simp_all add: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x*2-1" 2 _ g2, simplified o_def]) using s2 apply (auto simp: has_vector_derivative_def has_derivative_def bounded_linear_mult_left vector_derivative_works add_divide_distrib) done show ?thesis using s2 apply (auto simp: contour_integrable_on) apply (rule integrable_spike_finite [of "{0,1} \ s2", OF _ _ *]) apply (auto simp: joinpaths_def scaleR_conv_of_real g2) done qed lemma contour_integrable_join [simp]: shows "\valid_path g1; valid_path g2\ \ f contour_integrable_on (g1 +++ g2) \ f contour_integrable_on g1 \ f contour_integrable_on g2" using contour_integrable_joinD1 contour_integrable_joinD2 contour_integrable_joinI by blast lemma contour_integral_join [simp]: shows "\f contour_integrable_on g1; f contour_integrable_on g2; valid_path g1; valid_path g2\ \ contour_integral (g1 +++ g2) f = contour_integral g1 f + contour_integral g2 f" by (simp add: has_contour_integral_integral has_contour_integral_join contour_integral_unique) subsection\<^marker>\tag unimportant\ \Shifting the starting point of a (closed) path\ lemma shiftpath_alt_def: "shiftpath a f = (\x. if x \ 1-a then f (a + x) else f (a + x - 1))" by (auto simp: shiftpath_def) lemma valid_path_shiftpath [intro]: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "valid_path(shiftpath a g)" using assms apply (auto simp: valid_path_def shiftpath_alt_def) apply (rule piecewise_C1_differentiable_cases) apply (auto simp: algebra_simps) apply (rule piecewise_C1_differentiable_affine [of g 1 a, simplified o_def scaleR_one]) apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) apply (rule piecewise_C1_differentiable_affine [of g 1 "a-1", simplified o_def scaleR_one algebra_simps]) apply (auto simp: pathfinish_def pathstart_def elim: piecewise_C1_differentiable_on_subset) done lemma has_contour_integral_shiftpath: assumes f: "(f has_contour_integral i) g" "valid_path g" and a: "a \ {0..1}" shows "(f has_contour_integral i) (shiftpath a g)" proof - obtain s where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have *: "((\x. f (g x) * vector_derivative g (at x)) has_integral i) {0..1}" using assms by (auto simp: has_contour_integral) then have i: "i = integral {a..1} (\x. f (g x) * vector_derivative g (at x)) + integral {0..a} (\x. f (g x) * vector_derivative g (at x))" apply (rule has_integral_unique) apply (subst add.commute) apply (subst integral_combine) using assms * integral_unique by auto { fix x have "0 \ x \ x + a < 1 \ x \ (\x. x - a) ` s \ vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a))" unfolding shiftpath_def apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x))" and d = "dist(1-a) x"]]) apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x+a" 1 _ g, simplified o_def scaleR_one]) apply (intro derivative_eq_intros | simp)+ using g apply (drule_tac x="x+a" in bspec) using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) done } note vd1 = this { fix x have "1 < x + a \ x \ 1 \ x \ (\x. x - a + 1) ` s \ vector_derivative (shiftpath a g) (at x) = vector_derivative g (at (x + a - 1))" unfolding shiftpath_def apply (rule vector_derivative_at [OF has_vector_derivative_transform_within [where f = "(\x. g(a+x-1))" and d = "dist (1-a) x"]]) apply (auto simp: field_simps dist_real_def abs_if split: if_split_asm) apply (rule vector_diff_chain_at [of "\x. x+a-1" 1 _ g, simplified o_def scaleR_one]) apply (intro derivative_eq_intros | simp)+ using g apply (drule_tac x="x+a-1" in bspec) using a apply (auto simp: has_vector_derivative_def vector_derivative_works image_def add.commute) done } note vd2 = this have va1: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({a..1})" using * a by (fastforce intro: integrable_subinterval_real) have v0a: "(\x. f (g x) * vector_derivative g (at x)) integrable_on ({0..a})" apply (rule integrable_subinterval_real) using * a by auto have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) has_integral integral {a..1} (\x. f (g x) * vector_derivative g (at x))) {0..1 - a}" apply (rule has_integral_spike_finite [where S = "{1-a} \ (\x. x-a) ` s" and f = "\x. f(g(a+x)) * vector_derivative g (at(a+x))"]) using s apply blast using a apply (auto simp: algebra_simps vd1) apply (force simp: shiftpath_def add.commute) using has_integral_affinity [where m=1 and c=a, simplified, OF integrable_integral [OF va1]] apply (simp add: image_affinity_atLeastAtMost_diff [where m=1 and c=a, simplified] add.commute) done moreover have "((\x. f (shiftpath a g x) * vector_derivative (shiftpath a g) (at x)) has_integral integral {0..a} (\x. f (g x) * vector_derivative g (at x))) {1 - a..1}" apply (rule has_integral_spike_finite [where S = "{1-a} \ (\x. x-a+1) ` s" and f = "\x. f(g(a+x-1)) * vector_derivative g (at(a+x-1))"]) using s apply blast using a apply (auto simp: algebra_simps vd2) apply (force simp: shiftpath_def add.commute) using has_integral_affinity [where m=1 and c="a-1", simplified, OF integrable_integral [OF v0a]] apply (simp add: image_affinity_atLeastAtMost [where m=1 and c="1-a", simplified]) apply (simp add: algebra_simps) done ultimately show ?thesis using a by (auto simp: i has_contour_integral intro: has_integral_combine [where c = "1-a"]) qed lemma has_contour_integral_shiftpath_D: assumes "(f has_contour_integral i) (shiftpath a g)" "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "(f has_contour_integral i) g" proof - obtain s where s: "finite s" and g: "\x\{0..1} - s. g differentiable at x" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) { fix x assume x: "0 < x" "x < 1" "x \ s" then have gx: "g differentiable at x" using g by auto have "vector_derivative g (at x within {0..1}) = vector_derivative (shiftpath (1 - a) (shiftpath a g)) (at x within {0..1})" apply (rule vector_derivative_at_within_ivl [OF has_vector_derivative_transform_within_open [where f = "(shiftpath (1 - a) (shiftpath a g))" and S = "{0<..<1}-s"]]) using s g assms x apply (auto simp: finite_imp_closed open_Diff shiftpath_shiftpath at_within_interior [of _ "{0..1}"] vector_derivative_works [symmetric]) apply (rule differentiable_transform_within [OF gx, of "min x (1-x)"]) apply (auto simp: dist_real_def shiftpath_shiftpath abs_if split: if_split_asm) done } note vd = this have fi: "(f has_contour_integral i) (shiftpath (1 - a) (shiftpath a g))" using assms by (auto intro!: has_contour_integral_shiftpath) show ?thesis apply (simp add: has_contour_integral_def) apply (rule has_integral_spike_finite [of "{0,1} \ s", OF _ _ fi [unfolded has_contour_integral_def]]) using s assms vd apply (auto simp: Path_Connected.shiftpath_shiftpath) done qed lemma has_contour_integral_shiftpath_eq: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "(f has_contour_integral i) (shiftpath a g) \ (f has_contour_integral i) g" using assms has_contour_integral_shiftpath has_contour_integral_shiftpath_D by blast lemma contour_integrable_on_shiftpath_eq: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "f contour_integrable_on (shiftpath a g) \ f contour_integrable_on g" using assms contour_integrable_on_def has_contour_integral_shiftpath_eq by auto lemma contour_integral_shiftpath: assumes "valid_path g" "pathfinish g = pathstart g" "a \ {0..1}" shows "contour_integral (shiftpath a g) f = contour_integral g f" using assms by (simp add: contour_integral_def contour_integrable_on_def has_contour_integral_shiftpath_eq) subsection\<^marker>\tag unimportant\ \More about straight-line paths\ lemma has_vector_derivative_linepath_within: "(linepath a b has_vector_derivative (b - a)) (at x within s)" apply (simp add: linepath_def has_vector_derivative_def algebra_simps) apply (rule derivative_eq_intros | simp)+ done lemma vector_derivative_linepath_within: "x \ {0..1} \ vector_derivative (linepath a b) (at x within {0..1}) = b - a" apply (rule vector_derivative_within_cbox [of 0 "1::real", simplified]) apply (auto simp: has_vector_derivative_linepath_within) done lemma vector_derivative_linepath_at [simp]: "vector_derivative (linepath a b) (at x) = b - a" by (simp add: has_vector_derivative_linepath_within vector_derivative_at) lemma valid_path_linepath [iff]: "valid_path (linepath a b)" apply (simp add: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq continuous_on_linepath) apply (rule_tac x="{}" in exI) apply (simp add: differentiable_on_def differentiable_def) using has_vector_derivative_def has_vector_derivative_linepath_within apply (fastforce simp add: continuous_on_eq_continuous_within) done lemma has_contour_integral_linepath: shows "(f has_contour_integral i) (linepath a b) \ ((\x. f(linepath a b x) * (b - a)) has_integral i) {0..1}" by (simp add: has_contour_integral vector_derivative_linepath_at) lemma linepath_in_path: shows "x \ {0..1} \ linepath a b x \ closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_image_01: "linepath a b ` {0..1} = closed_segment a b" by (auto simp: segment linepath_def) lemma linepath_in_convex_hull: fixes x::real assumes a: "a \ convex hull s" and b: "b \ convex hull s" and x: "0\x" "x\1" shows "linepath a b x \ convex hull s" apply (rule closed_segment_subset_convex_hull [OF a b, THEN subsetD]) using x apply (auto simp: linepath_image_01 [symmetric]) done lemma Re_linepath: "Re(linepath (of_real a) (of_real b) x) = (1 - x)*a + x*b" by (simp add: linepath_def) lemma Im_linepath: "Im(linepath (of_real a) (of_real b) x) = 0" by (simp add: linepath_def) lemma has_contour_integral_trivial [iff]: "(f has_contour_integral 0) (linepath a a)" by (simp add: has_contour_integral_linepath) lemma has_contour_integral_trivial_iff [simp]: "(f has_contour_integral i) (linepath a a) \ i=0" using has_contour_integral_unique by blast lemma contour_integral_trivial [simp]: "contour_integral (linepath a a) f = 0" using has_contour_integral_trivial contour_integral_unique by blast lemma differentiable_linepath [intro]: "linepath a b differentiable at x within A" by (auto simp: linepath_def) lemma bounded_linear_linepath: assumes "bounded_linear f" shows "f (linepath a b x) = linepath (f a) (f b) x" proof - interpret f: bounded_linear f by fact show ?thesis by (simp add: linepath_def f.add f.scale) qed lemma bounded_linear_linepath': assumes "bounded_linear f" shows "f \ linepath a b = linepath (f a) (f b)" using bounded_linear_linepath[OF assms] by (simp add: fun_eq_iff) lemma cnj_linepath: "cnj (linepath a b x) = linepath (cnj a) (cnj b) x" by (simp add: linepath_def) lemma cnj_linepath': "cnj \ linepath a b = linepath (cnj a) (cnj b)" by (simp add: linepath_def fun_eq_iff) subsection\Relation to subpath construction\ lemma valid_path_subpath: fixes g :: "real \ 'a :: real_normed_vector" assumes "valid_path g" "u \ {0..1}" "v \ {0..1}" shows "valid_path(subpath u v g)" proof (cases "v=u") case True then show ?thesis unfolding valid_path_def subpath_def by (force intro: C1_differentiable_on_const C1_differentiable_imp_piecewise) next case False have "(g \ (\x. ((v-u) * x + u))) piecewise_C1_differentiable_on {0..1}" apply (rule piecewise_C1_differentiable_compose) apply (simp add: C1_differentiable_imp_piecewise) apply (simp add: image_affinity_atLeastAtMost) using assms False apply (auto simp: algebra_simps valid_path_def piecewise_C1_differentiable_on_subset) apply (subst Int_commute) apply (auto simp: inj_on_def algebra_simps crossproduct_eq finite_vimage_IntI) done then show ?thesis by (auto simp: o_def valid_path_def subpath_def) qed lemma has_contour_integral_subpath_refl [iff]: "(f has_contour_integral 0) (subpath u u g)" by (simp add: has_contour_integral subpath_def) lemma contour_integrable_subpath_refl [iff]: "f contour_integrable_on (subpath u u g)" using has_contour_integral_subpath_refl contour_integrable_on_def by blast lemma contour_integral_subpath_refl [simp]: "contour_integral (subpath u u g) f = 0" by (simp add: has_contour_integral_subpath_refl contour_integral_unique) lemma has_contour_integral_subpath: assumes f: "f contour_integrable_on g" and g: "valid_path g" and uv: "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(f has_contour_integral integral {u..v} (\x. f(g x) * vector_derivative g (at x))) (subpath u v g)" proof (cases "v=u") case True then show ?thesis using f by (simp add: contour_integrable_on_def subpath_def has_contour_integral) next case False obtain s where s: "\x. x \ {0..1} - s \ g differentiable at x" and fs: "finite s" using g unfolding piecewise_C1_differentiable_on_def C1_differentiable_on_eq valid_path_def by blast have *: "((\x. f (g ((v - u) * x + u)) * vector_derivative g (at ((v - u) * x + u))) has_integral (1 / (v - u)) * integral {u..v} (\t. f (g t) * vector_derivative g (at t))) {0..1}" using f uv apply (simp add: contour_integrable_on subpath_def has_contour_integral) apply (drule integrable_on_subcbox [where a=u and b=v, simplified]) apply (simp_all add: has_integral_integral) apply (drule has_integral_affinity [where m="v-u" and c=u, simplified]) apply (simp_all add: False image_affinity_atLeastAtMost_div_diff scaleR_conv_of_real) apply (simp add: divide_simps False) done { fix x have "x \ {0..1} \ x \ (\t. (v-u) *\<^sub>R t + u) -` s \ vector_derivative (\x. g ((v-u) * x + u)) (at x) = (v-u) *\<^sub>R vector_derivative g (at ((v-u) * x + u))" apply (rule vector_derivative_at [OF vector_diff_chain_at [simplified o_def]]) apply (intro derivative_eq_intros | simp)+ apply (cut_tac s [of "(v - u) * x + u"]) using uv mult_left_le [of x "v-u"] apply (auto simp: vector_derivative_works) done } note vd = this show ?thesis apply (cut_tac has_integral_cmul [OF *, where c = "v-u"]) using fs assms apply (simp add: False subpath_def has_contour_integral) apply (rule_tac S = "(\t. ((v-u) *\<^sub>R t + u)) -` s" in has_integral_spike_finite) apply (auto simp: inj_on_def False finite_vimageI vd scaleR_conv_of_real) done qed lemma contour_integrable_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" shows "f contour_integrable_on (subpath u v g)" apply (cases u v rule: linorder_class.le_cases) apply (metis contour_integrable_on_def has_contour_integral_subpath [OF assms]) apply (subst reversepath_subpath [symmetric]) apply (rule contour_integrable_reversepath) using assms apply (blast intro: valid_path_subpath) apply (simp add: contour_integrable_on_def) using assms apply (blast intro: has_contour_integral_subpath) done lemma has_integral_contour_integral_subpath: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "(((\x. f(g x) * vector_derivative g (at x))) has_integral contour_integral (subpath u v g) f) {u..v}" using assms apply (auto simp: has_integral_integrable_integral) apply (rule integrable_on_subcbox [where a=u and b=v and S = "{0..1}", simplified]) apply (auto simp: contour_integral_unique [OF has_contour_integral_subpath] contour_integrable_on) done lemma contour_integral_subcontour_integral: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "u \ v" shows "contour_integral (subpath u v g) f = integral {u..v} (\x. f(g x) * vector_derivative g (at x))" using assms has_contour_integral_subpath contour_integral_unique by blast lemma contour_integral_subpath_combine_less: assumes "f contour_integrable_on g" "valid_path g" "u \ {0..1}" "v \ {0..1}" "w \ {0..1}" "u {0..1}" "v \ {0..1}" "w \ {0..1}" shows "contour_integral (subpath u v g) f + contour_integral (subpath v w g) f = contour_integral (subpath u w g) f" proof (cases "u\v \ v\w \ u\w") case True have *: "subpath v u g = reversepath(subpath u v g) \ subpath w u g = reversepath(subpath u w g) \ subpath w v g = reversepath(subpath v w g)" by (auto simp: reversepath_subpath) have "u < v \ v < w \ u < w \ w < v \ v < u \ u < w \ v < w \ w < u \ w < u \ u < v \ w < v \ v < u" using True assms by linarith with assms show ?thesis using contour_integral_subpath_combine_less [of f g u v w] contour_integral_subpath_combine_less [of f g u w v] contour_integral_subpath_combine_less [of f g v u w] contour_integral_subpath_combine_less [of f g v w u] contour_integral_subpath_combine_less [of f g w u v] contour_integral_subpath_combine_less [of f g w v u] apply simp apply (elim disjE) apply (auto simp: * contour_integral_reversepath contour_integrable_subpath valid_path_reversepath valid_path_subpath algebra_simps) done next case False then show ?thesis apply (auto simp: contour_integral_subpath_refl) using assms by (metis eq_neg_iff_add_eq_0 contour_integrable_subpath contour_integral_reversepath reversepath_subpath valid_path_subpath) qed lemma contour_integral_integral: "contour_integral g f = integral {0..1} (\x. f (g x) * vector_derivative g (at x))" by (simp add: contour_integral_def integral_def has_contour_integral contour_integrable_on) lemma contour_integral_cong: assumes "g = g'" "\x. x \ path_image g \ f x = f' x" shows "contour_integral g f = contour_integral g' f'" unfolding contour_integral_integral using assms by (intro integral_cong) (auto simp: path_image_def) text \Contour integral along a segment on the real axis\ lemma has_contour_integral_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f has_contour_integral I) (linepath a b) \ ((\x. f (of_real x)) has_integral I) {Re a..Re b}" proof - from assms have [simp]: "of_real (Re a) = a" "of_real (Re b) = b" by (simp_all add: complex_eq_iff) from assms have "a \ b" by auto have "((\x. f (of_real x)) has_integral I) (cbox (Re a) (Re b)) \ ((\x. f (a + b * of_real x - a * of_real x)) has_integral I /\<^sub>R (Re b - Re a)) {0..1}" by (subst has_integral_affinity_iff [of "Re b - Re a" _ "Re a", symmetric]) (insert assms, simp_all add: field_simps scaleR_conv_of_real) also have "(\x. f (a + b * of_real x - a * of_real x)) = (\x. (f (a + b * of_real x - a * of_real x) * (b - a)) /\<^sub>R (Re b - Re a))" using \a \ b\ by (auto simp: field_simps fun_eq_iff scaleR_conv_of_real) also have "(\ has_integral I /\<^sub>R (Re b - Re a)) {0..1} \ ((\x. f (linepath a b x) * (b - a)) has_integral I) {0..1}" using assms by (subst has_integral_cmul_iff) (auto simp: linepath_def scaleR_conv_of_real algebra_simps) also have "\ \ (f has_contour_integral I) (linepath a b)" unfolding has_contour_integral_def by (intro has_integral_cong) (simp add: vector_derivative_linepath_within) finally show ?thesis by simp qed lemma contour_integrable_linepath_Reals_iff: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "(f contour_integrable_on linepath a b) \ (\x. f (of_real x)) integrable_on {Re a..Re b}" using has_contour_integral_linepath_Reals_iff[OF assms, of f] by (auto simp: contour_integrable_on_def integrable_on_def) lemma contour_integral_linepath_Reals_eq: fixes a b :: complex and f :: "complex \ complex" assumes "a \ Reals" "b \ Reals" "Re a < Re b" shows "contour_integral (linepath a b) f = integral {Re a..Re b} (\x. f (of_real x))" proof (cases "f contour_integrable_on linepath a b") case True thus ?thesis using has_contour_integral_linepath_Reals_iff[OF assms, of f] using has_contour_integral_integral has_contour_integral_unique by blast next case False thus ?thesis using contour_integrable_linepath_Reals_iff[OF assms, of f] by (simp add: not_integrable_contour_integral not_integrable_integral) qed text\Cauchy's theorem where there's a primitive\ lemma contour_integral_primitive_lemma: fixes f :: "complex \ complex" and g :: "real \ complex" assumes "a \ b" and "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "g piecewise_differentiable_on {a..b}" "\x. x \ {a..b} \ g x \ s" shows "((\x. f'(g x) * vector_derivative g (at x within {a..b})) has_integral (f(g b) - f(g a))) {a..b}" proof - obtain k where k: "finite k" "\x\{a..b} - k. g differentiable (at x within {a..b})" and cg: "continuous_on {a..b} g" using assms by (auto simp: piecewise_differentiable_on_def) have cfg: "continuous_on {a..b} (\x. f (g x))" apply (rule continuous_on_compose [OF cg, unfolded o_def]) using assms apply (metis field_differentiable_def field_differentiable_imp_continuous_at continuous_on_eq_continuous_within continuous_on_subset image_subset_iff) done { fix x::real assume a: "a < x" and b: "x < b" and xk: "x \ k" then have "g differentiable at x within {a..b}" using k by (simp add: differentiable_at_withinI) then have "(g has_vector_derivative vector_derivative g (at x within {a..b})) (at x within {a..b})" by (simp add: vector_derivative_works has_field_derivative_def scaleR_conv_of_real) then have gdiff: "(g has_derivative (\u. u * vector_derivative g (at x within {a..b}))) (at x within {a..b})" by (simp add: has_vector_derivative_def scaleR_conv_of_real) have "(f has_field_derivative (f' (g x))) (at (g x) within g ` {a..b})" using assms by (metis a atLeastAtMost_iff b DERIV_subset image_subset_iff less_eq_real_def) then have fdiff: "(f has_derivative (*) (f' (g x))) (at (g x) within g ` {a..b})" by (simp add: has_field_derivative_def) have "((\x. f (g x)) has_vector_derivative f' (g x) * vector_derivative g (at x within {a..b})) (at x within {a..b})" using diff_chain_within [OF gdiff fdiff] by (simp add: has_vector_derivative_def scaleR_conv_of_real o_def mult_ac) } note * = this show ?thesis apply (rule fundamental_theorem_of_calculus_interior_strong) using k assms cfg * apply (auto simp: at_within_Icc_at) done qed lemma contour_integral_primitive: assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "valid_path g" "path_image g \ s" shows "(f' has_contour_integral (f(pathfinish g) - f(pathstart g))) g" using assms apply (simp add: valid_path_def path_image_def pathfinish_def pathstart_def has_contour_integral_def) apply (auto intro!: piecewise_C1_imp_differentiable contour_integral_primitive_lemma [of 0 1 s]) done corollary Cauchy_theorem_primitive: assumes "\x. x \ s \ (f has_field_derivative f' x) (at x within s)" and "valid_path g" "path_image g \ s" "pathfinish g = pathstart g" shows "(f' has_contour_integral 0) g" using assms by (metis diff_self contour_integral_primitive) text\Existence of path integral for continuous function\ lemma contour_integrable_continuous_linepath: assumes "continuous_on (closed_segment a b) f" shows "f contour_integrable_on (linepath a b)" proof - have "continuous_on {0..1} ((\x. f x * (b - a)) \ linepath a b)" apply (rule continuous_on_compose [OF continuous_on_linepath], simp add: linepath_image_01) apply (rule continuous_intros | simp add: assms)+ done then show ?thesis apply (simp add: contour_integrable_on_def has_contour_integral_def integrable_on_def [symmetric]) apply (rule integrable_continuous [of 0 "1::real", simplified]) apply (rule continuous_on_eq [where f = "\x. f(linepath a b x)*(b - a)"]) apply (auto simp: vector_derivative_linepath_within) done qed lemma has_field_der_id: "((\x. x\<^sup>2 / 2) has_field_derivative x) (at x)" by (rule has_derivative_imp_has_field_derivative) (rule derivative_intros | simp)+ lemma contour_integral_id [simp]: "contour_integral (linepath a b) (\y. y) = (b^2 - a^2)/2" apply (rule contour_integral_unique) using contour_integral_primitive [of UNIV "\x. x^2/2" "\x. x" "linepath a b"] apply (auto simp: field_simps has_field_der_id) done lemma contour_integrable_on_const [iff]: "(\x. c) contour_integrable_on (linepath a b)" by (simp add: continuous_on_const contour_integrable_continuous_linepath) lemma contour_integrable_on_id [iff]: "(\x. x) contour_integrable_on (linepath a b)" by (simp add: continuous_on_id contour_integrable_continuous_linepath) subsection\<^marker>\tag unimportant\ \Arithmetical combining theorems\ lemma has_contour_integral_neg: "(f has_contour_integral i) g \ ((\x. -(f x)) has_contour_integral (-i)) g" by (simp add: has_integral_neg has_contour_integral_def) lemma has_contour_integral_add: "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ \ ((\x. f1 x + f2 x) has_contour_integral (i1 + i2)) g" by (simp add: has_integral_add has_contour_integral_def algebra_simps) lemma has_contour_integral_diff: "\(f1 has_contour_integral i1) g; (f2 has_contour_integral i2) g\ \ ((\x. f1 x - f2 x) has_contour_integral (i1 - i2)) g" by (simp add: has_integral_diff has_contour_integral_def algebra_simps) lemma has_contour_integral_lmul: "(f has_contour_integral i) g \ ((\x. c * (f x)) has_contour_integral (c*i)) g" apply (simp add: has_contour_integral_def) apply (drule has_integral_mult_right) apply (simp add: algebra_simps) done lemma has_contour_integral_rmul: "(f has_contour_integral i) g \ ((\x. (f x) * c) has_contour_integral (i*c)) g" apply (drule has_contour_integral_lmul) apply (simp add: mult.commute) done lemma has_contour_integral_div: "(f has_contour_integral i) g \ ((\x. f x/c) has_contour_integral (i/c)) g" by (simp add: field_class.field_divide_inverse) (metis has_contour_integral_rmul) lemma has_contour_integral_eq: "\(f has_contour_integral y) p; \x. x \ path_image p \ f x = g x\ \ (g has_contour_integral y) p" apply (simp add: path_image_def has_contour_integral_def) by (metis (no_types, lifting) image_eqI has_integral_eq) lemma has_contour_integral_bound_linepath: assumes "(f has_contour_integral i) (linepath a b)" "0 \ B" "\x. x \ closed_segment a b \ norm(f x) \ B" shows "norm i \ B * norm(b - a)" proof - { fix x::real assume x: "0 \ x" "x \ 1" have "norm (f (linepath a b x)) * norm (vector_derivative (linepath a b) (at x within {0..1})) \ B * norm (b - a)" by (auto intro: mult_mono simp: assms linepath_in_path of_real_linepath vector_derivative_linepath_within x) } note * = this have "norm i \ (B * norm (b - a)) * content (cbox 0 (1::real))" apply (rule has_integral_bound [of _ "\x. f (linepath a b x) * vector_derivative (linepath a b) (at x within {0..1})"]) using assms * unfolding has_contour_integral_def apply (auto simp: norm_mult) done then show ?thesis by (auto simp: content_real) qed (*UNUSED lemma has_contour_integral_bound_linepath_strong: fixes a :: real and f :: "complex \ real" assumes "(f has_contour_integral i) (linepath a b)" "finite k" "0 \ B" "\x::real. x \ closed_segment a b - k \ norm(f x) \ B" shows "norm i \ B*norm(b - a)" *) lemma has_contour_integral_const_linepath: "((\x. c) has_contour_integral c*(b - a))(linepath a b)" unfolding has_contour_integral_linepath by (metis content_real diff_0_right has_integral_const_real lambda_one of_real_1 scaleR_conv_of_real zero_le_one) lemma has_contour_integral_0: "((\x. 0) has_contour_integral 0) g" by (simp add: has_contour_integral_def) lemma has_contour_integral_is_0: "(\z. z \ path_image g \ f z = 0) \ (f has_contour_integral 0) g" by (rule has_contour_integral_eq [OF has_contour_integral_0]) auto lemma has_contour_integral_sum: "\finite s; \a. a \ s \ (f a has_contour_integral i a) p\ \ ((\x. sum (\a. f a x) s) has_contour_integral sum i s) p" by (induction s rule: finite_induct) (auto simp: has_contour_integral_0 has_contour_integral_add) subsection\<^marker>\tag unimportant\ \Operations on path integrals\ lemma contour_integral_const_linepath [simp]: "contour_integral (linepath a b) (\x. c) = c*(b - a)" by (rule contour_integral_unique [OF has_contour_integral_const_linepath]) lemma contour_integral_neg: "f contour_integrable_on g \ contour_integral g (\x. -(f x)) = -(contour_integral g f)" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_neg) lemma contour_integral_add: "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x + f2 x) = contour_integral g f1 + contour_integral g f2" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_add) lemma contour_integral_diff: "f1 contour_integrable_on g \ f2 contour_integrable_on g \ contour_integral g (\x. f1 x - f2 x) = contour_integral g f1 - contour_integral g f2" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_diff) lemma contour_integral_lmul: shows "f contour_integrable_on g \ contour_integral g (\x. c * f x) = c*contour_integral g f" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_lmul) lemma contour_integral_rmul: shows "f contour_integrable_on g \ contour_integral g (\x. f x * c) = contour_integral g f * c" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_rmul) lemma contour_integral_div: shows "f contour_integrable_on g \ contour_integral g (\x. f x / c) = contour_integral g f / c" by (simp add: contour_integral_unique has_contour_integral_integral has_contour_integral_div) lemma contour_integral_eq: "(\x. x \ path_image p \ f x = g x) \ contour_integral p f = contour_integral p g" apply (simp add: contour_integral_def) using has_contour_integral_eq by (metis contour_integral_unique has_contour_integral_integrable has_contour_integral_integral) lemma contour_integral_eq_0: "(\z. z \ path_image g \ f z = 0) \ contour_integral g f = 0" by (simp add: has_contour_integral_is_0 contour_integral_unique) lemma contour_integral_bound_linepath: shows "\f contour_integrable_on (linepath a b); 0 \ B; \x. x \ closed_segment a b \ norm(f x) \ B\ \ norm(contour_integral (linepath a b) f) \ B*norm(b - a)" apply (rule has_contour_integral_bound_linepath [of f]) apply (auto simp: has_contour_integral_integral) done lemma contour_integral_0 [simp]: "contour_integral g (\x. 0) = 0" by (simp add: contour_integral_unique has_contour_integral_0) lemma contour_integral_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ contour_integral p (\x. sum (\a. f a x) s) = sum (\a. contour_integral p (f a)) s" by (auto simp: contour_integral_unique has_contour_integral_sum has_contour_integral_integral) lemma contour_integrable_eq: "\f contour_integrable_on p; \x. x \ path_image p \ f x = g x\ \ g contour_integrable_on p" unfolding contour_integrable_on_def by (metis has_contour_integral_eq) subsection\<^marker>\tag unimportant\ \Arithmetic theorems for path integrability\ lemma contour_integrable_neg: "f contour_integrable_on g \ (\x. -(f x)) contour_integrable_on g" using has_contour_integral_neg contour_integrable_on_def by blast lemma contour_integrable_add: "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x + f2 x) contour_integrable_on g" using has_contour_integral_add contour_integrable_on_def by fastforce lemma contour_integrable_diff: "\f1 contour_integrable_on g; f2 contour_integrable_on g\ \ (\x. f1 x - f2 x) contour_integrable_on g" using has_contour_integral_diff contour_integrable_on_def by fastforce lemma contour_integrable_lmul: "f contour_integrable_on g \ (\x. c * f x) contour_integrable_on g" using has_contour_integral_lmul contour_integrable_on_def by fastforce lemma contour_integrable_rmul: "f contour_integrable_on g \ (\x. f x * c) contour_integrable_on g" using has_contour_integral_rmul contour_integrable_on_def by fastforce lemma contour_integrable_div: "f contour_integrable_on g \ (\x. f x / c) contour_integrable_on g" using has_contour_integral_div contour_integrable_on_def by fastforce lemma contour_integrable_sum: "\finite s; \a. a \ s \ (f a) contour_integrable_on p\ \ (\x. sum (\a. f a x) s) contour_integrable_on p" unfolding contour_integrable_on_def by (metis has_contour_integral_sum) subsection\<^marker>\tag unimportant\ \Reversing a path integral\ lemma has_contour_integral_reverse_linepath: "(f has_contour_integral i) (linepath a b) \ (f has_contour_integral (-i)) (linepath b a)" using has_contour_integral_reversepath valid_path_linepath by fastforce lemma contour_integral_reverse_linepath: "continuous_on (closed_segment a b) f \ contour_integral (linepath a b) f = - (contour_integral(linepath b a) f)" apply (rule contour_integral_unique) apply (rule has_contour_integral_reverse_linepath) by (simp add: closed_segment_commute contour_integrable_continuous_linepath has_contour_integral_integral) (* Splitting a path integral in a flat way.*) lemma has_contour_integral_split: assumes f: "(f has_contour_integral i) (linepath a c)" "(f has_contour_integral j) (linepath c b)" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "(f has_contour_integral (i + j)) (linepath a b)" proof (cases "k = 0 \ k = 1") case True then show ?thesis using assms by auto next case False then have k: "0 < k" "k < 1" "complex_of_real k \ 1" using assms by auto have c': "c = k *\<^sub>R (b - a) + a" by (metis diff_add_cancel c) have bc: "(b - c) = (1 - k) *\<^sub>R (b - a)" by (simp add: algebra_simps c') { assume *: "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R c) * (c - a)) has_integral i) {0..1}" have **: "\x. ((k - x) / k) *\<^sub>R a + (x / k) *\<^sub>R c = (1 - x) *\<^sub>R a + x *\<^sub>R b" using False apply (simp add: c' algebra_simps) apply (simp add: real_vector.scale_left_distrib [symmetric] field_split_simps) done have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral i) {0..k}" using k has_integral_affinity01 [OF *, of "inverse k" "0"] apply (simp add: divide_simps mult.commute [of _ "k"] image_affinity_atLeastAtMost ** c) apply (auto dest: has_integral_cmul [where c = "inverse k"]) done } note fi = this { assume *: "((\x. f ((1 - x) *\<^sub>R c + x *\<^sub>R b) * (b - c)) has_integral j) {0..1}" have **: "\x. (((1 - x) / (1 - k)) *\<^sub>R c + ((x - k) / (1 - k)) *\<^sub>R b) = ((1 - x) *\<^sub>R a + x *\<^sub>R b)" using k apply (simp add: c' field_simps) apply (simp add: scaleR_conv_of_real divide_simps) apply (simp add: field_simps) done have "((\x. f ((1 - x) *\<^sub>R a + x *\<^sub>R b) * (b - a)) has_integral j) {k..1}" using k has_integral_affinity01 [OF *, of "inverse(1 - k)" "-(k/(1 - k))"] apply (simp add: divide_simps mult.commute [of _ "1-k"] image_affinity_atLeastAtMost ** bc) apply (auto dest: has_integral_cmul [where k = "(1 - k) *\<^sub>R j" and c = "inverse (1 - k)"]) done } note fj = this show ?thesis using f k apply (simp add: has_contour_integral_linepath) apply (simp add: linepath_def) apply (rule has_integral_combine [OF _ _ fi fj], simp_all) done qed lemma continuous_on_closed_segment_transform: assumes f: "continuous_on (closed_segment a b) f" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "continuous_on (closed_segment a c) f" proof - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" using c by (simp add: algebra_simps) have "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) then show "continuous_on (closed_segment a c) f" by (rule continuous_on_subset [OF f]) qed lemma contour_integral_split: assumes f: "continuous_on (closed_segment a b) f" and k: "0 \ k" "k \ 1" and c: "c - a = k *\<^sub>R (b - a)" shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" proof - have c': "c = (1 - k) *\<^sub>R a + k *\<^sub>R b" using c by (simp add: algebra_simps) have "closed_segment a c \ closed_segment a b" by (metis c' ends_in_segment(1) in_segment(1) k subset_closed_segment) moreover have "closed_segment c b \ closed_segment a b" by (metis c' ends_in_segment(2) in_segment(1) k subset_closed_segment) ultimately have *: "continuous_on (closed_segment a c) f" "continuous_on (closed_segment c b) f" by (auto intro: continuous_on_subset [OF f]) show ?thesis by (rule contour_integral_unique) (meson "*" c contour_integrable_continuous_linepath has_contour_integral_integral has_contour_integral_split k) qed lemma contour_integral_split_linepath: assumes f: "continuous_on (closed_segment a b) f" and c: "c \ closed_segment a b" shows "contour_integral(linepath a b) f = contour_integral(linepath a c) f + contour_integral(linepath c b) f" using c by (auto simp: closed_segment_def algebra_simps intro!: contour_integral_split [OF f]) text\The special case of midpoints used in the main quadrisection\ lemma has_contour_integral_midpoint: assumes "(f has_contour_integral i) (linepath a (midpoint a b))" "(f has_contour_integral j) (linepath (midpoint a b) b)" shows "(f has_contour_integral (i + j)) (linepath a b)" apply (rule has_contour_integral_split [where c = "midpoint a b" and k = "1/2"]) using assms apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) done lemma contour_integral_midpoint: "continuous_on (closed_segment a b) f \ contour_integral (linepath a b) f = contour_integral (linepath a (midpoint a b)) f + contour_integral (linepath (midpoint a b) b) f" apply (rule contour_integral_split [where c = "midpoint a b" and k = "1/2"]) apply (auto simp: midpoint_def algebra_simps scaleR_conv_of_real) done text\A couple of special case lemmas that are useful below\ lemma triangle_linear_has_chain_integral: "((\x. m*x + d) has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" apply (rule Cauchy_theorem_primitive [of UNIV "\x. m/2 * x^2 + d*x"]) apply (auto intro!: derivative_eq_intros) done lemma has_chain_integral_chain_integral3: "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d) \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f = i" apply (subst contour_integral_unique [symmetric], assumption) apply (drule has_contour_integral_integrable) apply (simp add: valid_path_join) done lemma has_chain_integral_chain_integral4: "(f has_contour_integral i) (linepath a b +++ linepath b c +++ linepath c d +++ linepath d e) \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c d) f + contour_integral (linepath d e) f = i" apply (subst contour_integral_unique [symmetric], assumption) apply (drule has_contour_integral_integrable) apply (simp add: valid_path_join) done subsection\Reversing the order in a double path integral\ text\The condition is stronger than needed but it's often true in typical situations\ lemma fst_im_cbox [simp]: "cbox c d \ {} \ (fst ` cbox (a,c) (b,d)) = cbox a b" by (auto simp: cbox_Pair_eq) lemma snd_im_cbox [simp]: "cbox a b \ {} \ (snd ` cbox (a,c) (b,d)) = cbox c d" by (auto simp: cbox_Pair_eq) proposition contour_integral_swap: assumes fcon: "continuous_on (path_image g \ path_image h) (\(y1,y2). f y1 y2)" and vp: "valid_path g" "valid_path h" and gvcon: "continuous_on {0..1} (\t. vector_derivative g (at t))" and hvcon: "continuous_on {0..1} (\t. vector_derivative h (at t))" shows "contour_integral g (\w. contour_integral h (f w)) = contour_integral h (\z. contour_integral g (\w. f w z))" proof - have gcon: "continuous_on {0..1} g" and hcon: "continuous_on {0..1} h" using assms by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have fgh1: "\x. (\t. f (g x) (h t)) = (\(y1,y2). f y1 y2) \ (\t. (g x, h t))" by (rule ext) simp have fgh2: "\x. (\t. f (g t) (h x)) = (\(y1,y2). f y1 y2) \ (\t. (g t, h x))" by (rule ext) simp have fcon_im1: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g x, h t)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have fcon_im2: "\x. 0 \ x \ x \ 1 \ continuous_on ((\t. (g t, h x)) ` {0..1}) (\(x, y). f x y)" by (rule continuous_on_subset [OF fcon]) (auto simp: path_image_def) have "\y. y \ {0..1} \ continuous_on {0..1} (\x. f (g x) (h y))" by (subst fgh2) (rule fcon_im2 gcon continuous_intros | simp)+ then have vdg: "\y. y \ {0..1} \ (\x. f (g x) (h y) * vector_derivative g (at x)) integrable_on {0..1}" using continuous_on_mult gvcon integrable_continuous_real by blast have "(\z. vector_derivative g (at (fst z))) = (\x. vector_derivative g (at x)) \ fst" by auto then have gvcon': "continuous_on (cbox (0, 0) (1, 1::real)) (\x. vector_derivative g (at (fst x)))" apply (rule ssubst) apply (rule continuous_intros | simp add: gvcon)+ done have "(\z. vector_derivative h (at (snd z))) = (\x. vector_derivative h (at x)) \ snd" by auto then have hvcon': "continuous_on (cbox (0, 0) (1::real, 1)) (\x. vector_derivative h (at (snd x)))" apply (rule ssubst) apply (rule continuous_intros | simp add: hvcon)+ done have "(\x. f (g (fst x)) (h (snd x))) = (\(y1,y2). f y1 y2) \ (\w. ((g \ fst) w, (h \ snd) w))" by auto then have fgh: "continuous_on (cbox (0, 0) (1, 1)) (\x. f (g (fst x)) (h (snd x)))" apply (rule ssubst) apply (rule gcon hcon continuous_intros | simp)+ apply (auto simp: path_image_def intro: continuous_on_subset [OF fcon]) done have "integral {0..1} (\x. contour_integral h (f (g x)) * vector_derivative g (at x)) = integral {0..1} (\x. contour_integral h (\y. f (g x) y * vector_derivative g (at x)))" proof (rule integral_cong [OF contour_integral_rmul [symmetric]]) show "\x. x \ {0..1} \ f (g x) contour_integrable_on h" unfolding contour_integrable_on apply (rule integrable_continuous_real) apply (rule continuous_on_mult [OF _ hvcon]) apply (subst fgh1) apply (rule fcon_im1 hcon continuous_intros | simp)+ done qed also have "\ = integral {0..1} (\y. contour_integral g (\x. f x (h y) * vector_derivative h (at y)))" unfolding contour_integral_integral apply (subst integral_swap_continuous [where 'a = real and 'b = real, of 0 0 1 1, simplified]) apply (rule fgh gvcon' hvcon' continuous_intros | simp add: split_def)+ unfolding integral_mult_left [symmetric] apply (simp only: mult_ac) done also have "\ = contour_integral h (\z. contour_integral g (\w. f w z))" unfolding contour_integral_integral apply (rule integral_cong) unfolding integral_mult_left [symmetric] apply (simp add: algebra_simps) done finally show ?thesis by (simp add: contour_integral_integral) qed subsection\<^marker>\tag unimportant\ \The key quadrisection step\ lemma norm_sum_half: assumes "norm(a + b) \ e" shows "norm a \ e/2 \ norm b \ e/2" proof - have "e \ norm (- a - b)" by (simp add: add.commute assms norm_minus_commute) thus ?thesis using norm_triangle_ineq4 order_trans by fastforce qed lemma norm_sum_lemma: assumes "e \ norm (a + b + c + d)" shows "e / 4 \ norm a \ e / 4 \ norm b \ e / 4 \ norm c \ e / 4 \ norm d" proof - have "e \ norm ((a + b) + (c + d))" using assms by (simp add: algebra_simps) then show ?thesis by (auto dest!: norm_sum_half) qed lemma Cauchy_theorem_quadrisection: assumes f: "continuous_on (convex hull {a,b,c}) f" and dist: "dist a b \ K" "dist b c \ K" "dist c a \ K" and e: "e * K^2 \ norm (contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f)" shows "\a' b' c'. a' \ convex hull {a,b,c} \ b' \ convex hull {a,b,c} \ c' \ convex hull {a,b,c} \ dist a' b' \ K/2 \ dist b' c' \ K/2 \ dist c' a' \ K/2 \ e * (K/2)^2 \ norm(contour_integral(linepath a' b') f + contour_integral(linepath b' c') f + contour_integral(linepath c' a') f)" (is "\x y z. ?\ x y z") proof - note divide_le_eq_numeral1 [simp del] define a' where "a' = midpoint b c" define b' where "b' = midpoint c a" define c' where "c' = midpoint a b" have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ have fcont': "continuous_on (closed_segment c' b') f" "continuous_on (closed_segment a' c') f" "continuous_on (closed_segment b' a') f" unfolding a'_def b'_def c'_def by (rule continuous_on_subset [OF f], metis midpoints_in_convex_hull convex_hull_subset hull_subset insert_subset segment_convex_hull)+ let ?pathint = "\x y. contour_integral(linepath x y) f" have *: "?pathint a b + ?pathint b c + ?pathint c a = (?pathint a c' + ?pathint c' b' + ?pathint b' a) + (?pathint a' c' + ?pathint c' b + ?pathint b a') + (?pathint a' c + ?pathint c b' + ?pathint b' a') + (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" by (simp add: fcont' contour_integral_reverse_linepath) (simp add: a'_def b'_def c'_def contour_integral_midpoint fabc) have [simp]: "\x y. cmod (x * 2 - y * 2) = cmod (x - y) * 2" by (metis left_diff_distrib mult.commute norm_mult_numeral1) have [simp]: "\x y. cmod (x - y) = cmod (y - x)" by (simp add: norm_minus_commute) consider "e * K\<^sup>2 / 4 \ cmod (?pathint a c' + ?pathint c' b' + ?pathint b' a)" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' c' + ?pathint c' b + ?pathint b a')" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' c + ?pathint c b' + ?pathint b' a')" | "e * K\<^sup>2 / 4 \ cmod (?pathint a' b' + ?pathint b' c' + ?pathint c' a')" using assms unfolding * by (blast intro: that dest!: norm_sum_lemma) then show ?thesis proof cases case 1 then have "?\ a c' b'" using assms apply (clarsimp simp: c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 2 then have "?\ a' c' b" using assms apply (clarsimp simp: a'_def c'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 3 then have "?\ a' c b'" using assms apply (clarsimp simp: a'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast next case 4 then have "?\ a' b' c'" using assms apply (clarsimp simp: a'_def c'_def b'_def midpoints_in_convex_hull hull_subset [THEN subsetD]) apply (auto simp: midpoint_def dist_norm scaleR_conv_of_real field_split_simps) done then show ?thesis by blast qed qed subsection\<^marker>\tag unimportant\ \Cauchy's theorem for triangles\ lemma triangle_points_closer: fixes a::complex shows "\x \ convex hull {a,b,c}; y \ convex hull {a,b,c}\ \ norm(x - y) \ norm(a - b) \ norm(x - y) \ norm(b - c) \ norm(x - y) \ norm(c - a)" using simplex_extremal_le [of "{a,b,c}"] by (auto simp: norm_minus_commute) lemma holomorphic_point_small_triangle: assumes x: "x \ S" and f: "continuous_on S f" and cd: "f field_differentiable (at x within S)" and e: "0 < e" shows "\k>0. \a b c. dist a b \ k \ dist b c \ k \ dist c a \ k \ x \ convex hull {a,b,c} \ convex hull {a,b,c} \ S \ norm(contour_integral(linepath a b) f + contour_integral(linepath b c) f + contour_integral(linepath c a) f) \ e*(dist a b + dist b c + dist c a)^2" (is "\k>0. \a b c. _ \ ?normle a b c") proof - have le_of_3: "\a x y z. \0 \ x*y; 0 \ x*z; 0 \ y*z; a \ (e*(x + y + z))*x + (e*(x + y + z))*y + (e*(x + y + z))*z\ \ a \ e*(x + y + z)^2" by (simp add: algebra_simps power2_eq_square) have disj_le: "\x \ a \ x \ b \ x \ c; 0 \ a; 0 \ b; 0 \ c\ \ x \ a + b + c" for x::real and a b c by linarith have fabc: "f contour_integrable_on linepath a b" "f contour_integrable_on linepath b c" "f contour_integrable_on linepath c a" if "convex hull {a, b, c} \ S" for a b c using segments_subset_convex_hull that by (metis continuous_on_subset f contour_integrable_continuous_linepath)+ note path_bound = has_contour_integral_bound_linepath [simplified norm_minus_commute, OF has_contour_integral_integral] { fix f' a b c d assume d: "0 < d" and f': "\y. \cmod (y - x) \ d; y \ S\ \ cmod (f y - f x - f' * (y - x)) \ e * cmod (y - x)" and le: "cmod (a - b) \ d" "cmod (b - c) \ d" "cmod (c - a) \ d" and xc: "x \ convex hull {a, b, c}" and S: "convex hull {a, b, c} \ S" have pa: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = contour_integral (linepath a b) (\y. f y - f x - f'*(y - x)) + contour_integral (linepath b c) (\y. f y - f x - f'*(y - x)) + contour_integral (linepath c a) (\y. f y - f x - f'*(y - x))" apply (simp add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc [OF S]) apply (simp add: field_simps) done { fix y assume yc: "y \ convex hull {a,b,c}" have "cmod (f y - f x - f' * (y - x)) \ e*norm(y - x)" proof (rule f') show "cmod (y - x) \ d" by (metis triangle_points_closer [OF xc yc] le norm_minus_commute order_trans) qed (use S yc in blast) also have "\ \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" by (simp add: yc e xc disj_le [OF triangle_points_closer]) finally have "cmod (f y - f x - f' * (y - x)) \ e * (cmod (a - b) + cmod (b - c) + cmod (c - a))" . } note cm_le = this have "?normle a b c" unfolding dist_norm pa apply (rule le_of_3) using f' xc S e apply simp_all apply (intro norm_triangle_le add_mono path_bound) apply (simp_all add: contour_integral_diff contour_integral_lmul contour_integrable_lmul contour_integrable_diff fabc) apply (blast intro: cm_le elim: dest: segments_subset_convex_hull [THEN subsetD])+ done } note * = this show ?thesis using cd e apply (simp add: field_differentiable_def has_field_derivative_def has_derivative_within_alt approachable_lt_le2 Ball_def) apply (clarify dest!: spec mp) using * unfolding dist_norm apply blast done qed text\Hence the most basic theorem for a triangle.\ locale Chain = fixes x0 At Follows assumes At0: "At x0 0" and AtSuc: "\x n. At x n \ \x'. At x' (Suc n) \ Follows x' x" begin primrec f where "f 0 = x0" | "f (Suc n) = (SOME x. At x (Suc n) \ Follows x (f n))" lemma At: "At (f n) n" proof (induct n) case 0 show ?case by (simp add: At0) next case (Suc n) show ?case by (metis (no_types, lifting) AtSuc [OF Suc] f.simps(2) someI_ex) qed lemma Follows: "Follows (f(Suc n)) (f n)" by (metis (no_types, lifting) AtSuc [OF At [of n]] f.simps(2) someI_ex) declare f.simps(2) [simp del] end lemma Chain3: assumes At0: "At x0 y0 z0 0" and AtSuc: "\x y z n. At x y z n \ \x' y' z'. At x' y' z' (Suc n) \ Follows x' y' z' x y z" obtains f g h where "f 0 = x0" "g 0 = y0" "h 0 = z0" "\n. At (f n) (g n) (h n) n" "\n. Follows (f(Suc n)) (g(Suc n)) (h(Suc n)) (f n) (g n) (h n)" proof - interpret three: Chain "(x0,y0,z0)" "\(x,y,z). At x y z" "\(x',y',z'). \(x,y,z). Follows x' y' z' x y z" apply unfold_locales using At0 AtSuc by auto show ?thesis apply (rule that [of "\n. fst (three.f n)" "\n. fst (snd (three.f n))" "\n. snd (snd (three.f n))"]) using three.At three.Follows apply simp_all apply (simp_all add: split_beta') done qed proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle: assumes "f holomorphic_on (convex hull {a,b,c})" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof - have contf: "continuous_on (convex hull {a,b,c}) f" by (metis assms holomorphic_on_imp_continuous_on) let ?pathint = "\x y. contour_integral(linepath x y) f" { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" define K where "K = 1 + max (dist a b) (max (dist b c) (dist c a))" define e where "e = norm y / K^2" have K1: "K \ 1" by (simp add: K_def max.coboundedI1) then have K: "K > 0" by linarith have [iff]: "dist a b \ K" "dist b c \ K" "dist c a \ K" by (simp_all add: K_def) have e: "e > 0" unfolding e_def using ynz K1 by simp define At where "At x y z n \ convex hull {x,y,z} \ convex hull {a,b,c} \ dist x y \ K/2^n \ dist y z \ K/2^n \ dist z x \ K/2^n \ norm(?pathint x y + ?pathint y z + ?pathint z x) \ e*(K/2^n)^2" for x y z n have At0: "At a b c 0" using fy by (simp add: At_def e_def has_chain_integral_chain_integral3) { fix x y z n assume At: "At x y z n" then have contf': "continuous_on (convex hull {x,y,z}) f" using contf At_def continuous_on_subset by metis have "\x' y' z'. At x' y' z' (Suc n) \ convex hull {x',y',z'} \ convex hull {x,y,z}" using At Cauchy_theorem_quadrisection [OF contf', of "K/2^n" e] apply (simp add: At_def algebra_simps) apply (meson convex_hull_subset empty_subsetI insert_subset subsetCE) done } note AtSuc = this obtain fa fb fc where f0 [simp]: "fa 0 = a" "fb 0 = b" "fc 0 = c" and cosb: "\n. convex hull {fa n, fb n, fc n} \ convex hull {a,b,c}" and dist: "\n. dist (fa n) (fb n) \ K/2^n" "\n. dist (fb n) (fc n) \ K/2^n" "\n. dist (fc n) (fa n) \ K/2^n" and no: "\n. norm(?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) \ e * (K/2^n)^2" and conv_le: "\n. convex hull {fa(Suc n), fb(Suc n), fc(Suc n)} \ convex hull {fa n, fb n, fc n}" apply (rule Chain3 [of At, OF At0 AtSuc]) apply (auto simp: At_def) done obtain x where x: "\n. x \ convex hull {fa n, fb n, fc n}" proof (rule bounded_closed_nest) show "\n. closed (convex hull {fa n, fb n, fc n})" by (simp add: compact_imp_closed finite_imp_compact_convex_hull) show "\m n. m \ n \ convex hull {fa n, fb n, fc n} \ convex hull {fa m, fb m, fc m}" by (erule transitive_stepwise_le) (auto simp: conv_le) qed (fastforce intro: finite_imp_bounded_convex_hull)+ then have xin: "x \ convex hull {a,b,c}" using assms f0 by blast then have fx: "f field_differentiable at x within (convex hull {a,b,c})" using assms holomorphic_on_def by blast { fix k n assume k: "0 < k" and le: "\x' y' z'. \dist x' y' \ k; dist y' z' \ k; dist z' x' \ k; x \ convex hull {x',y',z'}; convex hull {x',y',z'} \ convex hull {a,b,c}\ \ cmod (?pathint x' y' + ?pathint y' z' + ?pathint z' x') * 10 \ e * (dist x' y' + dist y' z' + dist z' x')\<^sup>2" and Kk: "K / k < 2 ^ n" have "K / 2 ^ n < k" using Kk k by (auto simp: field_simps) then have DD: "dist (fa n) (fb n) \ k" "dist (fb n) (fc n) \ k" "dist (fc n) (fa n) \ k" using dist [of n] k by linarith+ have dle: "(dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ (3 * K / 2 ^ n)\<^sup>2" using dist [of n] e K by (simp add: abs_le_square_iff [symmetric]) have less10: "\x y::real. 0 < x \ y \ 9*x \ y < x*10" by linarith have "e * (dist (fa n) (fb n) + dist (fb n) (fc n) + dist (fc n) (fa n))\<^sup>2 \ e * (3 * K / 2 ^ n)\<^sup>2" using ynz dle e mult_le_cancel_left_pos by blast also have "\ < cmod (?pathint (fa n) (fb n) + ?pathint (fb n) (fc n) + ?pathint (fc n) (fa n)) * 10" using no [of n] e K apply (simp add: e_def field_simps) apply (simp only: zero_less_norm_iff [symmetric]) done finally have False using le [OF DD x cosb] by auto } then have ?thesis using holomorphic_point_small_triangle [OF xin contf fx, of "e/10"] e apply clarsimp apply (rule_tac y1="K/k" in exE [OF real_arch_pow[of 2]], force+) done } moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" by simp (meson contf continuous_on_subset contour_integrable_continuous_linepath segments_subset_convex_hull(1) segments_subset_convex_hull(3) segments_subset_convex_hull(5)) ultimately show ?thesis using has_contour_integral_integral by fastforce qed subsection\<^marker>\tag unimportant\ \Version needing function holomorphic in interior only\ lemma Cauchy_theorem_flat_lemma: assumes f: "continuous_on (convex hull {a,b,c}) f" and c: "c - a = k *\<^sub>R (b - a)" and k: "0 \ k" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ show ?thesis proof (cases "k \ 1") case True show ?thesis by (simp add: contour_integral_split [OF fabc(1) k True c] contour_integral_reverse_linepath fabc) next case False then show ?thesis using fabc c apply (subst contour_integral_split [of a c f "1/k" b, symmetric]) apply (metis closed_segment_commute fabc(3)) apply (auto simp: k contour_integral_reverse_linepath) done qed qed lemma Cauchy_theorem_flat: assumes f: "continuous_on (convex hull {a,b,c}) f" and c: "c - a = k *\<^sub>R (b - a)" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "0 \ k") case True with assms show ?thesis by (blast intro: Cauchy_theorem_flat_lemma) next case False have "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using f continuous_on_subset segments_subset_convex_hull by metis+ moreover have "contour_integral (linepath b a) f + contour_integral (linepath a c) f + contour_integral (linepath c b) f = 0" apply (rule Cauchy_theorem_flat_lemma [of b a c f "1-k"]) using False c apply (auto simp: f insert_commute scaleR_conv_of_real algebra_simps) done ultimately show ?thesis apply (auto simp: contour_integral_reverse_linepath) using add_eq_0_iff by force qed lemma Cauchy_theorem_triangle_interior: assumes contf: "continuous_on (convex hull {a,b,c}) f" and holf: "f holomorphic_on interior (convex hull {a,b,c})" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof - have fabc: "continuous_on (closed_segment a b) f" "continuous_on (closed_segment b c) f" "continuous_on (closed_segment c a) f" using contf continuous_on_subset segments_subset_convex_hull by metis+ have "bounded (f ` (convex hull {a,b,c}))" by (simp add: compact_continuous_image compact_convex_hull compact_imp_bounded contf) then obtain B where "0 < B" and Bnf: "\x. x \ convex hull {a,b,c} \ norm (f x) \ B" by (auto simp: dest!: bounded_pos [THEN iffD1]) have "bounded (convex hull {a,b,c})" by (simp add: bounded_convex_hull) then obtain C where C: "0 < C" and Cno: "\y. y \ convex hull {a,b,c} \ norm y < C" using bounded_pos_less by blast then have diff_2C: "norm(x - y) \ 2*C" if x: "x \ convex hull {a, b, c}" and y: "y \ convex hull {a, b, c}" for x y proof - have "cmod x \ C" using x by (meson Cno not_le not_less_iff_gr_or_eq) hence "cmod (x - y) \ C + C" using y by (meson Cno add_mono_thms_linordered_field(4) less_eq_real_def norm_triangle_ineq4 order_trans) thus "cmod (x - y) \ 2 * C" by (metis mult_2) qed have contf': "continuous_on (convex hull {b,a,c}) f" using contf by (simp add: insert_commute) { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" have pi_eq_y: "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = y" by (rule has_chain_integral_chain_integral3 [OF fy]) have ?thesis proof (cases "c=a \ a=b \ b=c") case True then show ?thesis using Cauchy_theorem_flat [OF contf, of 0] using has_chain_integral_chain_integral3 [OF fy] ynz by (force simp: fabc contour_integral_reverse_linepath) next case False then have car3: "card {a, b, c} = Suc (DIM(complex))" by auto { assume "interior(convex hull {a,b,c}) = {}" then have "collinear{a,b,c}" using interior_convex_hull_eq_empty [OF car3] by (simp add: collinear_3_eq_affine_dependent) with False obtain d where "c \ a" "a \ b" "b \ c" "c - b = d *\<^sub>R (a - b)" by (auto simp: collinear_3 collinear_lemma) then have "False" using False Cauchy_theorem_flat [OF contf'] pi_eq_y ynz by (simp add: fabc add_eq_0_iff contour_integral_reverse_linepath) } then obtain d where d: "d \ interior (convex hull {a, b, c})" by blast { fix d1 assume d1_pos: "0 < d1" and d1: "\x x'. \x\convex hull {a, b, c}; x'\convex hull {a, b, c}; cmod (x' - x) < d1\ \ cmod (f x' - f x) < cmod y / (24 * C)" define e where "e = min 1 (min (d1/(4*C)) ((norm y / 24 / C) / B))" define shrink where "shrink x = x - e *\<^sub>R (x - d)" for x let ?pathint = "\x y. contour_integral(linepath x y) f" have e: "0 < e" "e \ 1" "e \ d1 / (4 * C)" "e \ cmod y / 24 / C / B" using d1_pos \C>0\ \B>0\ ynz by (simp_all add: e_def) then have eCB: "24 * e * C * B \ cmod y" using \C>0\ \B>0\ by (simp add: field_simps) have e_le_d1: "e * (4 * C) \ d1" using e \C>0\ by (simp add: field_simps) have "shrink a \ interior(convex hull {a,b,c})" "shrink b \ interior(convex hull {a,b,c})" "shrink c \ interior(convex hull {a,b,c})" using d e by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) then have fhp0: "(f has_contour_integral 0) (linepath (shrink a) (shrink b) +++ linepath (shrink b) (shrink c) +++ linepath (shrink c) (shrink a))" by (simp add: Cauchy_theorem_triangle holomorphic_on_subset [OF holf] hull_minimal) then have f_0_shrink: "?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a) = 0" by (simp add: has_chain_integral_chain_integral3) have fpi_abc: "f contour_integrable_on linepath (shrink a) (shrink b)" "f contour_integrable_on linepath (shrink b) (shrink c)" "f contour_integrable_on linepath (shrink c) (shrink a)" using fhp0 by (auto simp: valid_path_join dest: has_contour_integral_integrable) have cmod_shr: "\x y. cmod (shrink y - shrink x - (y - x)) = e * cmod (x - y)" using e by (simp add: shrink_def real_vector.scale_right_diff_distrib [symmetric]) have sh_eq: "\a b d::complex. (b - e *\<^sub>R (b - d)) - (a - e *\<^sub>R (a - d)) - (b - a) = e *\<^sub>R (a - b)" by (simp add: algebra_simps) have "cmod y / (24 * C) \ cmod y / cmod (b - a) / 12" using False \C>0\ diff_2C [of b a] ynz by (auto simp: field_split_simps hull_inc) have less_C: "\u \ convex hull {a, b, c}; 0 \ x; x \ 1\ \ x * cmod u < C" for x u apply (cases "x=0", simp add: \0) using Cno [of u] mult_left_le_one_le [of "cmod u" x] le_less_trans norm_ge_zero by blast { fix u v assume uv: "u \ convex hull {a, b, c}" "v \ convex hull {a, b, c}" "u\v" and fpi_uv: "f contour_integrable_on linepath (shrink u) (shrink v)" have shr_uv: "shrink u \ interior(convex hull {a,b,c})" "shrink v \ interior(convex hull {a,b,c})" using d e uv by (auto simp: hull_inc mem_interior_convex_shrink shrink_def) have cmod_fuv: "\x. 0\x \ x\1 \ cmod (f (linepath (shrink u) (shrink v) x)) \ B" using shr_uv by (blast intro: Bnf linepath_in_convex_hull interior_subset [THEN subsetD]) have By_uv: "B * (12 * (e * cmod (u - v))) \ cmod y" apply (rule order_trans [OF _ eCB]) using e \B>0\ diff_2C [of u v] uv by (auto simp: field_simps) { fix x::real assume x: "0\x" "x\1" have cmod_less_4C: "cmod ((1 - x) *\<^sub>R u - (1 - x) *\<^sub>R d) + cmod (x *\<^sub>R v - x *\<^sub>R d) < (C+C) + (C+C)" apply (rule add_strict_mono; rule norm_triangle_half_l [of _ 0]) using uv x d interior_subset apply (auto simp: hull_inc intro!: less_C) done have ll: "linepath (shrink u) (shrink v) x - linepath u v x = -e * ((1 - x) *\<^sub>R (u - d) + x *\<^sub>R (v - d))" by (simp add: linepath_def shrink_def algebra_simps scaleR_conv_of_real) have cmod_less_dt: "cmod (linepath (shrink u) (shrink v) x - linepath u v x) < d1" apply (simp only: ll norm_mult scaleR_diff_right) using \e>0\ cmod_less_4C apply (force intro: norm_triangle_lt less_le_trans [OF _ e_le_d1]) done have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) < cmod y / (24 * C)" using x uv shr_uv cmod_less_dt by (auto simp: hull_inc intro: d1 interior_subset [THEN subsetD] linepath_in_convex_hull) also have "\ \ cmod y / cmod (v - u) / 12" using False uv \C>0\ diff_2C [of v u] ynz by (auto simp: field_split_simps hull_inc) finally have "cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / cmod (v - u) / 12" by simp then have cmod_12_le: "cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) * 12 \ cmod y" using uv False by (auto simp: field_simps) have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ B * (cmod y / 24 / C / B * 2 * C) + 2 * C * (cmod y / 24 / C)" apply (rule add_mono [OF mult_mono]) using By_uv e \0 < B\ \0 < C\ x apply (simp_all add: cmod_fuv cmod_shr cmod_12_le) apply (simp add: field_simps) done also have "\ \ cmod y / 6" by simp finally have "cmod (f (linepath (shrink u) (shrink v) x)) * cmod (shrink v - shrink u - (v - u)) + cmod (v - u) * cmod (f (linepath (shrink u) (shrink v) x) - f (linepath u v x)) \ cmod y / 6" . } note cmod_diff_le = this have f_uv: "continuous_on (closed_segment u v) f" by (blast intro: uv continuous_on_subset [OF contf closed_segment_subset_convex_hull]) have **: "\f' x' f x::complex. f'*x' - f*x = f'*(x' - x) + x*(f' - f)" by (simp add: algebra_simps) have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ (B*(norm y /24/C/B)*2*C + (2*C)*(norm y/24/C)) * content (cbox 0 (1::real))" apply (rule has_integral_bound [of _ "\x. f(linepath (shrink u) (shrink v) x) * (shrink v - shrink u) - f(linepath u v x)*(v - u)" _ 0 1]) using ynz \0 < B\ \0 < C\ apply (simp_all del: le_divide_eq_numeral1) apply (simp add: has_integral_diff has_contour_integral_linepath [symmetric] has_contour_integral_integral fpi_uv f_uv contour_integrable_continuous_linepath) apply (auto simp: ** norm_triangle_le norm_mult cmod_diff_le simp del: le_divide_eq_numeral1) done also have "\ \ norm y / 6" by simp finally have "norm (?pathint (shrink u) (shrink v) - ?pathint u v) \ norm y / 6" . } note * = this have "norm (?pathint (shrink a) (shrink b) - ?pathint a b) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) moreover have "norm (?pathint (shrink b) (shrink c) - ?pathint b c) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) moreover have "norm (?pathint (shrink c) (shrink a) - ?pathint c a) \ norm y / 6" using False fpi_abc by (rule_tac *) (auto simp: hull_inc) ultimately have "norm((?pathint (shrink a) (shrink b) - ?pathint a b) + (?pathint (shrink b) (shrink c) - ?pathint b c) + (?pathint (shrink c) (shrink a) - ?pathint c a)) \ norm y / 6 + norm y / 6 + norm y / 6" by (metis norm_triangle_le add_mono) also have "\ = norm y / 2" by simp finally have "norm((?pathint (shrink a) (shrink b) + ?pathint (shrink b) (shrink c) + ?pathint (shrink c) (shrink a)) - (?pathint a b + ?pathint b c + ?pathint c a)) \ norm y / 2" by (simp add: algebra_simps) then have "norm(?pathint a b + ?pathint b c + ?pathint c a) \ norm y / 2" by (simp add: f_0_shrink) (metis (mono_tags) add.commute minus_add_distrib norm_minus_cancel uminus_add_conv_diff) then have "False" using pi_eq_y ynz by auto } moreover have "uniformly_continuous_on (convex hull {a,b,c}) f" by (simp add: contf compact_convex_hull compact_uniformly_continuous) ultimately have "False" unfolding uniformly_continuous_on_def by (force simp: ynz \0 < C\ dist_norm) then show ?thesis .. qed } moreover have "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" using fabc contour_integrable_continuous_linepath by auto ultimately show ?thesis using has_contour_integral_integral by fastforce qed subsection\<^marker>\tag unimportant\ \Version allowing finite number of exceptional points\ proposition\<^marker>\tag unimportant\ Cauchy_theorem_triangle_cofinite: assumes "continuous_on (convex hull {a,b,c}) f" and "finite S" and "(\x. x \ interior(convex hull {a,b,c}) - S \ f field_differentiable (at x))" shows "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" using assms proof (induction "card S" arbitrary: a b c S rule: less_induct) case (less S a b c) show ?case proof (cases "S={}") case True with less show ?thesis by (fastforce simp: holomorphic_on_def field_differentiable_at_within Cauchy_theorem_triangle_interior) next case False then obtain d S' where d: "S = insert d S'" "d \ S'" by (meson Set.set_insert all_not_in_conv) then show ?thesis proof (cases "d \ convex hull {a,b,c}") case False show "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" proof (rule less.hyps) show "\x. x \ interior (convex hull {a, b, c}) - S' \ f field_differentiable at x" using False d interior_subset by (auto intro!: less.prems) qed (use d less.prems in auto) next case True have *: "convex hull {a, b, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have abd: "(f has_contour_integral 0) (linepath a b +++ linepath b d +++ linepath d a)" proof (rule less.hyps) show "\x. x \ interior (convex hull {a, b, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have *: "convex hull {b, c, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have bcd: "(f has_contour_integral 0) (linepath b c +++ linepath c d +++ linepath d b)" proof (rule less.hyps) show "\x. x \ interior (convex hull {b, c, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have *: "convex hull {c, a, d} \ convex hull {a, b, c}" by (meson True hull_subset insert_subset convex_hull_subset) have cad: "(f has_contour_integral 0) (linepath c a +++ linepath a d +++ linepath d c)" proof (rule less.hyps) show "\x. x \ interior (convex hull {c, a, d}) - S' \ f field_differentiable at x" using d not_in_interior_convex_hull_3 by (clarsimp intro!: less.prems) (metis * insert_absorb insert_subset interior_mono) qed (use d continuous_on_subset [OF _ *] less.prems in auto) have "f contour_integrable_on linepath a b" using less.prems abd contour_integrable_joinD1 contour_integrable_on_def by blast moreover have "f contour_integrable_on linepath b c" using less.prems bcd contour_integrable_joinD1 contour_integrable_on_def by blast moreover have "f contour_integrable_on linepath c a" using less.prems cad contour_integrable_joinD1 contour_integrable_on_def by blast ultimately have fpi: "f contour_integrable_on (linepath a b +++ linepath b c +++ linepath c a)" by auto { fix y::complex assume fy: "(f has_contour_integral y) (linepath a b +++ linepath b c +++ linepath c a)" and ynz: "y \ 0" have cont_ad: "continuous_on (closed_segment a d) f" by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(3)) have cont_bd: "continuous_on (closed_segment b d) f" by (meson True closed_segment_subset_convex_hull continuous_on_subset hull_subset insert_subset less.prems(1)) have cont_cd: "continuous_on (closed_segment c d) f" by (meson "*" continuous_on_subset less.prems(1) segments_subset_convex_hull(2)) have "contour_integral (linepath a b) f = - (contour_integral (linepath b d) f + (contour_integral (linepath d a) f))" "contour_integral (linepath b c) f = - (contour_integral (linepath c d) f + (contour_integral (linepath d b) f))" "contour_integral (linepath c a) f = - (contour_integral (linepath a d) f + contour_integral (linepath d c) f)" using has_chain_integral_chain_integral3 [OF abd] has_chain_integral_chain_integral3 [OF bcd] has_chain_integral_chain_integral3 [OF cad] by (simp_all add: algebra_simps add_eq_0_iff) then have ?thesis using cont_ad cont_bd cont_cd fy has_chain_integral_chain_integral3 contour_integral_reverse_linepath by fastforce } then show ?thesis using fpi contour_integrable_on_def by blast qed qed qed subsection\<^marker>\tag unimportant\ \Cauchy's theorem for an open starlike set\ lemma starlike_convex_subset: assumes S: "a \ S" "closed_segment b c \ S" and subs: "\x. x \ S \ closed_segment a x \ S" shows "convex hull {a,b,c} \ S" using S apply (clarsimp simp add: convex_hull_insert [of "{b,c}" a] segment_convex_hull) apply (meson subs convexD convex_closed_segment ends_in_segment(1) ends_in_segment(2) subsetCE) done lemma triangle_contour_integrals_starlike_primitive: assumes contf: "continuous_on S f" and S: "a \ S" "open S" and x: "x \ S" and subs: "\y. y \ S \ closed_segment a y \ S" and zer: "\b c. closed_segment b c \ S \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x)" proof - let ?pathint = "\x y. contour_integral(linepath x y) f" { fix e y assume e: "0 < e" and bxe: "ball x e \ S" and close: "cmod (y - x) < e" have y: "y \ S" using bxe close by (force simp: dist_norm norm_minus_commute) have cont_ayf: "continuous_on (closed_segment a y) f" using contf continuous_on_subset subs y by blast have xys: "closed_segment x y \ S" apply (rule order_trans [OF _ bxe]) using close by (auto simp: dist_norm ball_def norm_minus_commute dest: segment_bound) have "?pathint a y - ?pathint a x = ?pathint x y" using zer [OF xys] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force } note [simp] = this { fix e::real assume e: "0 < e" have cont_atx: "continuous (at x) f" using x S contf continuous_on_eq_continuous_at by blast then obtain d1 where d1: "d1>0" and d1_less: "\y. cmod (y - x) < d1 \ cmod (f y - f x) < e/2" unfolding continuous_at Lim_at dist_norm using e by (drule_tac x="e/2" in spec) force obtain d2 where d2: "d2>0" "ball x d2 \ S" using \open S\ x by (auto simp: open_contains_ball) have dpos: "min d1 d2 > 0" using d1 d2 by simp { fix y assume yx: "y \ x" and close: "cmod (y - x) < min d1 d2" have y: "y \ S" using d2 close by (force simp: dist_norm norm_minus_commute) have "closed_segment x y \ S" using close d2 by (auto simp: dist_norm norm_minus_commute dest!: segment_bound(1)) then have fxy: "f contour_integrable_on linepath x y" by (metis contour_integrable_continuous_linepath continuous_on_subset [OF contf]) then obtain i where i: "(f has_contour_integral i) (linepath x y)" by (auto simp: contour_integrable_on_def) then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" proof (rule has_contour_integral_bound_linepath) show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" by (meson close d1_less le_less_trans less_imp_le min.strict_boundedE segment_bound1) qed (use e in simp) also have "\ < e * cmod (y - x)" by (simp add: e yx) finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using i yx by (simp add: contour_integral_unique divide_less_eq) } then have "\d>0. \y. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using dpos by blast } then have *: "(\y. (?pathint x y - f x * (y - x)) /\<^sub>R cmod (y - x)) \x\ 0" by (simp add: Lim_at dist_norm inverse_eq_divide) show ?thesis apply (simp add: has_field_derivative_def has_derivative_at2 bounded_linear_mult_right) apply (rule Lim_transform [OF * tendsto_eventually]) using \open S\ x apply (force simp: dist_norm open_contains_ball inverse_eq_divide [symmetric] eventually_at) done qed (** Existence of a primitive.*) lemma holomorphic_starlike_primitive: fixes f :: "complex \ complex" assumes contf: "continuous_on S f" and S: "starlike S" and os: "open S" and k: "finite k" and fcd: "\x. x \ S - k \ f field_differentiable at x" shows "\g. \x \ S. (g has_field_derivative f x) (at x)" proof - obtain a where a: "a\S" and a_cs: "\x. x\S \ closed_segment a x \ S" using S by (auto simp: starlike_def) { fix x b c assume "x \ S" "closed_segment b c \ S" then have abcs: "convex hull {a, b, c} \ S" by (simp add: a a_cs starlike_convex_subset) then have "continuous_on (convex hull {a, b, c}) f" by (simp add: continuous_on_subset [OF contf]) then have "(f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" using abcs interior_subset by (force intro: fcd Cauchy_theorem_triangle_cofinite [OF _ k]) } note 0 = this show ?thesis apply (intro exI ballI) apply (rule triangle_contour_integrals_starlike_primitive [OF contf a os], assumption) apply (metis a_cs) apply (metis has_chain_integral_chain_integral3 0) done qed lemma Cauchy_theorem_starlike: "\open S; starlike S; finite k; continuous_on S f; \x. x \ S - k \ f field_differentiable at x; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (metis holomorphic_starlike_primitive Cauchy_theorem_primitive at_within_open) lemma Cauchy_theorem_starlike_simple: "\open S; starlike S; f holomorphic_on S; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" apply (rule Cauchy_theorem_starlike [OF _ _ finite.emptyI]) apply (simp_all add: holomorphic_on_imp_continuous_on) apply (metis at_within_open holomorphic_on_def) done subsection\Cauchy's theorem for a convex set\ text\For a convex set we can avoid assuming openness and boundary analyticity\ lemma triangle_contour_integrals_convex_primitive: assumes contf: "continuous_on S f" and S: "a \ S" "convex S" and x: "x \ S" and zer: "\b c. \b \ S; c \ S\ \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" shows "((\x. contour_integral(linepath a x) f) has_field_derivative f x) (at x within S)" proof - let ?pathint = "\x y. contour_integral(linepath x y) f" { fix y assume y: "y \ S" have cont_ayf: "continuous_on (closed_segment a y) f" using S y by (meson contf continuous_on_subset convex_contains_segment) have xys: "closed_segment x y \ S" (*?*) using convex_contains_segment S x y by auto have "?pathint a y - ?pathint a x = ?pathint x y" using zer [OF x y] contour_integral_reverse_linepath [OF cont_ayf] add_eq_0_iff by force } note [simp] = this { fix e::real assume e: "0 < e" have cont_atx: "continuous (at x within S) f" using x S contf by (simp add: continuous_on_eq_continuous_within) then obtain d1 where d1: "d1>0" and d1_less: "\y. \y \ S; cmod (y - x) < d1\ \ cmod (f y - f x) < e/2" unfolding continuous_within Lim_within dist_norm using e by (drule_tac x="e/2" in spec) force { fix y assume yx: "y \ x" and close: "cmod (y - x) < d1" and y: "y \ S" have fxy: "f contour_integrable_on linepath x y" using convex_contains_segment S x y by (blast intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) then obtain i where i: "(f has_contour_integral i) (linepath x y)" by (auto simp: contour_integrable_on_def) then have "((\w. f w - f x) has_contour_integral (i - f x * (y - x))) (linepath x y)" by (rule has_contour_integral_diff [OF _ has_contour_integral_const_linepath]) then have "cmod (i - f x * (y - x)) \ e / 2 * cmod (y - x)" proof (rule has_contour_integral_bound_linepath) show "\u. u \ closed_segment x y \ cmod (f u - f x) \ e / 2" by (meson assms(3) close convex_contains_segment d1_less le_less_trans less_imp_le segment_bound1 subset_iff x y) qed (use e in simp) also have "\ < e * cmod (y - x)" by (simp add: e yx) finally have "cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using i yx by (simp add: contour_integral_unique divide_less_eq) } then have "\d>0. \y\S. y \ x \ cmod (y-x) < d \ cmod (?pathint x y - f x * (y-x)) / cmod (y-x) < e" using d1 by blast } then have *: "((\y. (contour_integral (linepath x y) f - f x * (y - x)) /\<^sub>R cmod (y - x)) \ 0) (at x within S)" by (simp add: Lim_within dist_norm inverse_eq_divide) show ?thesis apply (simp add: has_field_derivative_def has_derivative_within bounded_linear_mult_right) apply (rule Lim_transform [OF * tendsto_eventually]) using linordered_field_no_ub apply (force simp: inverse_eq_divide [symmetric] eventually_at) done qed lemma contour_integral_convex_primitive: assumes "convex S" "continuous_on S f" "\a b c. \a \ S; b \ S; c \ S\ \ (f has_contour_integral 0) (linepath a b +++ linepath b c +++ linepath c a)" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (cases "S={}") case False with assms that show ?thesis by (blast intro: triangle_contour_integrals_convex_primitive has_chain_integral_chain_integral3) qed auto lemma holomorphic_convex_primitive: fixes f :: "complex \ complex" assumes "convex S" "finite K" and contf: "continuous_on S f" and fd: "\x. x \ interior S - K \ f field_differentiable at x" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (rule contour_integral_convex_primitive [OF \convex S\ contf Cauchy_theorem_triangle_cofinite]) have *: "convex hull {a, b, c} \ S" if "a \ S" "b \ S" "c \ S" for a b c by (simp add: \convex S\ hull_minimal that) show "continuous_on (convex hull {a, b, c}) f" if "a \ S" "b \ S" "c \ S" for a b c by (meson "*" contf continuous_on_subset that) show "f field_differentiable at x" if "a \ S" "b \ S" "c \ S" "x \ interior (convex hull {a, b, c}) - K" for a b c x by (metis "*" DiffD1 DiffD2 DiffI fd interior_mono subsetCE that) qed (use assms in \force+\) lemma holomorphic_convex_primitive': fixes f :: "complex \ complex" assumes "convex S" and "open S" and "f holomorphic_on S" obtains g where "\x. x \ S \ (g has_field_derivative f x) (at x within S)" proof (rule holomorphic_convex_primitive) fix x assume "x \ interior S - {}" with assms show "f field_differentiable at x" by (auto intro!: holomorphic_on_imp_differentiable_at simp: interior_open) qed (use assms in \auto intro: holomorphic_on_imp_continuous_on\) corollary\<^marker>\tag unimportant\ Cauchy_theorem_convex: "\continuous_on S f; convex S; finite K; \x. x \ interior S - K \ f field_differentiable at x; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (metis holomorphic_convex_primitive Cauchy_theorem_primitive) corollary Cauchy_theorem_convex_simple: "\f holomorphic_on S; convex S; valid_path g; path_image g \ S; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" apply (rule Cauchy_theorem_convex [where K = "{}"]) apply (simp_all add: holomorphic_on_imp_continuous_on) using at_within_interior holomorphic_on_def interior_subset by fastforce text\In particular for a disc\ corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc: "\finite K; continuous_on (cball a e) f; \x. x \ ball a e - K \ f field_differentiable at x; valid_path g; path_image g \ cball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (auto intro: Cauchy_theorem_convex) corollary\<^marker>\tag unimportant\ Cauchy_theorem_disc_simple: "\f holomorphic_on (ball a e); valid_path g; path_image g \ ball a e; pathfinish g = pathstart g\ \ (f has_contour_integral 0) g" by (simp add: Cauchy_theorem_convex_simple) subsection\<^marker>\tag unimportant\ \Generalize integrability to local primitives\ lemma contour_integral_local_primitive_lemma: fixes f :: "complex\complex" shows "\g piecewise_differentiable_on {a..b}; \x. x \ s \ (f has_field_derivative f' x) (at x within s); \x. x \ {a..b} \ g x \ s\ \ (\x. f' (g x) * vector_derivative g (at x within {a..b})) integrable_on {a..b}" apply (cases "cbox a b = {}", force) apply (simp add: integrable_on_def) apply (rule exI) apply (rule contour_integral_primitive_lemma, assumption+) using atLeastAtMost_iff by blast lemma contour_integral_local_primitive_any: fixes f :: "complex \ complex" assumes gpd: "g piecewise_differentiable_on {a..b}" and dh: "\x. x \ s \ \d h. 0 < d \ (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" and gs: "\x. x \ {a..b} \ g x \ s" shows "(\x. f(g x) * vector_derivative g (at x)) integrable_on {a..b}" proof - { fix x assume x: "a \ x" "x \ b" obtain d h where d: "0 < d" and h: "(\y. norm(y - g x) < d \ (h has_field_derivative f y) (at y within s))" using x gs dh by (metis atLeastAtMost_iff) have "continuous_on {a..b} g" using gpd piecewise_differentiable_on_def by blast then obtain e where e: "e>0" and lessd: "\x'. x' \ {a..b} \ \x' - x\ < e \ cmod (g x' - g x) < d" using x d apply (auto simp: dist_norm continuous_on_iff) apply (drule_tac x=x in bspec) using x apply simp apply (drule_tac x=d in spec, auto) done have "\d>0. \u v. u \ x \ x \ v \ {u..v} \ ball x d \ (u \ v \ a \ u \ v \ b) \ (\x. f (g x) * vector_derivative g (at x)) integrable_on {u..v}" apply (rule_tac x=e in exI) using e apply (simp add: integrable_on_localized_vector_derivative [symmetric], clarify) apply (rule_tac f = h and s = "g ` {u..v}" in contour_integral_local_primitive_lemma) apply (meson atLeastatMost_subset_iff gpd piecewise_differentiable_on_subset) apply (force simp: ball_def dist_norm intro: lessd gs DERIV_subset [OF h], force) done } then show ?thesis by (force simp: intro!: integrable_on_little_subintervals [of a b, simplified]) qed lemma contour_integral_local_primitive: fixes f :: "complex \ complex" assumes g: "valid_path g" "path_image g \ s" and dh: "\x. x \ s \ \d h. 0 < d \ (\y. norm(y - x) < d \ (h has_field_derivative f y) (at y within s))" shows "f contour_integrable_on g" using g apply (simp add: valid_path_def path_image_def contour_integrable_on_def has_contour_integral_def has_integral_localized_vector_derivative integrable_on_def [symmetric]) using contour_integral_local_primitive_any [OF _ dh] by (meson image_subset_iff piecewise_C1_imp_differentiable) text\In particular if a function is holomorphic\ lemma contour_integrable_holomorphic: assumes contf: "continuous_on s f" and os: "open s" and k: "finite k" and g: "valid_path g" "path_image g \ s" and fcd: "\x. x \ s - k \ f field_differentiable at x" shows "f contour_integrable_on g" proof - { fix z assume z: "z \ s" obtain d where "d>0" and d: "ball z d \ s" using \open s\ z by (auto simp: open_contains_ball) then have contfb: "continuous_on (ball z d) f" using contf continuous_on_subset by blast obtain h where "\y\ball z d. (h has_field_derivative f y) (at y within ball z d)" by (metis holomorphic_convex_primitive [OF convex_ball k contfb fcd] d interior_subset Diff_iff subsetD) then have "\y\ball z d. (h has_field_derivative f y) (at y within s)" by (metis open_ball at_within_open d os subsetCE) then have "\h. (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" by (force simp: dist_norm norm_minus_commute) then have "\d h. 0 < d \ (\y. cmod (y - z) < d \ (h has_field_derivative f y) (at y within s))" using \0 < d\ by blast } then show ?thesis by (rule contour_integral_local_primitive [OF g]) qed lemma contour_integrable_holomorphic_simple: assumes fh: "f holomorphic_on S" and os: "open S" and g: "valid_path g" "path_image g \ S" shows "f contour_integrable_on g" apply (rule contour_integrable_holomorphic [OF _ os Finite_Set.finite.emptyI g]) apply (simp add: fh holomorphic_on_imp_continuous_on) using fh by (simp add: field_differentiable_def holomorphic_on_open os) lemma continuous_on_inversediff: fixes z:: "'a::real_normed_field" shows "z \ S \ continuous_on S (\w. 1 / (w - z))" by (rule continuous_intros | force)+ lemma contour_integrable_inversediff: "\valid_path g; z \ path_image g\ \ (\w. 1 / (w-z)) contour_integrable_on g" apply (rule contour_integrable_holomorphic_simple [of _ "UNIV-{z}"]) apply (auto simp: holomorphic_on_open open_delete intro!: derivative_eq_intros) done text\Key fact that path integral is the same for a "nearby" path. This is the main lemma for the homotopy form of Cauchy's theorem and is also useful if we want "without loss of generality" to assume some nice properties of a path (e.g. smoothness). It can also be used to define the integrals of analytic functions over arbitrary continuous paths. This is just done for winding numbers now. \ text\A technical definition to avoid duplication of similar proofs, for paths joined at the ends versus looping paths\ definition linked_paths :: "bool \ (real \ 'a) \ (real \ 'a::topological_space) \ bool" where "linked_paths atends g h == (if atends then pathstart h = pathstart g \ pathfinish h = pathfinish g else pathfinish g = pathstart g \ pathfinish h = pathstart h)" text\This formulation covers two cases: \<^term>\g\ and \<^term>\h\ share their start and end points; \<^term>\g\ and \<^term>\h\ both loop upon themselves.\ lemma contour_integral_nearby: assumes os: "open S" and p: "path p" "path_image p \ S" shows "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ linked_paths atends g h \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" proof - have "\z. \e. z \ path_image p \ 0 < e \ ball z e \ S" using open_contains_ball os p(2) by blast then obtain ee where ee: "\z. z \ path_image p \ 0 < ee z \ ball z (ee z) \ S" by metis define cover where "cover = (\z. ball z (ee z/3)) ` (path_image p)" have "compact (path_image p)" by (metis p(1) compact_path_image) moreover have "path_image p \ (\c\path_image p. ball c (ee c / 3))" using ee by auto ultimately have "\D \ cover. finite D \ path_image p \ \D" by (simp add: compact_eq_Heine_Borel cover_def) then obtain D where D: "D \ cover" "finite D" "path_image p \ \D" by blast then obtain k where k: "k \ {0..1}" "finite k" and D_eq: "D = ((\z. ball z (ee z / 3)) \ p) ` k" apply (simp add: cover_def path_image_def image_comp) apply (blast dest!: finite_subset_image [OF \finite D\]) done then have kne: "k \ {}" using D by auto have pi: "\i. i \ k \ p i \ path_image p" using k by (auto simp: path_image_def) then have eepi: "\i. i \ k \ 0 < ee((p i))" by (metis ee) define e where "e = Min((ee \ p) ` k)" have fin_eep: "finite ((ee \ p) ` k)" using k by blast have "0 < e" using ee k by (simp add: kne e_def Min_gr_iff [OF fin_eep] eepi) have "uniformly_continuous_on {0..1} p" using p by (simp add: path_def compact_uniformly_continuous) then obtain d::real where d: "d>0" and de: "\x x'. \x' - x\ < d \ x\{0..1} \ x'\{0..1} \ cmod (p x' - p x) < e/3" unfolding uniformly_continuous_on_def dist_norm real_norm_def by (metis divide_pos_pos \0 < e\ zero_less_numeral) then obtain N::nat where N: "N>0" "inverse N < d" using real_arch_inverse [of d] by auto show ?thesis proof (intro exI conjI allI; clarify?) show "e/3 > 0" using \0 < e\ by simp fix g h assume g: "valid_path g" and ghp: "\t\{0..1}. cmod (g t - p t) < e / 3 \ cmod (h t - p t) < e / 3" and h: "valid_path h" and joins: "linked_paths atends g h" { fix t::real assume t: "0 \ t" "t \ 1" then obtain u where u: "u \ k" and ptu: "p t \ ball(p u) (ee(p u) / 3)" using \path_image p \ \D\ D_eq by (force simp: path_image_def) then have ele: "e \ ee (p u)" using fin_eep by (simp add: e_def) have "cmod (g t - p t) < e / 3" "cmod (h t - p t) < e / 3" using ghp t by auto with ele have "cmod (g t - p t) < ee (p u) / 3" "cmod (h t - p t) < ee (p u) / 3" by linarith+ then have "g t \ ball(p u) (ee(p u))" "h t \ ball(p u) (ee(p u))" using norm_diff_triangle_ineq [of "g t" "p t" "p t" "p u"] norm_diff_triangle_ineq [of "h t" "p t" "p t" "p u"] ptu eepi u by (force simp: dist_norm ball_def norm_minus_commute)+ then have "g t \ S" "h t \ S" using ee u k by (auto simp: path_image_def ball_def) } then have ghs: "path_image g \ S" "path_image h \ S" by (auto simp: path_image_def) moreover { fix f assume fhols: "f holomorphic_on S" then have fpa: "f contour_integrable_on g" "f contour_integrable_on h" using g ghs h holomorphic_on_imp_continuous_on os contour_integrable_holomorphic_simple by blast+ have contf: "continuous_on S f" by (simp add: fhols holomorphic_on_imp_continuous_on) { fix z assume z: "z \ path_image p" have "f holomorphic_on ball z (ee z)" using fhols ee z holomorphic_on_subset by blast then have "\ff. (\w \ ball z (ee z). (ff has_field_derivative f w) (at w))" using holomorphic_convex_primitive [of "ball z (ee z)" "{}" f, simplified] by (metis open_ball at_within_open holomorphic_on_def holomorphic_on_imp_continuous_on mem_ball) } then obtain ff where ff: "\z w. \z \ path_image p; w \ ball z (ee z)\ \ (ff z has_field_derivative f w) (at w)" by metis { fix n assume n: "n \ N" then have "contour_integral(subpath 0 (n/N) h) f - contour_integral(subpath 0 (n/N) g) f = contour_integral(linepath (g(n/N)) (h(n/N))) f - contour_integral(linepath (g 0) (h 0)) f" proof (induct n) case 0 show ?case by simp next case (Suc n) obtain t where t: "t \ k" and "p (n/N) \ ball(p t) (ee(p t) / 3)" using \path_image p \ \D\ [THEN subsetD, where c="p (n/N)"] D_eq N Suc.prems by (force simp: path_image_def) then have ptu: "cmod (p t - p (n/N)) < ee (p t) / 3" by (simp add: dist_norm) have e3le: "e/3 \ ee (p t) / 3" using fin_eep t by (simp add: e_def) { fix x assume x: "n/N \ x" "x \ (1 + n)/N" then have nN01: "0 \ n/N" "(1 + n)/N \ 1" using Suc.prems by auto then have x01: "0 \ x" "x \ 1" using x by linarith+ have "cmod (p t - p x) < ee (p t) / 3 + e/3" proof (rule norm_diff_triangle_less [OF ptu de]) show "\real n / real N - x\ < d" using x N by (auto simp: field_simps) qed (use x01 Suc.prems in auto) then have ptx: "cmod (p t - p x) < 2*ee (p t)/3" using e3le eepi [OF t] by simp have "cmod (p t - g x) < 2*ee (p t)/3 + e/3 " apply (rule norm_diff_triangle_less [OF ptx]) using ghp x01 by (simp add: norm_minus_commute) also have "\ \ ee (p t)" using e3le eepi [OF t] by simp finally have gg: "cmod (p t - g x) < ee (p t)" . have "cmod (p t - h x) < 2*ee (p t)/3 + e/3 " apply (rule norm_diff_triangle_less [OF ptx]) using ghp x01 by (simp add: norm_minus_commute) also have "\ \ ee (p t)" using e3le eepi [OF t] by simp finally have "cmod (p t - g x) < ee (p t)" "cmod (p t - h x) < ee (p t)" using gg by auto } note ptgh_ee = this have "closed_segment (g (real n / real N)) (h (real n / real N)) = path_image (linepath (h (n/N)) (g (n/N)))" by (simp add: closed_segment_commute) also have pi_hgn: "\ \ ball (p t) (ee (p t))" using ptgh_ee [of "n/N"] Suc.prems by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) finally have gh_ns: "closed_segment (g (n/N)) (h (n/N)) \ S" using ee pi t by blast have pi_ghn': "path_image (linepath (g ((1 + n) / N)) (h ((1 + n) / N))) \ ball (p t) (ee (p t))" using ptgh_ee [of "(1+n)/N"] Suc.prems by (auto simp: field_simps dist_norm dest: segment_furthest_le [where y="p t"]) then have gh_n's: "closed_segment (g ((1 + n) / N)) (h ((1 + n) / N)) \ S" using \N>0\ Suc.prems ee pi t by (auto simp: Path_Connected.path_image_join field_simps) have pi_subset_ball: "path_image (subpath (n/N) ((1+n) / N) g +++ linepath (g ((1+n) / N)) (h ((1+n) / N)) +++ subpath ((1+n) / N) (n/N) h +++ linepath (h (n/N)) (g (n/N))) \ ball (p t) (ee (p t))" apply (intro subset_path_image_join pi_hgn pi_ghn') using \N>0\ Suc.prems apply (auto simp: path_image_subpath dist_norm field_simps closed_segment_eq_real_ivl ptgh_ee) done have pi0: "(f has_contour_integral 0) (subpath (n/ N) ((Suc n)/N) g +++ linepath(g ((Suc n) / N)) (h((Suc n) / N)) +++ subpath ((Suc n) / N) (n/N) h +++ linepath(h (n/N)) (g (n/N)))" apply (rule Cauchy_theorem_primitive [of "ball(p t) (ee(p t))" "ff (p t)" "f"]) apply (metis ff open_ball at_within_open pi t) using Suc.prems pi_subset_ball apply (simp_all add: valid_path_join valid_path_subpath g h) done have fpa1: "f contour_integrable_on subpath (real n / real N) (real (Suc n) / real N) g" using Suc.prems by (simp add: contour_integrable_subpath g fpa) have fpa2: "f contour_integrable_on linepath (g (real (Suc n) / real N)) (h (real (Suc n) / real N))" using gh_n's by (auto intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) have fpa3: "f contour_integrable_on linepath (h (real n / real N)) (g (real n / real N))" using gh_ns by (auto simp: closed_segment_commute intro!: contour_integrable_continuous_linepath continuous_on_subset [OF contf]) have eq0: "contour_integral (subpath (n/N) ((Suc n) / real N) g) f + contour_integral (linepath (g ((Suc n) / N)) (h ((Suc n) / N))) f + contour_integral (subpath ((Suc n) / N) (n/N) h) f + contour_integral (linepath (h (n/N)) (g (n/N))) f = 0" using contour_integral_unique [OF pi0] Suc.prems by (simp add: g h fpa valid_path_subpath contour_integrable_subpath fpa1 fpa2 fpa3 algebra_simps del: of_nat_Suc) have *: "\hn he hn' gn gd gn' hgn ghn gh0 ghn'. \hn - gn = ghn - gh0; gd + ghn' + he + hgn = (0::complex); hn - he = hn'; gn + gd = gn'; hgn = -ghn\ \ hn' - gn' = ghn' - gh0" by (auto simp: algebra_simps) have "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = contour_integral (subpath 0 (n/N) h) f + contour_integral (subpath (n/N) ((Suc n) / N) h) f" unfolding reversepath_subpath [symmetric, of "((Suc n) / N)"] using Suc.prems by (simp add: h fpa contour_integral_reversepath valid_path_subpath contour_integrable_subpath) also have "\ = contour_integral (subpath 0 ((Suc n) / N) h) f" using Suc.prems by (simp add: contour_integral_subpath_combine h fpa) finally have pi0_eq: "contour_integral (subpath 0 (n/N) h) f - contour_integral (subpath ((Suc n) / N) (n/N) h) f = contour_integral (subpath 0 ((Suc n) / N) h) f" . show ?case apply (rule * [OF Suc.hyps eq0 pi0_eq]) using Suc.prems apply (simp_all add: g h fpa contour_integral_subpath_combine contour_integral_reversepath [symmetric] contour_integrable_continuous_linepath continuous_on_subset [OF contf gh_ns]) done qed } note ind = this have "contour_integral h f = contour_integral g f" using ind [OF order_refl] N joins by (simp add: linked_paths_def pathstart_def pathfinish_def split: if_split_asm) } ultimately show "path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f)" by metis qed qed lemma assumes "open S" "path p" "path_image p \ S" shows contour_integral_nearby_ends: "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ pathstart h = pathstart g \ pathfinish h = pathfinish g \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" and contour_integral_nearby_loops: "\d. 0 < d \ (\g h. valid_path g \ valid_path h \ (\t \ {0..1}. norm(g t - p t) < d \ norm(h t - p t) < d) \ pathfinish g = pathstart g \ pathfinish h = pathstart h \ path_image g \ S \ path_image h \ S \ (\f. f holomorphic_on S \ contour_integral h f = contour_integral g f))" using contour_integral_nearby [OF assms, where atends=True] using contour_integral_nearby [OF assms, where atends=False] unfolding linked_paths_def by simp_all lemma C1_differentiable_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ p C1_differentiable_on S" by (metis continuous_on_polymonial_function C1_differentiable_on_def has_vector_derivative_polynomial_function) lemma valid_path_polynomial_function: fixes p :: "real \ 'a::euclidean_space" shows "polynomial_function p \ valid_path p" by (force simp: valid_path_def piecewise_C1_differentiable_on_def continuous_on_polymonial_function C1_differentiable_polynomial_function) lemma valid_path_subpath_trivial [simp]: fixes g :: "real \ 'a::euclidean_space" shows "z \ g x \ valid_path (subpath x x g)" by (simp add: subpath_def valid_path_polynomial_function) lemma contour_integral_bound_exists: assumes S: "open S" and g: "valid_path g" and pag: "path_image g \ S" shows "\L. 0 < L \ (\f B. f holomorphic_on S \ (\z \ S. norm(f z) \ B) \ norm(contour_integral g f) \ L*B)" proof - have "path g" using g by (simp add: valid_path_imp_path) then obtain d::real and p where d: "0 < d" and p: "polynomial_function p" "path_image p \ S" and pi: "\f. f holomorphic_on S \ contour_integral g f = contour_integral p f" using contour_integral_nearby_ends [OF S \path g\ pag] apply clarify apply (drule_tac x=g in spec) apply (simp only: assms) apply (force simp: valid_path_polynomial_function dest: path_approx_polynomial_function) done then obtain p' where p': "polynomial_function p'" "\x. (p has_vector_derivative (p' x)) (at x)" by (blast intro: has_vector_derivative_polynomial_function that) then have "bounded(p' ` {0..1})" using continuous_on_polymonial_function by (force simp: intro!: compact_imp_bounded compact_continuous_image) then obtain L where L: "L>0" and nop': "\x. \0 \ x; x \ 1\ \ norm (p' x) \ L" by (force simp: bounded_pos) { fix f B assume f: "f holomorphic_on S" and B: "\z. z\S \ cmod (f z) \ B" then have "f contour_integrable_on p \ valid_path p" using p S by (blast intro: valid_path_polynomial_function contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on) moreover have "cmod (vector_derivative p (at x)) * cmod (f (p x)) \ L * B" if "0 \ x" "x \ 1" for x proof (rule mult_mono) show "cmod (vector_derivative p (at x)) \ L" by (metis nop' p'(2) that vector_derivative_at) show "cmod (f (p x)) \ B" by (metis B atLeastAtMost_iff imageI p(2) path_defs(4) subset_eq that) qed (use \L>0\ in auto) ultimately have "cmod (contour_integral g f) \ L * B" apply (simp only: pi [OF f]) apply (simp only: contour_integral_integral) apply (rule order_trans [OF integral_norm_bound_integral]) apply (auto simp: mult.commute integral_norm_bound_integral contour_integrable_on [symmetric] norm_mult) done } then show ?thesis by (force simp: L contour_integral_integral) qed text\We can treat even non-rectifiable paths as having a "length" for bounds on analytic functions in open sets.\ subsection \Winding Numbers\ definition\<^marker>\tag important\ winding_number_prop :: "[real \ complex, complex, real, real \ complex, complex] \ bool" where "winding_number_prop \ z e p n \ valid_path p \ z \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t \ {0..1}. norm(\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" definition\<^marker>\tag important\ winding_number:: "[real \ complex, complex] \ complex" where "winding_number \ z \ SOME n. \e > 0. \p. winding_number_prop \ z e p n" lemma winding_number: assumes "path \" "z \ path_image \" "0 < e" shows "\p. winding_number_prop \ z e p (winding_number \ z)" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain d where d: "d>0" and pi_eq: "\h1 h2. valid_path h1 \ valid_path h2 \ (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d) \ pathstart h2 = pathstart h1 \ pathfinish h2 = pathfinish h1 \ path_image h1 \ UNIV - {z} \ path_image h2 \ UNIV - {z} \ (\f. f holomorphic_on UNIV - {z} \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) then obtain h where h: "polynomial_function h \ pathstart h = pathstart \ \ pathfinish h = pathfinish \ \ (\t \ {0..1}. norm(h t - \ t) < d/2)" using path_approx_polynomial_function [OF \path \\, of "d/2"] d by auto define nn where "nn = 1/(2* pi*\) * contour_integral h (\w. 1/(w - z))" have "\n. \e > 0. \p. winding_number_prop \ z e p n" proof (rule_tac x=nn in exI, clarify) fix e::real assume e: "e>0" obtain p where p: "polynomial_function p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (p t - \ t) < min e (d/2))" using path_approx_polynomial_function [OF \path \\, of "min e (d/2)"] d \0 by auto have "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto simp: intro!: holomorphic_intros) then show "\p. winding_number_prop \ z e p nn" apply (rule_tac x=p in exI) using pi_eq [of h p] h p d apply (auto simp: valid_path_polynomial_function norm_minus_commute nn_def winding_number_prop_def) done qed then show ?thesis unfolding winding_number_def by (rule someI2_ex) (blast intro: \0) qed lemma winding_number_unique: assumes \: "path \" "z \ path_image \" and pi: "\e. e>0 \ \p. winding_number_prop \ z e p n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_ends [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "winding_number_prop \ z e p n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by (auto simp: winding_number_prop_def) also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed (*NB not winding_number_prop here due to the loop in p*) lemma winding_number_unique_loop: assumes \: "path \" "z \ path_image \" and loop: "pathfinish \ = pathstart \" and pi: "\e. e>0 \ \p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" shows "winding_number \ z = n" proof - have "path_image \ \ UNIV - {z}" using assms by blast then obtain e where e: "e>0" and pi_eq: "\h1 h2 f. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < e \ cmod (h2 t - \ t) < e); pathfinish h1 = pathstart h1; pathfinish h2 = pathstart h2; f holomorphic_on UNIV - {z}\ \ contour_integral h2 f = contour_integral h1 f" using contour_integral_nearby_loops [of "UNIV - {z}" \] assms by (auto simp: open_delete) obtain p where p: "valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t \ {0..1}. norm (\ t - p t) < e) \ contour_integral p (\w. 1/(w - z)) = 2 * pi * \ * n" using pi [OF e] by blast obtain q where q: "winding_number_prop \ z e q (winding_number \ z)" using winding_number [OF \ e] by blast have "2 * complex_of_real pi * \ * n = contour_integral p (\w. 1 / (w - z))" using p by auto also have "\ = contour_integral q (\w. 1 / (w - z))" proof (rule pi_eq) show "(\w. 1 / (w - z)) holomorphic_on UNIV - {z}" by (auto intro!: holomorphic_intros) qed (use p q loop in \auto simp: winding_number_prop_def norm_minus_commute\) also have "\ = 2 * complex_of_real pi * \ * winding_number \ z" using q by (auto simp: winding_number_prop_def) finally have "2 * complex_of_real pi * \ * n = 2 * complex_of_real pi * \ * winding_number \ z" . then show ?thesis by simp qed proposition winding_number_valid_path: assumes "valid_path \" "z \ path_image \" shows "winding_number \ z = 1/(2*pi*\) * contour_integral \ (\w. 1/(w - z))" by (rule winding_number_unique) (use assms in \auto simp: valid_path_imp_path winding_number_prop_def\) proposition has_contour_integral_winding_number: assumes \: "valid_path \" "z \ path_image \" shows "((\w. 1/(w - z)) has_contour_integral (2*pi*\*winding_number \ z)) \" by (simp add: winding_number_valid_path has_contour_integral_integral contour_integrable_inversediff assms) lemma winding_number_trivial [simp]: "z \ a \ winding_number(linepath a a) z = 0" by (simp add: winding_number_valid_path) lemma winding_number_subpath_trivial [simp]: "z \ g x \ winding_number (subpath x x g) z = 0" by (simp add: path_image_subpath winding_number_valid_path) lemma winding_number_join: assumes \1: "path \1" "z \ path_image \1" and \2: "path \2" "z \ path_image \2" and "pathfinish \1 = pathstart \2" shows "winding_number(\1 +++ \2) z = winding_number \1 z + winding_number \2 z" proof (rule winding_number_unique) show "\p. winding_number_prop (\1 +++ \2) z e p (winding_number \1 z + winding_number \2 z)" if "e > 0" for e proof - obtain p1 where "winding_number_prop \1 z e p1 (winding_number \1 z)" using \0 < e\ \1 winding_number by blast moreover obtain p2 where "winding_number_prop \2 z e p2 (winding_number \2 z)" using \0 < e\ \2 winding_number by blast ultimately have "winding_number_prop (\1+++\2) z e (p1+++p2) (winding_number \1 z + winding_number \2 z)" using assms apply (simp add: winding_number_prop_def not_in_path_image_join contour_integrable_inversediff algebra_simps) apply (auto simp: joinpaths_def) done then show ?thesis by blast qed qed (use assms in \auto simp: not_in_path_image_join\) lemma winding_number_reversepath: assumes "path \" "z \ path_image \" shows "winding_number(reversepath \) z = - (winding_number \ z)" proof (rule winding_number_unique) show "\p. winding_number_prop (reversepath \) z e p (- winding_number \ z)" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then have "winding_number_prop (reversepath \) z e (reversepath p) (- winding_number \ z)" using assms apply (simp add: winding_number_prop_def contour_integral_reversepath contour_integrable_inversediff valid_path_imp_reverse) apply (auto simp: reversepath_def) done then show ?thesis by blast qed qed (use assms in auto) lemma winding_number_shiftpath: assumes \: "path \" "z \ path_image \" and "pathfinish \ = pathstart \" "a \ {0..1}" shows "winding_number(shiftpath a \) z = winding_number \ z" proof (rule winding_number_unique_loop) show "\p. valid_path p \ z \ path_image p \ pathfinish p = pathstart p \ (\t\{0..1}. cmod (shiftpath a \ t - p t) < e) \ contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" if "e > 0" for e proof - obtain p where "winding_number_prop \ z e p (winding_number \ z)" using \0 < e\ assms winding_number by blast then show ?thesis apply (rule_tac x="shiftpath a p" in exI) using assms that apply (auto simp: winding_number_prop_def path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath contour_integral_shiftpath) apply (simp add: shiftpath_def) done qed qed (use assms in \auto simp: path_shiftpath path_image_shiftpath pathfinish_shiftpath pathstart_shiftpath\) lemma winding_number_split_linepath: assumes "c \ closed_segment a b" "z \ closed_segment a b" shows "winding_number(linepath a b) z = winding_number(linepath a c) z + winding_number(linepath c b) z" proof - have "z \ closed_segment a c" "z \ closed_segment c b" using assms by (meson convex_contains_segment convex_segment ends_in_segment subsetCE)+ then show ?thesis using assms by (simp add: winding_number_valid_path contour_integral_split_linepath [symmetric] continuous_on_inversediff field_simps) qed lemma winding_number_cong: "(\t. \0 \ t; t \ 1\ \ p t = q t) \ winding_number p z = winding_number q z" by (simp add: winding_number_def winding_number_prop_def pathstart_def pathfinish_def) lemma winding_number_constI: assumes "c\z" "\t. \0\t; t\1\ \ g t = c" shows "winding_number g z = 0" proof - have "winding_number g z = winding_number (linepath c c) z" apply (rule winding_number_cong) using assms unfolding linepath_def by auto moreover have "winding_number (linepath c c) z =0" apply (rule winding_number_trivial) using assms by auto ultimately show ?thesis by auto qed lemma winding_number_offset: "winding_number p z = winding_number (\w. p w - z) 0" unfolding winding_number_def proof (intro ext arg_cong [where f = Eps] arg_cong [where f = All] imp_cong refl, safe) fix n e g assume "0 < e" and g: "winding_number_prop p z e g n" then show "\r. winding_number_prop (\w. p w - z) 0 e r n" by (rule_tac x="\t. g t - z" in exI) (force simp: winding_number_prop_def contour_integral_integral valid_path_def path_defs vector_derivative_def has_vector_derivative_diff_const piecewise_C1_differentiable_diff C1_differentiable_imp_piecewise) next fix n e g assume "0 < e" and g: "winding_number_prop (\w. p w - z) 0 e g n" then show "\r. winding_number_prop p z e r n" apply (rule_tac x="\t. g t + z" in exI) apply (simp add: winding_number_prop_def contour_integral_integral valid_path_def path_defs piecewise_C1_differentiable_add vector_derivative_def has_vector_derivative_add_const C1_differentiable_imp_piecewise) apply (force simp: algebra_simps) done qed subsubsection\<^marker>\tag unimportant\ \Some lemmas about negating a path\ lemma valid_path_negatepath: "valid_path \ \ valid_path (uminus \ \)" unfolding o_def using piecewise_C1_differentiable_neg valid_path_def by blast lemma has_contour_integral_negatepath: assumes \: "valid_path \" and cint: "((\z. f (- z)) has_contour_integral - i) \" shows "(f has_contour_integral i) (uminus \ \)" proof - obtain S where cont: "continuous_on {0..1} \" and "finite S" and diff: "\ C1_differentiable_on {0..1} - S" using \ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def) have "((\x. - (f (- \ x) * vector_derivative \ (at x within {0..1}))) has_integral i) {0..1}" using cint by (auto simp: has_contour_integral_def dest: has_integral_neg) then have "((\x. f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1})) has_integral i) {0..1}" proof (rule rev_iffD1 [OF _ has_integral_spike_eq]) show "negligible S" by (simp add: \finite S\ negligible_finite) show "f (- \ x) * vector_derivative (uminus \ \) (at x within {0..1}) = - (f (- \ x) * vector_derivative \ (at x within {0..1}))" if "x \ {0..1} - S" for x proof - have "vector_derivative (uminus \ \) (at x within cbox 0 1) = - vector_derivative \ (at x within cbox 0 1)" proof (rule vector_derivative_within_cbox) show "(uminus \ \ has_vector_derivative - vector_derivative \ (at x within cbox 0 1)) (at x within cbox 0 1)" using that unfolding o_def by (metis C1_differentiable_on_eq UNIV_I diff differentiable_subset has_vector_derivative_minus subsetI that vector_derivative_works) qed (use that in auto) then show ?thesis by simp qed qed then show ?thesis by (simp add: has_contour_integral_def) qed lemma winding_number_negatepath: assumes \: "valid_path \" and 0: "0 \ path_image \" shows "winding_number(uminus \ \) 0 = winding_number \ 0" proof - have "(/) 1 contour_integrable_on \" using "0" \ contour_integrable_inversediff by fastforce then have "((\z. 1/z) has_contour_integral contour_integral \ ((/) 1)) \" by (rule has_contour_integral_integral) then have "((\z. 1 / - z) has_contour_integral - contour_integral \ ((/) 1)) \" using has_contour_integral_neg by auto then show ?thesis using assms apply (simp add: winding_number_valid_path valid_path_negatepath image_def path_defs) apply (simp add: contour_integral_unique has_contour_integral_negatepath) done qed lemma contour_integrable_negatepath: assumes \: "valid_path \" and pi: "(\z. f (- z)) contour_integrable_on \" shows "f contour_integrable_on (uminus \ \)" by (metis \ add.inverse_inverse contour_integrable_on_def has_contour_integral_negatepath pi) (* A combined theorem deducing several things piecewise.*) lemma winding_number_join_pos_combined: "\valid_path \1; z \ path_image \1; 0 < Re(winding_number \1 z); valid_path \2; z \ path_image \2; 0 < Re(winding_number \2 z); pathfinish \1 = pathstart \2\ \ valid_path(\1 +++ \2) \ z \ path_image(\1 +++ \2) \ 0 < Re(winding_number(\1 +++ \2) z)" by (simp add: valid_path_join path_image_join winding_number_join valid_path_imp_path) subsubsection\<^marker>\tag unimportant\ \Useful sufficient conditions for the winding number to be positive\ lemma Re_winding_number: "\valid_path \; z \ path_image \\ \ Re(winding_number \ z) = Im(contour_integral \ (\w. 1/(w - z))) / (2*pi)" by (simp add: winding_number_valid_path field_simps Re_divide power2_eq_square) lemma winding_number_pos_le: assumes \: "valid_path \" "z \ path_image \" and ge: "\x. \0 < x; x < 1\ \ 0 \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 \ Re(winding_number \ z)" proof - have ge0: "0 \ Im (vector_derivative \ (at x) / (\ x - z))" if x: "0 < x" "x < 1" for x using ge by (simp add: Complex.Im_divide algebra_simps x) let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have hi: "(?vd has_integral ?int z) (cbox 0 1)" unfolding box_real apply (subst has_contour_integral [symmetric]) using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) have "0 \ Im (?int z)" proof (rule has_integral_component_nonneg [of \, simplified]) show "\x. x \ cbox 0 1 \ 0 \ Im (if 0 < x \ x < 1 then ?vd x else 0)" by (force simp: ge0) show "((\x. if 0 < x \ x < 1 then ?vd x else 0) has_integral ?int z) (cbox 0 1)" by (rule has_integral_spike_interior [OF hi]) simp qed then show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt_lemma: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) / (\ x - z))" shows "0 < Re(winding_number \ z)" proof - let ?vd = "\x. 1 / (\ x - z) * vector_derivative \ (at x)" let ?int = "\z. contour_integral \ (\w. 1 / (w - z))" have hi: "(?vd has_integral ?int z) (cbox 0 1)" unfolding box_real apply (subst has_contour_integral [symmetric]) using \ by (simp add: contour_integrable_inversediff has_contour_integral_integral) have "e \ Im (contour_integral \ (\w. 1 / (w - z)))" proof (rule has_integral_component_le [of \ "\x. \*e" "\*e" "{0..1}", simplified]) show "((\x. if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e) has_integral ?int z) {0..1}" by (rule has_integral_spike_interior [OF hi, simplified box_real]) (use e in simp) show "\x. 0 \ x \ x \ 1 \ e \ Im (if 0 < x \ x < 1 then ?vd x else \ * complex_of_real e)" by (simp add: ge) qed (use has_integral_const_real [of _ 0 1] in auto) with e show ?thesis by (simp add: Re_winding_number [OF \] field_simps) qed lemma winding_number_pos_lt: assumes \: "valid_path \" "z \ path_image \" and e: "0 < e" and ge: "\x. \0 < x; x < 1\ \ e \ Im (vector_derivative \ (at x) * cnj(\ x - z))" shows "0 < Re (winding_number \ z)" proof - have bm: "bounded ((\w. w - z) ` (path_image \))" using bounded_translation [of _ "-z"] \ by (simp add: bounded_valid_path_image) then obtain B where B: "B > 0" and Bno: "\x. x \ (\w. w - z) ` (path_image \) \ norm x \ B" using bounded_pos [THEN iffD1, OF bm] by blast { fix x::real assume x: "0 < x" "x < 1" then have B2: "cmod (\ x - z)^2 \ B^2" using Bno [of "\ x - z"] by (simp add: path_image_def power2_eq_square mult_mono') with x have "\ x \ z" using \ using path_image_def by fastforce then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) * cnj (\ x - z)) / (cmod (\ x - z))\<^sup>2" using B ge [OF x] B2 e apply (rule_tac y="e / (cmod (\ x - z))\<^sup>2" in order_trans) apply (auto simp: divide_left_mono divide_right_mono) done then have "e / B\<^sup>2 \ Im (vector_derivative \ (at x) / (\ x - z))" by (simp add: complex_div_cnj [of _ "\ x - z" for x] del: complex_cnj_diff times_complex.sel) } note * = this show ?thesis using e B by (simp add: * winding_number_pos_lt_lemma [OF \, of "e/B^2"]) qed subsection\The winding number is an integer\ text\Proof from the book Complex Analysis by Lars V. Ahlfors, Chapter 4, section 2.1, Also on page 134 of Serge Lang's book with the name title, etc.\ lemma exp_fg: fixes z::complex assumes g: "(g has_vector_derivative g') (at x within s)" and f: "(f has_vector_derivative (g' / (g x - z))) (at x within s)" and z: "g x \ z" shows "((\x. exp(-f x) * (g x - z)) has_vector_derivative 0) (at x within s)" proof - have *: "(exp \ (\x. (- f x)) has_vector_derivative - (g' / (g x - z)) * exp (- f x)) (at x within s)" using assms unfolding has_vector_derivative_def scaleR_conv_of_real by (auto intro!: derivative_eq_intros) show ?thesis apply (rule has_vector_derivative_eq_rhs) using z apply (auto intro!: derivative_eq_intros * [unfolded o_def] g) done qed lemma winding_number_exp_integral: fixes z::complex assumes \: "\ piecewise_C1_differentiable_on {a..b}" and ab: "a \ b" and z: "z \ \ ` {a..b}" shows "(\x. vector_derivative \ (at x) / (\ x - z)) integrable_on {a..b}" (is "?thesis1") "exp (- (integral {a..b} (\x. vector_derivative \ (at x) / (\ x - z)))) * (\ b - z) = \ a - z" (is "?thesis2") proof - let ?D\ = "\x. vector_derivative \ (at x)" have [simp]: "\x. a \ x \ x \ b \ \ x \ z" using z by force have cong: "continuous_on {a..b} \" using \ by (simp add: piecewise_C1_differentiable_on_def) obtain k where fink: "finite k" and g_C1_diff: "\ C1_differentiable_on ({a..b} - k)" using \ by (force simp: piecewise_C1_differentiable_on_def) have \: "open ({a<..finite k\ by (simp add: finite_imp_closed open_Diff) moreover have "{a<.. {a..b} - k" by force ultimately have g_diff_at: "\x. \x \ k; x \ {a<.. \ \ differentiable at x" by (metis Diff_iff differentiable_on_subset C1_diff_imp_diff [OF g_C1_diff] differentiable_on_def at_within_open) { fix w assume "w \ z" have "continuous_on (ball w (cmod (w - z))) (\w. 1 / (w - z))" by (auto simp: dist_norm intro!: continuous_intros) moreover have "\x. cmod (w - x) < cmod (w - z) \ \f'. ((\w. 1 / (w - z)) has_field_derivative f') (at x)" by (auto simp: intro!: derivative_eq_intros) ultimately have "\h. \y. norm(y - w) < norm(w - z) \ (h has_field_derivative 1/(y - z)) (at y)" using holomorphic_convex_primitive [of "ball w (norm(w - z))" "{}" "\w. 1/(w - z)"] by (force simp: field_differentiable_def Ball_def dist_norm at_within_open_NO_MATCH norm_minus_commute) } then obtain h where h: "\w y. w \ z \ norm(y - w) < norm(w - z) \ (h w has_field_derivative 1/(y - z)) (at y)" by meson have exy: "\y. ((\x. inverse (\ x - z) * ?D\ x) has_integral y) {a..b}" unfolding integrable_on_def [symmetric] proof (rule contour_integral_local_primitive_any [OF piecewise_C1_imp_differentiable [OF \]]) show "\d h. 0 < d \ (\y. cmod (y - w) < d \ (h has_field_derivative inverse (y - z))(at y within - {z}))" if "w \ - {z}" for w apply (rule_tac x="norm(w - z)" in exI) using that inverse_eq_divide has_field_derivative_at_within h by (metis Compl_insert DiffD2 insertCI right_minus_eq zero_less_norm_iff) qed simp have vg_int: "(\x. ?D\ x / (\ x - z)) integrable_on {a..b}" unfolding box_real [symmetric] divide_inverse_commute by (auto intro!: exy integrable_subinterval simp add: integrable_on_def ab) with ab show ?thesis1 by (simp add: divide_inverse_commute integral_def integrable_on_def) { fix t assume t: "t \ {a..b}" have cball: "continuous_on (ball (\ t) (dist (\ t) z)) (\x. inverse (x - z))" using z by (auto intro!: continuous_intros simp: dist_norm) have icd: "\x. cmod (\ t - x) < cmod (\ t - z) \ (\w. inverse (w - z)) field_differentiable at x" unfolding field_differentiable_def by (force simp: intro!: derivative_eq_intros) obtain h where h: "\x. cmod (\ t - x) < cmod (\ t - z) \ (h has_field_derivative inverse (x - z)) (at x within {y. cmod (\ t - y) < cmod (\ t - z)})" using holomorphic_convex_primitive [where f = "\w. inverse(w - z)", OF convex_ball finite.emptyI cball icd] by simp (auto simp: ball_def dist_norm that) { fix x D assume x: "x \ k" "a < x" "x < b" then have "x \ interior ({a..b} - k)" using open_subset_interior [OF \] by fastforce then have con: "isCont ?D\ x" using g_C1_diff x by (auto simp: C1_differentiable_on_eq intro: continuous_on_interior) then have con_vd: "continuous (at x within {a..b}) (\x. ?D\ x)" by (rule continuous_at_imp_continuous_within) have gdx: "\ differentiable at x" using x by (simp add: g_diff_at) have "\d. \x \ k; a < x; x < b; (\ has_vector_derivative d) (at x); a \ t; t \ b\ \ ((\x. integral {a..x} (\x. ?D\ x / (\ x - z))) has_vector_derivative d / (\ x - z)) (at x within {a..b})" apply (rule has_vector_derivative_eq_rhs) apply (rule integral_has_vector_derivative_continuous_at [where S = "{}", simplified]) apply (rule con_vd continuous_intros cong vg_int | simp add: continuous_at_imp_continuous_within has_vector_derivative_continuous vector_derivative_at)+ done then have "((\c. exp (- integral {a..c} (\x. ?D\ x / (\ x - z))) * (\ c - z)) has_derivative (\h. 0)) (at x within {a..b})" using x gdx t apply (clarsimp simp add: differentiable_iff_scaleR) apply (rule exp_fg [unfolded has_vector_derivative_def, simplified], blast intro: has_derivative_at_withinI) apply (simp_all add: has_vector_derivative_def [symmetric]) done } note * = this have "exp (- (integral {a..t} (\x. ?D\ x / (\ x - z)))) * (\ t - z) =\ a - z" apply (rule has_derivative_zero_unique_strong_interval [of "{a,b} \ k" a b]) using t apply (auto intro!: * continuous_intros fink cong indefinite_integral_continuous_1 [OF vg_int] simp add: ab)+ done } with ab show ?thesis2 by (simp add: divide_inverse_commute integral_def) qed lemma winding_number_exp_2pi: "\path p; z \ path_image p\ \ pathfinish p - z = exp (2 * pi * \ * winding_number p z) * (pathstart p - z)" using winding_number [of p z 1] unfolding valid_path_def path_image_def pathstart_def pathfinish_def winding_number_prop_def by (force dest: winding_number_exp_integral(2) [of _ 0 1 z] simp: field_simps contour_integral_integral exp_minus) lemma integer_winding_number_eq: assumes \: "path \" and z: "z \ path_image \" shows "winding_number \ z \ \ \ pathfinish \ = pathstart \" proof - obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart \" "pathfinish p = pathfinish \" and eq: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF assms, of 1] unfolding winding_number_prop_def by auto then have wneq: "winding_number \ z = winding_number p z" using eq winding_number_valid_path by force have iff: "(winding_number \ z \ \) \ (exp (contour_integral p (\w. 1 / (w - z))) = 1)" using eq by (simp add: exp_eq_1 complex_is_Int_iff) have "exp (contour_integral p (\w. 1 / (w - z))) = (\ 1 - z) / (\ 0 - z)" using p winding_number_exp_integral(2) [of p 0 1 z] apply (simp add: valid_path_def path_defs contour_integral_integral exp_minus field_split_simps) by (metis path_image_def pathstart_def pathstart_in_path_image) then have "winding_number p z \ \ \ pathfinish p = pathstart p" using p wneq iff by (auto simp: path_defs) then show ?thesis using p eq by (auto simp: winding_number_valid_path) qed theorem integer_winding_number: "\path \; pathfinish \ = pathstart \; z \ path_image \\ \ winding_number \ z \ \" by (metis integer_winding_number_eq) text\If the winding number's magnitude is at least one, then the path must contain points in every direction.*) We can thus bound the winding number of a path that doesn't intersect a given ray. \ lemma winding_number_pos_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and 1: "Re (winding_number \ z) \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" proof - have [simp]: "\x. 0 \ x \ x \ 1 \ \ x \ z" using z by (auto simp: path_image_def) have [simp]: "z \ \ ` {0..1}" using path_image_def z by auto have gpd: "\ piecewise_C1_differentiable_on {0..1}" using \ valid_path_def by blast define r where "r = (w - z) / (\ 0 - z)" have [simp]: "r \ 0" using w z by (auto simp: r_def) have cont: "continuous_on {0..1} (\x. Im (integral {0..x} (\x. vector_derivative \ (at x) / (\ x - z))))" by (intro continuous_intros indefinite_integral_continuous_1 winding_number_exp_integral [OF gpd]; simp) have "Arg2pi r \ 2*pi" by (simp add: Arg2pi less_eq_real_def) also have "\ \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" using 1 apply (simp add: winding_number_valid_path [OF \ z] contour_integral_integral) apply (simp add: Complex.Re_divide field_simps power2_eq_square) done finally have "Arg2pi r \ Im (integral {0..1} (\x. vector_derivative \ (at x) / (\ x - z)))" . then have "\t. t \ {0..1} \ Im(integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by (simp add: Arg2pi_ge_0 cont IVT') then obtain t where t: "t \ {0..1}" and eqArg: "Im (integral {0..t} (\x. vector_derivative \ (at x)/(\ x - z))) = Arg2pi r" by blast define i where "i = integral {0..t} (\x. vector_derivative \ (at x) / (\ x - z))" have iArg: "Arg2pi r = Im i" using eqArg by (simp add: i_def) have gpdt: "\ piecewise_C1_differentiable_on {0..t}" by (metis atLeastAtMost_iff atLeastatMost_subset_iff order_refl piecewise_C1_differentiable_on_subset gpd t) have "exp (- i) * (\ t - z) = \ 0 - z" unfolding i_def apply (rule winding_number_exp_integral [OF gpdt]) using t z unfolding path_image_def by force+ then have *: "\ t - z = exp i * (\ 0 - z)" by (simp add: exp_minus field_simps) then have "(w - z) = r * (\ 0 - z)" by (simp add: r_def) then have "z + complex_of_real (exp (Re i)) * (w - z) / complex_of_real (cmod r) = \ t" apply simp apply (subst Complex_Transcendental.Arg2pi_eq [of r]) apply (simp add: iArg) using * apply (simp add: exp_eq_polar field_simps) done with t show ?thesis by (rule_tac x="exp(Re i) / norm r" in exI) (auto simp: path_image_def) qed lemma winding_number_big_meets: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and "\Re (winding_number \ z)\ \ 1" and w: "w \ z" shows "\a::real. 0 < a \ z + a*(w - z) \ path_image \" proof - { assume "Re (winding_number \ z) \ - 1" then have "Re (winding_number (reversepath \) z) \ 1" by (simp add: \ valid_path_imp_path winding_number_reversepath z) moreover have "valid_path (reversepath \)" using \ valid_path_imp_reverse by auto moreover have "z \ path_image (reversepath \)" by (simp add: z) ultimately have "\a::real. 0 < a \ z + a*(w - z) \ path_image (reversepath \)" using winding_number_pos_meets w by blast then have ?thesis by simp } then show ?thesis using assms by (simp add: abs_if winding_number_pos_meets split: if_split_asm) qed lemma winding_number_less_1: fixes z::complex shows "\valid_path \; z \ path_image \; w \ z; \a::real. 0 < a \ z + a*(w - z) \ path_image \\ \ Re(winding_number \ z) < 1" by (auto simp: not_less dest: winding_number_big_meets) text\One way of proving that WN=1 for a loop.\ lemma winding_number_eq_1: fixes z::complex assumes \: "valid_path \" and z: "z \ path_image \" and loop: "pathfinish \ = pathstart \" and 0: "0 < Re(winding_number \ z)" and 2: "Re(winding_number \ z) < 2" shows "winding_number \ z = 1" proof - have "winding_number \ z \ Ints" by (simp add: \ integer_winding_number loop valid_path_imp_path z) then show ?thesis using 0 2 by (auto simp: Ints_def) qed subsection\Continuity of winding number and invariance on connected sets\ lemma continuous_at_winding_number: fixes z::complex assumes \: "path \" and z: "z \ path_image \" shows "continuous (at z) (winding_number \)" proof - obtain e where "e>0" and cbg: "cball z e \ - path_image \" using open_contains_cball [of "- path_image \"] z by (force simp: closed_def [symmetric] closed_path_image [OF \]) then have ppag: "path_image \ \ - cball z (e/2)" by (force simp: cball_def dist_norm) have oc: "open (- cball z (e / 2))" by (simp add: closed_def [symmetric]) obtain d where "d>0" and pi_eq: "\h1 h2. \valid_path h1; valid_path h2; (\t\{0..1}. cmod (h1 t - \ t) < d \ cmod (h2 t - \ t) < d); pathstart h2 = pathstart h1; pathfinish h2 = pathfinish h1\ \ path_image h1 \ - cball z (e / 2) \ path_image h2 \ - cball z (e / 2) \ (\f. f holomorphic_on - cball z (e / 2) \ contour_integral h2 f = contour_integral h1 f)" using contour_integral_nearby_ends [OF oc \ ppag] by metis obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart \ \ pathfinish p = pathfinish \" and pg: "\t. t\{0..1} \ cmod (\ t - p t) < min d e / 2" and pi: "contour_integral p (\x. 1 / (x - z)) = complex_of_real (2 * pi) * \ * winding_number \ z" using winding_number [OF \ z, of "min d e / 2"] \d>0\ \e>0\ by (auto simp: winding_number_prop_def) { fix w assume d2: "cmod (w - z) < d/2" and e2: "cmod (w - z) < e/2" then have wnotp: "w \ path_image p" using cbg \d>0\ \e>0\ apply (simp add: path_image_def cball_def dist_norm, clarify) apply (frule pg) apply (drule_tac c="\ x" in subsetD) apply (auto simp: less_eq_real_def norm_minus_commute norm_triangle_half_l) done have wnotg: "w \ path_image \" using cbg e2 \e>0\ by (force simp: dist_norm norm_minus_commute) { fix k::real assume k: "k>0" then obtain q where q: "valid_path q" "w \ path_image q" "pathstart q = pathstart \ \ pathfinish q = pathfinish \" and qg: "\t. t \ {0..1} \ cmod (\ t - q t) < min k (min d e) / 2" and qi: "contour_integral q (\u. 1 / (u - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" using winding_number [OF \ wnotg, of "min k (min d e) / 2"] \d>0\ \e>0\ k by (force simp: min_divide_distrib_right winding_number_prop_def) have "contour_integral p (\u. 1 / (u - w)) = contour_integral q (\u. 1 / (u - w))" apply (rule pi_eq [OF \valid_path q\ \valid_path p\, THEN conjunct2, THEN conjunct2, rule_format]) apply (frule pg) apply (frule qg) using p q \d>0\ e2 apply (auto simp: dist_norm norm_minus_commute intro!: holomorphic_intros) done then have "contour_integral p (\x. 1 / (x - w)) = complex_of_real (2 * pi) * \ * winding_number \ w" by (simp add: pi qi) } note pip = this have "path p" using p by (simp add: valid_path_imp_path) then have "winding_number p w = winding_number \ w" apply (rule winding_number_unique [OF _ wnotp]) apply (rule_tac x=p in exI) apply (simp add: p wnotp min_divide_distrib_right pip winding_number_prop_def) done } note wnwn = this obtain pe where "pe>0" and cbp: "cball z (3 / 4 * pe) \ - path_image p" using p open_contains_cball [of "- path_image p"] by (force simp: closed_def [symmetric] closed_path_image [OF valid_path_imp_path]) obtain L where "L>0" and L: "\f B. \f holomorphic_on - cball z (3 / 4 * pe); \z \ - cball z (3 / 4 * pe). cmod (f z) \ B\ \ cmod (contour_integral p f) \ L * B" using contour_integral_bound_exists [of "- cball z (3/4*pe)" p] cbp \valid_path p\ by force { fix e::real and w::complex assume e: "0 < e" and w: "cmod (w - z) < pe/4" "cmod (w - z) < e * pe\<^sup>2 / (8 * L)" then have [simp]: "w \ path_image p" using cbp p(2) \0 < pe\ by (force simp: dist_norm norm_minus_commute path_image_def cball_def) have [simp]: "contour_integral p (\x. 1/(x - w)) - contour_integral p (\x. 1/(x - z)) = contour_integral p (\x. 1/(x - w) - 1/(x - z))" by (simp add: p contour_integrable_inversediff contour_integral_diff) { fix x assume pe: "3/4 * pe < cmod (z - x)" have "cmod (w - x) < pe/4 + cmod (z - x)" by (meson add_less_cancel_right norm_diff_triangle_le order_refl order_trans_rules(21) w(1)) then have wx: "cmod (w - x) < 4/3 * cmod (z - x)" using pe by simp have "cmod (z - x) \ cmod (z - w) + cmod (w - x)" using norm_diff_triangle_le by blast also have "\ < pe/4 + cmod (w - x)" using w by (simp add: norm_minus_commute) finally have "pe/2 < cmod (w - x)" using pe by auto then have "(pe/2)^2 < cmod (w - x) ^ 2" apply (rule power_strict_mono) using \pe>0\ by auto then have pe2: "pe^2 < 4 * cmod (w - x) ^ 2" by (simp add: power_divide) have "8 * L * cmod (w - z) < e * pe\<^sup>2" using w \L>0\ by (simp add: field_simps) also have "\ < e * 4 * cmod (w - x) * cmod (w - x)" using pe2 \e>0\ by (simp add: power2_eq_square) also have "\ < e * 4 * cmod (w - x) * (4/3 * cmod (z - x))" using wx apply (rule mult_strict_left_mono) using pe2 e not_less_iff_gr_or_eq by fastforce finally have "L * cmod (w - z) < 2/3 * e * cmod (w - x) * cmod (z - x)" by simp also have "\ \ e * cmod (w - x) * cmod (z - x)" using e by simp finally have Lwz: "L * cmod (w - z) < e * cmod (w - x) * cmod (z - x)" . have "L * cmod (1 / (x - w) - 1 / (x - z)) \ e" apply (cases "x=z \ x=w") using pe \pe>0\ w \L>0\ apply (force simp: norm_minus_commute) using wx w(2) \L>0\ pe pe2 Lwz apply (auto simp: divide_simps mult_less_0_iff norm_minus_commute norm_divide norm_mult power2_eq_square) done } note L_cmod_le = this have *: "cmod (contour_integral p (\x. 1 / (x - w) - 1 / (x - z))) \ L * (e * pe\<^sup>2 / L / 4 * (inverse (pe / 2))\<^sup>2)" apply (rule L) using \pe>0\ w apply (force simp: dist_norm norm_minus_commute intro!: holomorphic_intros) using \pe>0\ w \L>0\ apply (auto simp: cball_def dist_norm field_simps L_cmod_le simp del: less_divide_eq_numeral1 le_divide_eq_numeral1) done have "cmod (contour_integral p (\x. 1 / (x - w)) - contour_integral p (\x. 1 / (x - z))) < 2*e" apply simp apply (rule le_less_trans [OF *]) using \L>0\ e apply (force simp: field_simps) done then have "cmod (winding_number p w - winding_number p z) < e" using pi_ge_two e by (force simp: winding_number_valid_path p field_simps norm_divide norm_mult intro: less_le_trans) } note cmod_wn_diff = this then have "isCont (winding_number p) z" apply (simp add: continuous_at_eps_delta, clarify) apply (rule_tac x="min (pe/4) (e/2*pe^2/L/4)" in exI) using \pe>0\ \L>0\ apply (simp add: dist_norm cmod_wn_diff) done then show ?thesis apply (rule continuous_transform_within [where d = "min d e / 2"]) apply (auto simp: \d>0\ \e>0\ dist_norm wnwn) done qed corollary continuous_on_winding_number: "path \ \ continuous_on (- path_image \) (\w. winding_number \ w)" by (simp add: continuous_at_imp_continuous_on continuous_at_winding_number) subsection\<^marker>\tag unimportant\ \The winding number is constant on a connected region\ lemma winding_number_constant: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cs: "connected S" and sg: "S \ path_image \ = {}" shows "winding_number \ constant_on S" proof - have *: "1 \ cmod (winding_number \ y - winding_number \ z)" if ne: "winding_number \ y \ winding_number \ z" and "y \ S" "z \ S" for y z proof - have "winding_number \ y \ \" "winding_number \ z \ \" using that integer_winding_number [OF \ loop] sg \y \ S\ by auto with ne show ?thesis by (auto simp: Ints_def simp flip: of_int_diff) qed have cont: "continuous_on S (\w. winding_number \ w)" using continuous_on_winding_number [OF \] sg by (meson continuous_on_subset disjoint_eq_subset_Compl) show ?thesis using "*" zero_less_one by (blast intro: continuous_discrete_range_constant [OF cs cont]) qed lemma winding_number_eq: "\path \; pathfinish \ = pathstart \; w \ S; z \ S; connected S; S \ path_image \ = {}\ \ winding_number \ w = winding_number \ z" using winding_number_constant by (metis constant_on_def) lemma open_winding_number_levelsets: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "open {z. z \ path_image \ \ winding_number \ z = k}" proof - have opn: "open (- path_image \)" by (simp add: closed_path_image \ open_Compl) { fix z assume z: "z \ path_image \" and k: "k = winding_number \ z" obtain e where e: "e>0" "ball z e \ - path_image \" using open_contains_ball [of "- path_image \"] opn z by blast have "\e>0. \y. dist y z < e \ y \ path_image \ \ winding_number \ y = winding_number \ z" apply (rule_tac x=e in exI) using e apply (simp add: dist_norm ball_def norm_minus_commute) apply (auto simp: dist_norm norm_minus_commute intro!: winding_number_eq [OF assms, where S = "ball z e"]) done } then show ?thesis by (auto simp: open_dist) qed subsection\Winding number is zero "outside" a curve, in various senses\ proposition winding_number_zero_in_outside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and z: "z \ outside (path_image \)" shows "winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto obtain w::complex where w: "w \ ball 0 (B + 1)" by (metis abs_of_nonneg le_less less_irrefl mem_ball_0 norm_of_real) have "- ball 0 (B + 1) \ outside (path_image \)" apply (rule outside_subset_convex) using B subset_ball by auto then have wout: "w \ outside (path_image \)" using w by blast moreover have "winding_number \ constant_on outside (path_image \)" using winding_number_constant [OF \ loop, of "outside(path_image \)"] connected_outside by (metis DIM_complex bounded_path_image dual_order.refl \ outside_no_overlap) ultimately have "winding_number \ z = winding_number \ w" by (metis (no_types, hide_lams) constant_on_def z) also have "\ = 0" proof - have wnot: "w \ path_image \" using wout by (simp add: outside_def) { fix e::real assume "0" "pathfinish p = pathfinish \" and pg1: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < 1)" and pge: "(\t. \0 \ t; t \ 1\ \ cmod (p t - \ t) < e)" using path_approx_polynomial_function [OF \, of "min 1 e"] \e>0\ by force have pip: "path_image p \ ball 0 (B + 1)" using B apply (clarsimp simp add: path_image_def dist_norm ball_def) apply (frule (1) pg1) apply (fastforce dest: norm_add_less) done then have "w \ path_image p" using w by blast then have "\p. valid_path p \ w \ path_image p \ pathstart p = pathstart \ \ pathfinish p = pathfinish \ \ (\t\{0..1}. cmod (\ t - p t) < e) \ contour_integral p (\wa. 1 / (wa - w)) = 0" apply (rule_tac x=p in exI) apply (simp add: p valid_path_polynomial_function) apply (intro conjI) using pge apply (simp add: norm_minus_commute) apply (rule contour_integral_unique [OF Cauchy_theorem_convex_simple [OF _ convex_ball [of 0 "B+1"]]]) apply (rule holomorphic_intros | simp add: dist_norm)+ using mem_ball_0 w apply blast using p apply (simp_all add: valid_path_polynomial_function loop pip) done } then show ?thesis by (auto intro: winding_number_unique [OF \] simp add: winding_number_prop_def wnot) qed finally show ?thesis . qed corollary\<^marker>\tag unimportant\ winding_number_zero_const: "a \ z \ winding_number (\t. a) z = 0" by (rule winding_number_zero_in_outside) (auto simp: pathfinish_def pathstart_def path_polynomial_function) corollary\<^marker>\tag unimportant\ winding_number_zero_outside: "\path \; convex s; pathfinish \ = pathstart \; z \ s; path_image \ \ s\ \ winding_number \ z = 0" by (meson convex_in_outside outside_mono subsetCE winding_number_zero_in_outside) lemma winding_number_zero_at_infinity: assumes \: "path \" and loop: "pathfinish \ = pathstart \" shows "\B. \z. B \ norm z \ winding_number \ z = 0" proof - obtain B::real where "0 < B" and B: "path_image \ \ ball 0 B" using bounded_subset_ballD [OF bounded_path_image [OF \]] by auto then show ?thesis apply (rule_tac x="B+1" in exI, clarify) apply (rule winding_number_zero_outside [OF \ convex_cball [of 0 B] loop]) apply (meson less_add_one mem_cball_0 not_le order_trans) using ball_subset_cball by blast qed lemma winding_number_zero_point: "\path \; convex s; pathfinish \ = pathstart \; open s; path_image \ \ s\ \ \z. z \ s \ winding_number \ z = 0" using outside_compact_in_open [of "path_image \" s] path_image_nonempty winding_number_zero_in_outside by (fastforce simp add: compact_path_image) text\If a path winds round a set, it winds rounds its inside.\ lemma winding_number_around_inside: assumes \: "path \" and loop: "pathfinish \ = pathstart \" and cls: "closed s" and cos: "connected s" and s_disj: "s \ path_image \ = {}" and z: "z \ s" and wn_nz: "winding_number \ z \ 0" and w: "w \ s \ inside s" shows "winding_number \ w = winding_number \ z" proof - have ssb: "s \ inside(path_image \)" proof fix x :: complex assume "x \ s" hence "x \ path_image \" by (meson disjoint_iff_not_equal s_disj) thus "x \ inside (path_image \)" using \x \ s\ by (metis (no_types) ComplI UnE cos \ loop s_disj union_with_outside winding_number_eq winding_number_zero_in_outside wn_nz z) qed show ?thesis apply (rule winding_number_eq [OF \ loop w]) using z apply blast apply (simp add: cls connected_with_inside cos) apply (simp add: Int_Un_distrib2 s_disj, safe) by (meson ssb inside_inside_compact_connected [OF cls, of "path_image \"] compact_path_image connected_path_image contra_subsetD disjoint_iff_not_equal \ inside_no_overlap) qed text\Bounding a WN by 1/2 for a path and point in opposite halfspaces.\ lemma winding_number_subpath_continuous: assumes \: "valid_path \" and z: "z \ path_image \" shows "continuous_on {0..1} (\x. winding_number(subpath 0 x \) z)" proof - have *: "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = winding_number (subpath 0 x \) z" if x: "0 \ x" "x \ 1" for x proof - have "integral {0..x} (\t. vector_derivative \ (at t) / (\ t - z)) / (2 * of_real pi * \) = 1 / (2*pi*\) * contour_integral (subpath 0 x \) (\w. 1/(w - z))" using assms x apply (simp add: contour_integral_subcontour_integral [OF contour_integrable_inversediff]) done also have "\ = winding_number (subpath 0 x \) z" apply (subst winding_number_valid_path) using assms x apply (simp_all add: path_image_subpath valid_path_subpath) by (force simp: path_image_def) finally show ?thesis . qed show ?thesis apply (rule continuous_on_eq [where f = "\x. 1 / (2*pi*\) * integral {0..x} (\t. 1/(\ t - z) * vector_derivative \ (at t))"]) apply (rule continuous_intros)+ apply (rule indefinite_integral_continuous_1) apply (rule contour_integrable_inversediff [OF assms, unfolded contour_integrable_on]) using assms apply (simp add: *) done qed lemma winding_number_ivt_pos: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ Re(winding_number \ z)" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" apply (rule ivt_increasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) apply (rule winding_number_subpath_continuous [OF \ z]) using assms apply (auto simp: path_image_def image_def) done lemma winding_number_ivt_neg: assumes \: "valid_path \" and z: "z \ path_image \" and "Re(winding_number \ z) \ w" "w \ 0" shows "\t \ {0..1}. Re(winding_number(subpath 0 t \) z) = w" apply (rule ivt_decreasing_component_on_1 [of 0 1, where ?k = "1::complex", simplified complex_inner_1_right], simp) apply (rule winding_number_subpath_continuous [OF \ z]) using assms apply (auto simp: path_image_def image_def) done lemma winding_number_ivt_abs: assumes \: "valid_path \" and z: "z \ path_image \" and "0 \ w" "w \ \Re(winding_number \ z)\" shows "\t \ {0..1}. \Re (winding_number (subpath 0 t \) z)\ = w" using assms winding_number_ivt_pos [of \ z w] winding_number_ivt_neg [of \ z "-w"] by force lemma winding_number_lt_half_lemma: assumes \: "valid_path \" and z: "z \ path_image \" and az: "a \ z \ b" and pag: "path_image \ \ {w. a \ w > b}" shows "Re(winding_number \ z) < 1/2" proof - { assume "Re(winding_number \ z) \ 1/2" then obtain t::real where t: "0 \ t" "t \ 1" and sub12: "Re (winding_number (subpath 0 t \) z) = 1/2" using winding_number_ivt_pos [OF \ z, of "1/2"] by auto have gt: "\ t - z = - (of_real (exp (- (2 * pi * Im (winding_number (subpath 0 t \) z)))) * (\ 0 - z))" using winding_number_exp_2pi [of "subpath 0 t \" z] apply (simp add: t \ valid_path_imp_path) using closed_segment_eq_real_ivl path_image_def t z by (fastforce simp: path_image_subpath Euler sub12) have "b < a \ \ 0" proof - have "\ 0 \ {c. b < a \ c}" by (metis (no_types) pag atLeastAtMost_iff image_subset_iff order_refl path_image_def zero_le_one) thus ?thesis by blast qed moreover have "b < a \ \ t" proof - have "\ t \ {c. b < a \ c}" by (metis (no_types) pag atLeastAtMost_iff image_subset_iff path_image_def t) thus ?thesis by blast qed ultimately have "0 < a \ (\ 0 - z)" "0 < a \ (\ t - z)" using az by (simp add: inner_diff_right)+ then have False by (simp add: gt inner_mult_right mult_less_0_iff) } then show ?thesis by force qed lemma winding_number_lt_half: assumes "valid_path \" "a \ z \ b" "path_image \ \ {w. a \ w > b}" shows "\Re (winding_number \ z)\ < 1/2" proof - have "z \ path_image \" using assms by auto with assms show ?thesis apply (simp add: winding_number_lt_half_lemma abs_if del: less_divide_eq_numeral1) apply (metis complex_inner_1_right winding_number_lt_half_lemma [OF valid_path_imp_reverse, of \ z a b] winding_number_reversepath valid_path_imp_path inner_minus_left path_image_reversepath) done qed lemma winding_number_le_half: assumes \: "valid_path \" and z: "z \ path_image \" and anz: "a \ 0" and azb: "a \ z \ b" and pag: "path_image \ \ {w. a \ w \ b}" shows "\Re (winding_number \ z)\ \ 1/2" proof - { assume wnz_12: "\Re (winding_number \ z)\ > 1/2" have "isCont (winding_number \) z" by (metis continuous_at_winding_number valid_path_imp_path \ z) then obtain d where "d>0" and d: "\x'. dist x' z < d \ dist (winding_number \ x') (winding_number \ z) < \Re(winding_number \ z)\ - 1/2" using continuous_at_eps_delta wnz_12 diff_gt_0_iff_gt by blast define z' where "z' = z - (d / (2 * cmod a)) *\<^sub>R a" have *: "a \ z' \ b - d / 3 * cmod a" unfolding z'_def inner_mult_right' divide_inverse apply (simp add: field_split_simps algebra_simps dot_square_norm power2_eq_square anz) apply (metis \0 < d\ add_increasing azb less_eq_real_def mult_nonneg_nonneg mult_right_mono norm_ge_zero norm_numeral) done have "cmod (winding_number \ z' - winding_number \ z) < \Re (winding_number \ z)\ - 1/2" using d [of z'] anz \d>0\ by (simp add: dist_norm z'_def) then have "1/2 < \Re (winding_number \ z)\ - cmod (winding_number \ z' - winding_number \ z)" by simp then have "1/2 < \Re (winding_number \ z)\ - \Re (winding_number \ z') - Re (winding_number \ z)\" using abs_Re_le_cmod [of "winding_number \ z' - winding_number \ z"] by simp then have wnz_12': "\Re (winding_number \ z')\ > 1/2" by linarith moreover have "\Re (winding_number \ z')\ < 1/2" apply (rule winding_number_lt_half [OF \ *]) using azb \d>0\ pag apply (auto simp: add_strict_increasing anz field_split_simps algebra_simps dest!: subsetD) done ultimately have False by simp } then show ?thesis by force qed lemma winding_number_lt_half_linepath: "z \ closed_segment a b \ \Re (winding_number (linepath a b) z)\ < 1/2" using separating_hyperplane_closed_point [of "closed_segment a b" z] apply auto apply (simp add: closed_segment_def) apply (drule less_imp_le) apply (frule winding_number_lt_half [OF valid_path_linepath [of a b]]) apply (auto simp: segment) done text\ Positivity of WN for a linepath.\ lemma winding_number_linepath_pos_lt: assumes "0 < Im ((b - a) * cnj (b - z))" shows "0 < Re(winding_number(linepath a b) z)" proof - have z: "z \ path_image (linepath a b)" using assms by (simp add: closed_segment_def) (force simp: algebra_simps) show ?thesis apply (rule winding_number_pos_lt [OF valid_path_linepath z assms]) apply (simp add: linepath_def algebra_simps) done qed subsection\Cauchy's integral formula, again for a convex enclosing set\ lemma Cauchy_integral_formula_weak: assumes s: "convex s" and "finite k" and conf: "continuous_on s f" and fcd: "(\x. x \ interior s - k \ f field_differentiable at x)" and z: "z \ interior s - k" and vpg: "valid_path \" and pasz: "path_image \ \ s - {z}" and loop: "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - obtain f' where f': "(f has_field_derivative f') (at z)" using fcd [OF z] by (auto simp: field_differentiable_def) have pas: "path_image \ \ s" and znotin: "z \ path_image \" using pasz by blast+ have c: "continuous (at x within s) (\w. if w = z then f' else (f w - f z) / (w - z))" if "x \ s" for x proof (cases "x = z") case True then show ?thesis apply (simp add: continuous_within) apply (rule Lim_transform_away_within [of _ "z+1" _ "\w::complex. (f w - f z)/(w - z)"]) using has_field_derivative_at_within has_field_derivative_iff f' apply (fastforce simp add:)+ done next case False then have dxz: "dist x z > 0" by auto have cf: "continuous (at x within s) f" using conf continuous_on_eq_continuous_within that by blast have "continuous (at x within s) (\w. (f w - f z) / (w - z))" by (rule cf continuous_intros | simp add: False)+ then show ?thesis apply (rule continuous_transform_within [OF _ dxz that, of "\w::complex. (f w - f z)/(w - z)"]) apply (force simp: dist_commute) done qed have fink': "finite (insert z k)" using \finite k\ by blast have *: "((\w. if w = z then f' else (f w - f z) / (w - z)) has_contour_integral 0) \" apply (rule Cauchy_theorem_convex [OF _ s fink' _ vpg pas loop]) using c apply (force simp: continuous_on_eq_continuous_within) apply (rename_tac w) apply (rule_tac d="dist w z" and f = "\w. (f w - f z)/(w - z)" in field_differentiable_transform_within) apply (simp_all add: dist_pos_lt dist_commute) apply (metis less_irrefl) apply (rule derivative_intros fcd | simp)+ done show ?thesis apply (rule has_contour_integral_eq) using znotin has_contour_integral_add [OF has_contour_integral_lmul [OF has_contour_integral_winding_number [OF vpg znotin], of "f z"] *] apply (auto simp: ac_simps divide_simps) done qed theorem Cauchy_integral_formula_convex_simple: "\convex s; f holomorphic_on s; z \ interior s; valid_path \; path_image \ \ s - {z}; pathfinish \ = pathstart \\ \ ((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" apply (rule Cauchy_integral_formula_weak [where k = "{}"]) using holomorphic_on_imp_continuous_on by auto (metis at_within_interior holomorphic_on_def interiorE subsetCE) subsection\Homotopy forms of Cauchy's theorem\ lemma Cauchy_theorem_homotopic: assumes hom: "if atends then homotopic_paths s g h else homotopic_loops s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" proof - have pathsf: "linked_paths atends g h" using hom by (auto simp: linked_paths_def homotopic_paths_imp_pathstart homotopic_paths_imp_pathfinish homotopic_loops_imp_loop) obtain k :: "real \ real \ complex" where contk: "continuous_on ({0..1} \ {0..1}) k" and ks: "k ` ({0..1} \ {0..1}) \ s" and k [simp]: "\x. k (0, x) = g x" "\x. k (1, x) = h x" and ksf: "\t\{0..1}. linked_paths atends g (\x. k (t, x))" using hom pathsf by (auto simp: linked_paths_def homotopic_paths_def homotopic_loops_def homotopic_with_def split: if_split_asm) have ucontk: "uniformly_continuous_on ({0..1} \ {0..1}) k" by (blast intro: compact_Times compact_uniformly_continuous [OF contk]) { fix t::real assume t: "t \ {0..1}" have pak: "path (k \ (\u. (t, u)))" unfolding path_def apply (rule continuous_intros continuous_on_subset [OF contk])+ using t by force have pik: "path_image (k \ Pair t) \ s" using ks t by (auto simp: path_image_def) obtain e where "e>0" and e: "\g h. \valid_path g; valid_path h; \u\{0..1}. cmod (g u - (k \ Pair t) u) < e \ cmod (h u - (k \ Pair t) u) < e; linked_paths atends g h\ \ contour_integral h f = contour_integral g f" using contour_integral_nearby [OF \open s\ pak pik, of atends] f by metis obtain d where "d>0" and d: "\x x'. \x \ {0..1} \ {0..1}; x' \ {0..1} \ {0..1}; norm (x'-x) < d\ \ norm (k x' - k x) < e/4" by (rule uniformly_continuous_onE [OF ucontk, of "e/4"]) (auto simp: dist_norm \e>0\) { fix t1 t2 assume t1: "0 \ t1" "t1 \ 1" and t2: "0 \ t2" "t2 \ 1" and ltd: "\t1 - t\ < d" "\t2 - t\ < d" have no2: "\g1 k1 kt. \norm(g1 - k1) < e/4; norm(k1 - kt) < e/4\ \ norm(g1 - kt) < e" using \e > 0\ apply (rule_tac y = k1 in norm_triangle_half_l) apply (auto simp: norm_minus_commute intro: order_less_trans) done have "\d>0. \g1 g2. valid_path g1 \ valid_path g2 \ (\u\{0..1}. cmod (g1 u - k (t1, u)) < d \ cmod (g2 u - k (t2, u)) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f" apply (rule_tac x="e/4" in exI) using t t1 t2 ltd \e > 0\ apply (auto intro!: e simp: d no2 simp del: less_divide_eq_numeral1) done } then have "\e. 0 < e \ (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < e \ \t2 - t\ < e \ (\d. 0 < d \ (\g1 g2. valid_path g1 \ valid_path g2 \ (\u \ {0..1}. norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f)))" by (rule_tac x=d in exI) (simp add: \d > 0\) } then obtain ee where ee: "\t. t \ {0..1} \ ee t > 0 \ (\t1 t2. t1 \ {0..1} \ t2 \ {0..1} \ \t1 - t\ < ee t \ \t2 - t\ < ee t \ (\d. 0 < d \ (\g1 g2. valid_path g1 \ valid_path g2 \ (\u \ {0..1}. norm(g1 u - k((t1,u))) < d \ norm(g2 u - k((t2,u))) < d) \ linked_paths atends g1 g2 \ contour_integral g2 f = contour_integral g1 f)))" by metis note ee_rule = ee [THEN conjunct2, rule_format] define C where "C = (\t. ball t (ee t / 3)) ` {0..1}" obtain C' where C': "C' \ C" "finite C'" and C'01: "{0..1} \ \C'" proof (rule compactE [OF compact_interval]) show "{0..1} \ \C" using ee [THEN conjunct1] by (auto simp: C_def dist_norm) qed (use C_def in auto) define kk where "kk = {t \ {0..1}. ball t (ee t / 3) \ C'}" have kk01: "kk \ {0..1}" by (auto simp: kk_def) define e where "e = Min (ee ` kk)" have C'_eq: "C' = (\t. ball t (ee t / 3)) ` kk" using C' by (auto simp: kk_def C_def) have ee_pos[simp]: "\t. t \ {0..1} \ ee t > 0" by (simp add: kk_def ee) moreover have "finite kk" using \finite C'\ kk01 by (force simp: C'_eq inj_on_def ball_eq_ball_iff dest: ee_pos finite_imageD) moreover have "kk \ {}" using \{0..1} \ \C'\ C'_eq by force ultimately have "e > 0" using finite_less_Inf_iff [of "ee ` kk" 0] kk01 by (force simp: e_def) then obtain N::nat where "N > 0" and N: "1/N < e/3" by (meson divide_pos_pos nat_approx_posE zero_less_Suc zero_less_numeral) have e_le_ee: "\i. i \ kk \ e \ ee i" using \finite kk\ by (simp add: e_def Min_le_iff [of "ee ` kk"]) have plus: "\t \ kk. x \ ball t (ee t / 3)" if "x \ {0..1}" for x using C' subsetD [OF C'01 that] unfolding C'_eq by blast have [OF order_refl]: "\d. 0 < d \ (\j. valid_path j \ (\u \ {0..1}. norm(j u - k (n/N, u)) < d) \ linked_paths atends g j \ contour_integral j f = contour_integral g f)" if "n \ N" for n using that proof (induct n) case 0 show ?case using ee_rule [of 0 0 0] apply clarsimp apply (rule_tac x=d in exI, safe) by (metis diff_self vpg norm_zero) next case (Suc n) then have N01: "n/N \ {0..1}" "(Suc n)/N \ {0..1}" by auto then obtain t where t: "t \ kk" "n/N \ ball t (ee t / 3)" using plus [of "n/N"] by blast then have nN_less: "\n/N - t\ < ee t" by (simp add: dist_norm del: less_divide_eq_numeral1) have n'N_less: "\real (Suc n) / real N - t\ < ee t" using t N \N > 0\ e_le_ee [of t] by (simp add: dist_norm add_divide_distrib abs_diff_less_iff del: less_divide_eq_numeral1) (simp add: field_simps) have t01: "t \ {0..1}" using \kk \ {0..1}\ \t \ kk\ by blast obtain d1 where "d1 > 0" and d1: "\g1 g2. \valid_path g1; valid_path g2; \u\{0..1}. cmod (g1 u - k (n/N, u)) < d1 \ cmod (g2 u - k ((Suc n) / N, u)) < d1; linked_paths atends g1 g2\ \ contour_integral g2 f = contour_integral g1 f" using ee [THEN conjunct2, rule_format, OF t01 N01 nN_less n'N_less] by fastforce have "n \ N" using Suc.prems by auto with Suc.hyps obtain d2 where "d2 > 0" and d2: "\j. \valid_path j; \u\{0..1}. cmod (j u - k (n/N, u)) < d2; linked_paths atends g j\ \ contour_integral j f = contour_integral g f" by auto have "continuous_on {0..1} (k \ (\u. (n/N, u)))" apply (rule continuous_intros continuous_on_subset [OF contk])+ using N01 by auto then have pkn: "path (\u. k (n/N, u))" by (simp add: path_def) have min12: "min d1 d2 > 0" by (simp add: \0 < d1\ \0 < d2\) obtain p where "polynomial_function p" and psf: "pathstart p = pathstart (\u. k (n/N, u))" "pathfinish p = pathfinish (\u. k (n/N, u))" and pk_le: "\t. t\{0..1} \ cmod (p t - k (n/N, t)) < min d1 d2" using path_approx_polynomial_function [OF pkn min12] by blast then have vpp: "valid_path p" using valid_path_polynomial_function by blast have lpa: "linked_paths atends g p" by (metis (mono_tags, lifting) N01(1) ksf linked_paths_def pathfinish_def pathstart_def psf) show ?case proof (intro exI; safe) fix j assume "valid_path j" "linked_paths atends g j" and "\u\{0..1}. cmod (j u - k (real (Suc n) / real N, u)) < min d1 d2" then have "contour_integral j f = contour_integral p f" using pk_le N01(1) ksf by (force intro!: vpp d1 simp add: linked_paths_def psf) also have "... = contour_integral g f" using pk_le by (force intro!: vpp d2 lpa) finally show "contour_integral j f = contour_integral g f" . qed (simp add: \0 < d1\ \0 < d2\) qed then obtain d where "0 < d" "\j. valid_path j \ (\u \ {0..1}. norm(j u - k (1,u)) < d) \ linked_paths atends g j \ contour_integral j f = contour_integral g f" using \N>0\ by auto then have "linked_paths atends g h \ contour_integral h f = contour_integral g f" using \N>0\ vph by fastforce then show ?thesis by (simp add: pathsf) qed proposition Cauchy_theorem_homotopic_paths: assumes hom: "homotopic_paths s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" using Cauchy_theorem_homotopic [of True s g h] assms by simp proposition Cauchy_theorem_homotopic_loops: assumes hom: "homotopic_loops s g h" and "open s" and f: "f holomorphic_on s" and vpg: "valid_path g" and vph: "valid_path h" shows "contour_integral g f = contour_integral h f" using Cauchy_theorem_homotopic [of False s g h] assms by simp lemma has_contour_integral_newpath: "\(f has_contour_integral y) h; f contour_integrable_on g; contour_integral g f = contour_integral h f\ \ (f has_contour_integral y) g" using has_contour_integral_integral contour_integral_unique by auto lemma Cauchy_theorem_null_homotopic: "\f holomorphic_on s; open s; valid_path g; homotopic_loops s g (linepath a a)\ \ (f has_contour_integral 0) g" apply (rule has_contour_integral_newpath [where h = "linepath a a"], simp) using contour_integrable_holomorphic_simple apply (blast dest: holomorphic_on_imp_continuous_on homotopic_loops_imp_subset) by (simp add: Cauchy_theorem_homotopic_loops) subsection\<^marker>\tag unimportant\ \More winding number properties\ text\including the fact that it's +-1 inside a simple closed curve.\ lemma winding_number_homotopic_paths: assumes "homotopic_paths (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_paths_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_paths_imp_subset [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathstart p = pathstart g; pathfinish p = pathfinish g; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_paths (-{z}) g p" and e: "\q. \path q; pathstart q = pathstart h; pathfinish q = pathfinish h; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_paths (-{z}) h q" using homotopic_nearby_paths [of g "-{z}"] homotopic_nearby_paths [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have "homotopic_paths (- {z}) g p" by (simp add: d p valid_path_imp_path norm_minus_commute gp_less) moreover have "homotopic_paths (- {z}) h q" by (simp add: e q valid_path_imp_path norm_minus_commute hq_less) ultimately have "homotopic_paths (- {z}) p q" by (blast intro: homotopic_paths_trans homotopic_paths_sym assms) then have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" by (rule Cauchy_theorem_homotopic_paths) (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_homotopic_loops: assumes "homotopic_loops (-{z}) g h" shows "winding_number g z = winding_number h z" proof - have "path g" "path h" using homotopic_loops_imp_path [OF assms] by auto moreover have pag: "z \ path_image g" and pah: "z \ path_image h" using homotopic_loops_imp_subset [OF assms] by auto moreover have gloop: "pathfinish g = pathstart g" and hloop: "pathfinish h = pathstart h" using homotopic_loops_imp_loop [OF assms] by auto ultimately obtain d e where "d > 0" "e > 0" and d: "\p. \path p; pathfinish p = pathstart p; \t\{0..1}. norm (p t - g t) < d\ \ homotopic_loops (-{z}) g p" and e: "\q. \path q; pathfinish q = pathstart q; \t\{0..1}. norm (q t - h t) < e\ \ homotopic_loops (-{z}) h q" using homotopic_nearby_loops [of g "-{z}"] homotopic_nearby_loops [of h "-{z}"] by force obtain p where p: "valid_path p" "z \ path_image p" "pathstart p = pathstart g" "pathfinish p = pathfinish g" and gp_less:"\t\{0..1}. cmod (g t - p t) < d" and pap: "contour_integral p (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number g z" using winding_number [OF \path g\ pag \0 < d\] unfolding winding_number_prop_def by blast obtain q where q: "valid_path q" "z \ path_image q" "pathstart q = pathstart h" "pathfinish q = pathfinish h" and hq_less: "\t\{0..1}. cmod (h t - q t) < e" and paq: "contour_integral q (\w. 1 / (w - z)) = complex_of_real (2 * pi) * \ * winding_number h z" using winding_number [OF \path h\ pah \0 < e\] unfolding winding_number_prop_def by blast have gp: "homotopic_loops (- {z}) g p" by (simp add: gloop d gp_less norm_minus_commute p valid_path_imp_path) have hq: "homotopic_loops (- {z}) h q" by (simp add: e hloop hq_less norm_minus_commute q valid_path_imp_path) have "contour_integral p (\w. 1/(w - z)) = contour_integral q (\w. 1/(w - z))" proof (rule Cauchy_theorem_homotopic_loops) show "homotopic_loops (- {z}) p q" by (blast intro: homotopic_loops_trans homotopic_loops_sym gp hq assms) qed (auto intro!: holomorphic_intros simp: p q) then show ?thesis by (simp add: pap paq) qed lemma winding_number_paths_linear_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_paths_linear winding_number_homotopic_paths) lemma winding_number_loops_linear_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ z \ closed_segment (g t) (h t)\ \ winding_number h z = winding_number g z" by (blast intro: sym homotopic_loops_linear winding_number_homotopic_loops) lemma winding_number_nearby_paths_eq: "\path g; path h; pathstart h = pathstart g; pathfinish h = pathfinish g; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_paths_linear_eq) lemma winding_number_nearby_loops_eq: "\path g; path h; pathfinish g = pathstart g; pathfinish h = pathstart h; \t. t \ {0..1} \ norm(h t - g t) < norm(g t - z)\ \ winding_number h z = winding_number g z" by (metis segment_bound(2) norm_minus_commute not_le winding_number_loops_linear_eq) lemma winding_number_subpath_combine: "\path g; z \ path_image g; u \ {0..1}; v \ {0..1}; w \ {0..1}\ \ winding_number (subpath u v g) z + winding_number (subpath v w g) z = winding_number (subpath u w g) z" apply (rule trans [OF winding_number_join [THEN sym] winding_number_homotopic_paths [OF homotopic_join_subpaths]]) using path_image_subpath_subset by auto subsection\Partial circle path\ definition\<^marker>\tag important\ part_circlepath :: "[complex, real, real, real, real] \ complex" where "part_circlepath z r s t \ \x. z + of_real r * exp (\ * of_real (linepath s t x))" lemma pathstart_part_circlepath [simp]: "pathstart(part_circlepath z r s t) = z + r*exp(\ * s)" by (metis part_circlepath_def pathstart_def pathstart_linepath) lemma pathfinish_part_circlepath [simp]: "pathfinish(part_circlepath z r s t) = z + r*exp(\*t)" by (metis part_circlepath_def pathfinish_def pathfinish_linepath) lemma reversepath_part_circlepath[simp]: "reversepath (part_circlepath z r s t) = part_circlepath z r t s" unfolding part_circlepath_def reversepath_def linepath_def by (auto simp:algebra_simps) lemma has_vector_derivative_part_circlepath [derivative_intros]: "((part_circlepath z r s t) has_vector_derivative (\ * r * (of_real t - of_real s) * exp(\ * linepath s t x))) (at x within X)" apply (simp add: part_circlepath_def linepath_def scaleR_conv_of_real) apply (rule has_vector_derivative_real_field) apply (rule derivative_eq_intros | simp)+ done lemma differentiable_part_circlepath: "part_circlepath c r a b differentiable at x within A" using has_vector_derivative_part_circlepath[of c r a b x A] differentiableI_vector by blast lemma vector_derivative_part_circlepath: "vector_derivative (part_circlepath z r s t) (at x) = \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" using has_vector_derivative_part_circlepath vector_derivative_at by blast lemma vector_derivative_part_circlepath01: "\0 \ x; x \ 1\ \ vector_derivative (part_circlepath z r s t) (at x within {0..1}) = \ * r * (of_real t - of_real s) * exp(\ * linepath s t x)" using has_vector_derivative_part_circlepath by (auto simp: vector_derivative_at_within_ivl) lemma valid_path_part_circlepath [simp]: "valid_path (part_circlepath z r s t)" apply (simp add: valid_path_def) apply (rule C1_differentiable_imp_piecewise) apply (auto simp: C1_differentiable_on_eq vector_derivative_works vector_derivative_part_circlepath has_vector_derivative_part_circlepath intro!: continuous_intros) done lemma path_part_circlepath [simp]: "path (part_circlepath z r s t)" by (simp add: valid_path_imp_path) proposition path_image_part_circlepath: assumes "s \ t" shows "path_image (part_circlepath z r s t) = {z + r * exp(\ * of_real x) | x. s \ x \ x \ t}" proof - { fix z::real assume "0 \ z" "z \ 1" with \s \ t\ have "\x. (exp (\ * linepath s t z) = exp (\ * of_real x)) \ s \ x \ x \ t" apply (rule_tac x="(1 - z) * s + z * t" in exI) apply (simp add: linepath_def scaleR_conv_of_real algebra_simps) apply (rule conjI) using mult_right_mono apply blast using affine_ineq by (metis "mult.commute") } moreover { fix z assume "s \ z" "z \ t" then have "z + of_real r * exp (\ * of_real z) \ (\x. z + of_real r * exp (\ * linepath s t x)) ` {0..1}" apply (rule_tac x="(z - s)/(t - s)" in image_eqI) apply (simp add: linepath_def scaleR_conv_of_real divide_simps exp_eq) apply (auto simp: field_split_simps) done } ultimately show ?thesis by (fastforce simp add: path_image_def part_circlepath_def) qed lemma path_image_part_circlepath': "path_image (part_circlepath z r s t) = (\x. z + r * cis x) ` closed_segment s t" proof - have "path_image (part_circlepath z r s t) = (\x. z + r * exp(\ * of_real x)) ` linepath s t ` {0..1}" by (simp add: image_image path_image_def part_circlepath_def) also have "linepath s t ` {0..1} = closed_segment s t" by (rule linepath_image_01) finally show ?thesis by (simp add: cis_conv_exp) qed lemma path_image_part_circlepath_subset: "\s \ t; 0 \ r\ \ path_image(part_circlepath z r s t) \ sphere z r" by (auto simp: path_image_part_circlepath sphere_def dist_norm algebra_simps norm_mult) lemma in_path_image_part_circlepath: assumes "w \ path_image(part_circlepath z r s t)" "s \ t" "0 \ r" shows "norm(w - z) = r" proof - have "w \ {c. dist z c = r}" by (metis (no_types) path_image_part_circlepath_subset sphere_def subset_eq assms) thus ?thesis by (simp add: dist_norm norm_minus_commute) qed lemma path_image_part_circlepath_subset': assumes "r \ 0" shows "path_image (part_circlepath z r s t) \ sphere z r" proof (cases "s \ t") case True thus ?thesis using path_image_part_circlepath_subset[of s t r z] assms by simp next case False thus ?thesis using path_image_part_circlepath_subset[of t s r z] assms by (subst reversepath_part_circlepath [symmetric], subst path_image_reversepath) simp_all qed lemma part_circlepath_cnj: "cnj (part_circlepath c r a b x) = part_circlepath (cnj c) r (-a) (-b) x" by (simp add: part_circlepath_def exp_cnj linepath_def algebra_simps) lemma contour_integral_bound_part_circlepath: assumes "f contour_integrable_on part_circlepath c r a b" assumes "B \ 0" "r \ 0" "\x. x \ path_image (part_circlepath c r a b) \ norm (f x) \ B" shows "norm (contour_integral (part_circlepath c r a b) f) \ B * r * \b - a\" proof - let ?I = "integral {0..1} (\x. f (part_circlepath c r a b x) * \ * of_real (r * (b - a)) * exp (\ * linepath a b x))" have "norm ?I \ integral {0..1} (\x::real. B * 1 * (r * \b - a\) * 1)" proof (rule integral_norm_bound_integral, goal_cases) case 1 with assms(1) show ?case by (simp add: contour_integrable_on vector_derivative_part_circlepath mult_ac) next case (3 x) with assms(2-) show ?case unfolding norm_mult norm_of_real abs_mult by (intro mult_mono) (auto simp: path_image_def) qed auto also have "?I = contour_integral (part_circlepath c r a b) f" by (simp add: contour_integral_integral vector_derivative_part_circlepath mult_ac) finally show ?thesis by simp qed lemma has_contour_integral_part_circlepath_iff: assumes "a < b" shows "(f has_contour_integral I) (part_circlepath c r a b) \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" proof - have "(f has_contour_integral I) (part_circlepath c r a b) \ ((\x. f (part_circlepath c r a b x) * vector_derivative (part_circlepath c r a b) (at x within {0..1})) has_integral I) {0..1}" unfolding has_contour_integral_def .. also have "\ \ ((\x. f (part_circlepath c r a b x) * r * (b - a) * \ * cis (linepath a b x)) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_part_circlepath01) (simp_all add: cis_conv_exp) also have "\ \ ((\x. f (c + r * exp (\ * linepath (of_real a) (of_real b) x)) * r * \ * exp (\ * linepath (of_real a) (of_real b) x) * vector_derivative (linepath (of_real a) (of_real b)) (at x within {0..1})) has_integral I) {0..1}" by (intro has_integral_cong, subst vector_derivative_linepath_within) (auto simp: part_circlepath_def cis_conv_exp of_real_linepath [symmetric]) also have "\ \ ((\z. f (c + r * exp (\ * z)) * r * \ * exp (\ * z)) has_contour_integral I) (linepath (of_real a) (of_real b))" by (simp add: has_contour_integral_def) also have "\ \ ((\t. f (c + r * cis t) * r * \ * cis t) has_integral I) {a..b}" using assms by (subst has_contour_integral_linepath_Reals_iff) (simp_all add: cis_conv_exp) finally show ?thesis . qed lemma contour_integrable_part_circlepath_iff: assumes "a < b" shows "f contour_integrable_on (part_circlepath c r a b) \ (\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (auto simp: contour_integrable_on_def integrable_on_def has_contour_integral_part_circlepath_iff) lemma contour_integral_part_circlepath_eq: assumes "a < b" shows "contour_integral (part_circlepath c r a b) f = integral {a..b} (\t. f (c + r * cis t) * r * \ * cis t)" proof (cases "f contour_integrable_on part_circlepath c r a b") case True hence "(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with True show ?thesis using has_contour_integral_part_circlepath_iff[OF assms] contour_integral_unique has_integral_integrable_integral by blast next case False hence "\(\t. f (c + r * cis t) * r * \ * cis t) integrable_on {a..b}" using assms by (simp add: contour_integrable_part_circlepath_iff) with False show ?thesis by (simp add: not_integrable_contour_integral not_integrable_integral) qed lemma contour_integral_part_circlepath_reverse: "contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (subst reversepath_part_circlepath [symmetric], subst contour_integral_reversepath) simp_all lemma contour_integral_part_circlepath_reverse': "b < a \ contour_integral (part_circlepath c r a b) f = -contour_integral (part_circlepath c r b a) f" by (rule contour_integral_part_circlepath_reverse) lemma finite_bounded_log: "finite {z::complex. norm z \ b \ exp z = w}" proof (cases "w = 0") case True then show ?thesis by auto next case False have *: "finite {x. cmod (complex_of_real (2 * real_of_int x * pi) * \) \ b + cmod (Ln w)}" apply (simp add: norm_mult finite_int_iff_bounded_le) apply (rule_tac x="\(b + cmod (Ln w)) / (2*pi)\" in exI) apply (auto simp: field_split_simps le_floor_iff) done have [simp]: "\P f. {z. P z \ (\n. z = f n)} = f ` {n. P (f n)}" by blast show ?thesis apply (subst exp_Ln [OF False, symmetric]) apply (simp add: exp_eq) using norm_add_leD apply (fastforce intro: finite_subset [OF _ *]) done qed lemma finite_bounded_log2: fixes a::complex assumes "a \ 0" shows "finite {z. norm z \ b \ exp(a*z) = w}" proof - have *: "finite ((\z. z / a) ` {z. cmod z \ b * cmod a \ exp z = w})" by (rule finite_imageI [OF finite_bounded_log]) show ?thesis by (rule finite_subset [OF _ *]) (force simp: assms norm_mult) qed lemma has_contour_integral_bound_part_circlepath_strong: assumes fi: "(f has_contour_integral i) (part_circlepath z r s t)" and "finite k" and le: "0 \ B" "0 < r" "s \ t" and B: "\x. x \ path_image(part_circlepath z r s t) - k \ norm(f x) \ B" shows "cmod i \ B * r * (t - s)" proof - consider "s = t" | "s < t" using \s \ t\ by linarith then show ?thesis proof cases case 1 with fi [unfolded has_contour_integral] have "i = 0" by (simp add: vector_derivative_part_circlepath) with assms show ?thesis by simp next case 2 have [simp]: "\r\ = r" using \r > 0\ by linarith have [simp]: "cmod (complex_of_real t - complex_of_real s) = t-s" by (metis "2" abs_of_pos diff_gt_0_iff_gt norm_of_real of_real_diff) have "finite (part_circlepath z r s t -` {y} \ {0..1})" if "y \ k" for y proof - define w where "w = (y - z)/of_real r / exp(\ * of_real s)" have fin: "finite (of_real -` {z. cmod z \ 1 \ exp (\ * complex_of_real (t - s) * z) = w})" apply (rule finite_vimageI [OF finite_bounded_log2]) using \s < t\ apply (auto simp: inj_of_real) done show ?thesis apply (simp add: part_circlepath_def linepath_def vimage_def) apply (rule finite_subset [OF _ fin]) using le apply (auto simp: w_def algebra_simps scaleR_conv_of_real exp_add exp_diff) done qed then have fin01: "finite ((part_circlepath z r s t) -` k \ {0..1})" by (rule finite_finite_vimage_IntI [OF \finite k\]) have **: "((\x. if (part_circlepath z r s t x) \ k then 0 else f(part_circlepath z r s t x) * vector_derivative (part_circlepath z r s t) (at x)) has_integral i) {0..1}" by (rule has_integral_spike [OF negligible_finite [OF fin01]]) (use fi has_contour_integral in auto) have *: "\x. \0 \ x; x \ 1; part_circlepath z r s t x \ k\ \ cmod (f (part_circlepath z r s t x)) \ B" by (auto intro!: B [unfolded path_image_def image_def, simplified]) show ?thesis apply (rule has_integral_bound [where 'a=real, simplified, OF _ **, simplified]) using assms apply force apply (simp add: norm_mult vector_derivative_part_circlepath) using le * "2" \r > 0\ by auto qed qed lemma has_contour_integral_bound_part_circlepath: "\(f has_contour_integral i) (part_circlepath z r s t); 0 \ B; 0 < r; s \ t; \x. x \ path_image(part_circlepath z r s t) \ norm(f x) \ B\ \ norm i \ B*r*(t - s)" by (auto intro: has_contour_integral_bound_part_circlepath_strong) lemma contour_integrable_continuous_part_circlepath: "continuous_on (path_image (part_circlepath z r s t)) f \ f contour_integrable_on (part_circlepath z r s t)" apply (simp add: contour_integrable_on has_contour_integral_def vector_derivative_part_circlepath path_image_def) apply (rule integrable_continuous_real) apply (fast intro: path_part_circlepath [unfolded path_def] continuous_intros continuous_on_compose2 [where g=f, OF _ _ order_refl]) done proposition winding_number_part_circlepath_pos_less: assumes "s < t" and no: "norm(w - z) < r" shows "0 < Re (winding_number(part_circlepath z r s t) w)" proof - have "0 < r" by (meson no norm_not_less_zero not_le order.strict_trans2) note valid_path_part_circlepath moreover have " w \ path_image (part_circlepath z r s t)" using assms by (auto simp: path_image_def image_def part_circlepath_def norm_mult linepath_def) moreover have "0 < r * (t - s) * (r - cmod (w - z))" using assms by (metis \0 < r\ diff_gt_0_iff_gt mult_pos_pos) ultimately show ?thesis apply (rule winding_number_pos_lt [where e = "r*(t - s)*(r - norm(w - z))"]) apply (simp add: vector_derivative_part_circlepath right_diff_distrib [symmetric] mult_ac) apply (rule mult_left_mono)+ using Re_Im_le_cmod [of "w-z" "linepath s t x" for x] apply (simp add: exp_Euler cos_of_real sin_of_real part_circlepath_def algebra_simps cos_squared_eq [unfolded power2_eq_square]) using assms \0 < r\ by auto qed lemma simple_path_part_circlepath: "simple_path(part_circlepath z r s t) \ (r \ 0 \ s \ t \ \s - t\ \ 2*pi)" proof (cases "r = 0 \ s = t") case True then show ?thesis unfolding part_circlepath_def simple_path_def by (rule disjE) (force intro: bexI [where x = "1/4"] bexI [where x = "1/3"])+ next case False then have "r \ 0" "s \ t" by auto have *: "\x y z s t. \*((1 - x) * s + x * t) = \*(((1 - y) * s + y * t)) + z \ \*(x - y) * (t - s) = z" by (simp add: algebra_simps) have abs01: "\x y::real. 0 \ x \ x \ 1 \ 0 \ y \ y \ 1 \ (x = y \ x = 0 \ y = 1 \ x = 1 \ y = 0 \ \x - y\ \ {0,1})" by auto have **: "\x y. (\n. (complex_of_real x - of_real y) * (of_real t - of_real s) = 2 * (of_int n * of_real pi)) \ (\n. \x - y\ * (t - s) = 2 * (of_int n * pi))" by (force simp: algebra_simps abs_if dest: arg_cong [where f=Re] arg_cong [where f=complex_of_real] intro: exI [where x = "-n" for n]) have 1: "\s - t\ \ 2 * pi" if "\x. 0 \ x \ x \ 1 \ (\n. x * (t - s) = 2 * (real_of_int n * pi)) \ x = 0 \ x = 1" proof (rule ccontr) assume "\ \s - t\ \ 2 * pi" then have *: "\n. t - s \ of_int n * \s - t\" using False that [of "2*pi / \t - s\"] by (simp add: abs_minus_commute divide_simps) show False using * [of 1] * [of "-1"] by auto qed have 2: "\s - t\ = \2 * (real_of_int n * pi) / x\" if "x \ 0" "x * (t - s) = 2 * (real_of_int n * pi)" for x n proof - have "t-s = 2 * (real_of_int n * pi)/x" using that by (simp add: field_simps) then show ?thesis by (metis abs_minus_commute) qed have abs_away: "\P. (\x\{0..1}. \y\{0..1}. P \x - y\) \ (\x::real. 0 \ x \ x \ 1 \ P x)" by force show ?thesis using False apply (simp add: simple_path_def) apply (simp add: part_circlepath_def linepath_def exp_eq * ** abs01 del: Set.insert_iff) apply (subst abs_away) apply (auto simp: 1) apply (rule ccontr) apply (auto simp: 2 field_split_simps abs_mult dest: of_int_leD) done qed lemma arc_part_circlepath: assumes "r \ 0" "s \ t" "\s - t\ < 2*pi" shows "arc (part_circlepath z r s t)" proof - have *: "x = y" if eq: "\ * (linepath s t x) = \ * (linepath s t y) + 2 * of_int n * complex_of_real pi * \" and x: "x \ {0..1}" and y: "y \ {0..1}" for x y n proof (rule ccontr) assume "x \ y" have "(linepath s t x) = (linepath s t y) + 2 * of_int n * complex_of_real pi" by (metis add_divide_eq_iff complex_i_not_zero mult.commute nonzero_mult_div_cancel_left eq) then have "s*y + t*x = s*x + (t*y + of_int n * (pi * 2))" by (force simp: algebra_simps linepath_def dest: arg_cong [where f=Re]) with \x \ y\ have st: "s-t = (of_int n * (pi * 2) / (y-x))" by (force simp: field_simps) have "\real_of_int n\ < \y - x\" using assms \x \ y\ by (simp add: st abs_mult field_simps) then show False using assms x y st by (auto dest: of_int_lessD) qed show ?thesis using assms apply (simp add: arc_def) apply (simp add: part_circlepath_def inj_on_def exp_eq) apply (blast intro: *) done qed subsection\Special case of one complete circle\ definition\<^marker>\tag important\ circlepath :: "[complex, real, real] \ complex" where "circlepath z r \ part_circlepath z r 0 (2*pi)" lemma circlepath: "circlepath z r = (\x. z + r * exp(2 * of_real pi * \ * of_real x))" by (simp add: circlepath_def part_circlepath_def linepath_def algebra_simps) lemma pathstart_circlepath [simp]: "pathstart (circlepath z r) = z + r" by (simp add: circlepath_def) lemma pathfinish_circlepath [simp]: "pathfinish (circlepath z r) = z + r" by (simp add: circlepath_def) (metis exp_two_pi_i mult.commute) lemma circlepath_minus: "circlepath z (-r) x = circlepath z r (x + 1/2)" proof - have "z + of_real r * exp (2 * pi * \ * (x + 1/2)) = z + of_real r * exp (2 * pi * \ * x + pi * \)" by (simp add: divide_simps) (simp add: algebra_simps) also have "\ = z - r * exp (2 * pi * \ * x)" by (simp add: exp_add) finally show ?thesis by (simp add: circlepath path_image_def sphere_def dist_norm) qed lemma circlepath_add1: "circlepath z r (x+1) = circlepath z r x" using circlepath_minus [of z r "x+1/2"] circlepath_minus [of z "-r" x] by (simp add: add.commute) lemma circlepath_add_half: "circlepath z r (x + 1/2) = circlepath z r (x - 1/2)" using circlepath_add1 [of z r "x-1/2"] by (simp add: add.commute) lemma path_image_circlepath_minus_subset: "path_image (circlepath z (-r)) \ path_image (circlepath z r)" apply (simp add: path_image_def image_def circlepath_minus, clarify) apply (case_tac "xa \ 1/2", force) apply (force simp: circlepath_add_half)+ done lemma path_image_circlepath_minus: "path_image (circlepath z (-r)) = path_image (circlepath z r)" using path_image_circlepath_minus_subset by fastforce lemma has_vector_derivative_circlepath [derivative_intros]: "((circlepath z r) has_vector_derivative (2 * pi * \ * r * exp (2 * of_real pi * \ * of_real x))) (at x within X)" apply (simp add: circlepath_def scaleR_conv_of_real) apply (rule derivative_eq_intros) apply (simp add: algebra_simps) done lemma vector_derivative_circlepath: "vector_derivative (circlepath z r) (at x) = 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" using has_vector_derivative_circlepath vector_derivative_at by blast lemma vector_derivative_circlepath01: "\0 \ x; x \ 1\ \ vector_derivative (circlepath z r) (at x within {0..1}) = 2 * pi * \ * r * exp(2 * of_real pi * \ * x)" using has_vector_derivative_circlepath by (auto simp: vector_derivative_at_within_ivl) lemma valid_path_circlepath [simp]: "valid_path (circlepath z r)" by (simp add: circlepath_def) lemma path_circlepath [simp]: "path (circlepath z r)" by (simp add: valid_path_imp_path) lemma path_image_circlepath_nonneg: assumes "0 \ r" shows "path_image (circlepath z r) = sphere z r" proof - have *: "x \ (\u. z + (cmod (x - z)) * exp (\ * (of_real u * (of_real pi * 2)))) ` {0..1}" for x proof (cases "x = z") case True then show ?thesis by force next case False define w where "w = x - z" then have "w \ 0" by (simp add: False) have **: "\t. \Re w = cos t * cmod w; Im w = sin t * cmod w\ \ w = of_real (cmod w) * exp (\ * t)" using cis_conv_exp complex_eq_iff by auto show ?thesis apply (rule sincos_total_2pi [of "Re(w/of_real(norm w))" "Im(w/of_real(norm w))"]) apply (simp add: divide_simps \w \ 0\ cmod_power2 [symmetric]) apply (rule_tac x="t / (2*pi)" in image_eqI) apply (simp add: field_simps \w \ 0\) using False ** apply (auto simp: w_def) done qed show ?thesis unfolding circlepath path_image_def sphere_def dist_norm by (force simp: assms algebra_simps norm_mult norm_minus_commute intro: *) qed lemma path_image_circlepath [simp]: "path_image (circlepath z r) = sphere z \r\" using path_image_circlepath_minus by (force simp: path_image_circlepath_nonneg abs_if) lemma has_contour_integral_bound_circlepath_strong: "\(f has_contour_integral i) (circlepath z r); finite k; 0 \ B; 0 < r; \x. \norm(x - z) = r; x \ k\ \ norm(f x) \ B\ \ norm i \ B*(2*pi*r)" unfolding circlepath_def by (auto simp: algebra_simps in_path_image_part_circlepath dest!: has_contour_integral_bound_part_circlepath_strong) lemma has_contour_integral_bound_circlepath: "\(f has_contour_integral i) (circlepath z r); 0 \ B; 0 < r; \x. norm(x - z) = r \ norm(f x) \ B\ \ norm i \ B*(2*pi*r)" by (auto intro: has_contour_integral_bound_circlepath_strong) lemma contour_integrable_continuous_circlepath: "continuous_on (path_image (circlepath z r)) f \ f contour_integrable_on (circlepath z r)" by (simp add: circlepath_def contour_integrable_continuous_part_circlepath) lemma simple_path_circlepath: "simple_path(circlepath z r) \ (r \ 0)" by (simp add: circlepath_def simple_path_part_circlepath) lemma notin_path_image_circlepath [simp]: "cmod (w - z) < r \ w \ path_image (circlepath z r)" by (simp add: sphere_def dist_norm norm_minus_commute) lemma contour_integral_circlepath: assumes "r > 0" shows "contour_integral (circlepath z r) (\w. 1 / (w - z)) = 2 * complex_of_real pi * \" proof (rule contour_integral_unique) show "((\w. 1 / (w - z)) has_contour_integral 2 * complex_of_real pi * \) (circlepath z r)" unfolding has_contour_integral_def using assms apply (subst has_integral_cong) apply (simp add: vector_derivative_circlepath01) using has_integral_const_real [of _ 0 1] apply (force simp: circlepath) done qed lemma winding_number_circlepath_centre: "0 < r \ winding_number (circlepath z r) z = 1" apply (rule winding_number_unique_loop) apply (simp_all add: sphere_def valid_path_imp_path) apply (rule_tac x="circlepath z r" in exI) apply (simp add: sphere_def contour_integral_circlepath) done proposition winding_number_circlepath: assumes "norm(w - z) < r" shows "winding_number(circlepath z r) w = 1" proof (cases "w = z") case True then show ?thesis using assms winding_number_circlepath_centre by auto next case False have [simp]: "r > 0" using assms le_less_trans norm_ge_zero by blast define r' where "r' = norm(w - z)" have "r' < r" by (simp add: assms r'_def) have disjo: "cball z r' \ sphere z r = {}" using \r' < r\ by (force simp: cball_def sphere_def) have "winding_number(circlepath z r) w = winding_number(circlepath z r) z" proof (rule winding_number_around_inside [where s = "cball z r'"]) show "winding_number (circlepath z r) z \ 0" by (simp add: winding_number_circlepath_centre) show "cball z r' \ path_image (circlepath z r) = {}" by (simp add: disjo less_eq_real_def) qed (auto simp: r'_def dist_norm norm_minus_commute) also have "\ = 1" by (simp add: winding_number_circlepath_centre) finally show ?thesis . qed text\ Hence the Cauchy formula for points inside a circle.\ theorem Cauchy_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on (ball z r)" and wz: "norm(w - z) < r" shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) (circlepath z r)" proof - have "r > 0" using assms le_less_trans norm_ge_zero by blast have "((\u. f u / (u - w)) has_contour_integral (2 * pi) * \ * winding_number (circlepath z r) w * f w) (circlepath z r)" proof (rule Cauchy_integral_formula_weak [where s = "cball z r" and k = "{}"]) show "\x. x \ interior (cball z r) - {} \ f field_differentiable at x" using holf holomorphic_on_imp_differentiable_at by auto have "w \ sphere z r" by simp (metis dist_commute dist_norm not_le order_refl wz) then show "path_image (circlepath z r) \ cball z r - {w}" using \r > 0\ by (auto simp add: cball_def sphere_def) qed (use wz in \simp_all add: dist_norm norm_minus_commute contf\) then show ?thesis by (simp add: winding_number_circlepath assms) qed corollary\<^marker>\tag unimportant\ Cauchy_integral_circlepath_simple: assumes "f holomorphic_on cball z r" "norm(w - z) < r" shows "((\u. f u/(u - w)) has_contour_integral (2 * of_real pi * \ * f w)) (circlepath z r)" using assms by (force simp: holomorphic_on_imp_continuous_on holomorphic_on_subset Cauchy_integral_circlepath) lemma no_bounded_connected_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" and nb: "\z. bounded (connected_component_set (- s) z) \ z \ s" shows "winding_number g z = 0" apply (rule winding_number_zero_in_outside) apply (simp_all add: assms) by (metis nb [of z] \path_image g \ s\ \z \ s\ contra_subsetD mem_Collect_eq outside outside_mono) lemma no_bounded_path_component_imp_winding_number_zero: assumes g: "path g" "path_image g \ s" "pathfinish g = pathstart g" "z \ s" and nb: "\z. bounded (path_component_set (- s) z) \ z \ s" shows "winding_number g z = 0" apply (rule no_bounded_connected_component_imp_winding_number_zero [OF g]) by (simp add: bounded_subset nb path_component_subset_connected_component) subsection\ Uniform convergence of path integral\ text\Uniform convergence when the derivative of the path is bounded, and in particular for the special case of a circle.\ proposition contour_integral_uniform_limit: assumes ev_fint: "eventually (\n::'a. (f n) contour_integrable_on \) F" and ul_f: "uniform_limit (path_image \) f l F" and noleB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" and \: "valid_path \" and [simp]: "\ trivial_limit F" shows "l contour_integrable_on \" "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" proof - have "0 \ B" by (meson noleB [of 0] atLeastAtMost_iff norm_ge_zero order_refl order_trans zero_le_one) { fix e::real assume "0 < e" then have "0 < e / (\B\ + 1)" by simp then have "\\<^sub>F n in F. \x\path_image \. cmod (f n x - l x) < e / (\B\ + 1)" using ul_f [unfolded uniform_limit_iff dist_norm] by auto with ev_fint obtain a where fga: "\x. x \ {0..1} \ cmod (f a (\ x) - l (\ x)) < e / (\B\ + 1)" and inta: "(\t. f a (\ t) * vector_derivative \ (at t)) integrable_on {0..1}" using eventually_happens [OF eventually_conj] by (fastforce simp: contour_integrable_on path_image_def) have Ble: "B * e / (\B\ + 1) \ e" using \0 \ B\ \0 < e\ by (simp add: field_split_simps) have "\h. (\x\{0..1}. cmod (l (\ x) * vector_derivative \ (at x) - h x) \ e) \ h integrable_on {0..1}" proof (intro exI conjI ballI) show "cmod (l (\ x) * vector_derivative \ (at x) - f a (\ x) * vector_derivative \ (at x)) \ e" if "x \ {0..1}" for x apply (rule order_trans [OF _ Ble]) using noleB [OF that] fga [OF that] \0 \ B\ \0 < e\ apply (simp add: norm_mult left_diff_distrib [symmetric] norm_minus_commute divide_simps) apply (fastforce simp: mult_ac dest: mult_mono [OF less_imp_le]) done qed (rule inta) } then show lintg: "l contour_integrable_on \" unfolding contour_integrable_on by (metis (mono_tags, lifting)integrable_uniform_limit_real) { fix e::real define B' where "B' = B + 1" have B': "B' > 0" "B' > B" using \0 \ B\ by (auto simp: B'_def) assume "0 < e" then have ev_no': "\\<^sub>F n in F. \x\path_image \. 2 * cmod (f n x - l x) < e / B'" using ul_f [unfolded uniform_limit_iff dist_norm, rule_format, of "e / B' / 2"] B' by (simp add: field_simps) have ie: "integral {0..1::real} (\x. e / 2) < e" using \0 < e\ by simp have *: "cmod (f x (\ t) * vector_derivative \ (at t) - l (\ t) * vector_derivative \ (at t)) \ e / 2" if t: "t\{0..1}" and leB': "2 * cmod (f x (\ t) - l (\ t)) < e / B'" for x t proof - have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) \ e * (B/ B')" using mult_mono [OF less_imp_le [OF leB'] noleB] B' \0 < e\ t by auto also have "\ < e" by (simp add: B' \0 < e\ mult_imp_div_pos_less) finally have "2 * cmod (f x (\ t) - l (\ t)) * cmod (vector_derivative \ (at t)) < e" . then show ?thesis by (simp add: left_diff_distrib [symmetric] norm_mult) qed have le_e: "\x. \\xa\{0..1}. 2 * cmod (f x (\ xa) - l (\ xa)) < e / B'; f x contour_integrable_on \\ \ cmod (integral {0..1} (\u. f x (\ u) * vector_derivative \ (at u) - l (\ u) * vector_derivative \ (at u))) < e" apply (rule le_less_trans [OF integral_norm_bound_integral ie]) apply (simp add: lintg integrable_diff contour_integrable_on [symmetric]) apply (blast intro: *)+ done have "\\<^sub>F x in F. dist (contour_integral \ (f x)) (contour_integral \ l) < e" apply (rule eventually_mono [OF eventually_conj [OF ev_no' ev_fint]]) apply (simp add: dist_norm contour_integrable_on path_image_def contour_integral_integral) apply (simp add: lintg integral_diff [symmetric] contour_integrable_on [symmetric] le_e) done } then show "((\n. contour_integral \ (f n)) \ contour_integral \ l) F" by (rule tendstoI) qed corollary\<^marker>\tag unimportant\ contour_integral_uniform_limit_circlepath: assumes "\\<^sub>F n::'a in F. (f n) contour_integrable_on (circlepath z r)" and "uniform_limit (sphere z r) f l F" and "\ trivial_limit F" "0 < r" shows "l contour_integrable_on (circlepath z r)" "((\n. contour_integral (circlepath z r) (f n)) \ contour_integral (circlepath z r) l) F" using assms by (auto simp: vector_derivative_circlepath norm_mult intro!: contour_integral_uniform_limit) subsection\<^marker>\tag unimportant\ \General stepping result for derivative formulas\ lemma Cauchy_next_derivative: assumes "continuous_on (path_image \) f'" and leB: "\t. t \ {0..1} \ norm (vector_derivative \ (at t)) \ B" and int: "\w. w \ s - path_image \ \ ((\u. f' u / (u - w)^k) has_contour_integral f w) \" and k: "k \ 0" and "open s" and \: "valid_path \" and w: "w \ s - path_image \" shows "(\u. f' u / (u - w)^(Suc k)) contour_integrable_on \" and "(f has_field_derivative (k * contour_integral \ (\u. f' u/(u - w)^(Suc k)))) (at w)" (is "?thes2") proof - have "open (s - path_image \)" using \open s\ closed_valid_path_image \ by blast then obtain d where "d>0" and d: "ball w d \ s - path_image \" using w using open_contains_ball by blast have [simp]: "\n. cmod (1 + of_nat n) = 1 + of_nat n" by (metis norm_of_nat of_nat_Suc) have cint: "\x. \x \ w; cmod (x - w) < d\ \ (\z. (f' z / (z - x) ^ k - f' z / (z - w) ^ k) / (x * k - w * k)) contour_integrable_on \" apply (rule contour_integrable_div [OF contour_integrable_diff]) using int w d by (force simp: dist_norm norm_minus_commute intro!: has_contour_integral_integrable)+ have 1: "\\<^sub>F n in at w. (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) contour_integrable_on \" unfolding eventually_at apply (rule_tac x=d in exI) apply (simp add: \d > 0\ dist_norm field_simps cint) done have bim_g: "bounded (image f' (path_image \))" by (simp add: compact_imp_bounded compact_continuous_image compact_valid_path_image assms) then obtain C where "C > 0" and C: "\x. \0 \ x; x \ 1\ \ cmod (f' (\ x)) \ C" by (force simp: bounded_pos path_image_def) have twom: "\\<^sub>F n in at w. \x\path_image \. cmod ((inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k - inverse (x - w) ^ Suc k) < e" if "0 < e" for e proof - have *: "cmod ((inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k) - inverse (x - w) ^ Suc k) < e" if x: "x \ path_image \" and "u \ w" and uwd: "cmod (u - w) < d/2" and uw_less: "cmod (u - w) < e * (d/2) ^ (k+2) / (1 + real k)" for u x proof - define ff where [abs_def]: "ff n w = (if n = 0 then inverse(x - w)^k else if n = 1 then k / (x - w)^(Suc k) else (k * of_real(Suc k)) / (x - w)^(k + 2))" for n :: nat and w have km1: "\z::complex. z \ 0 \ z ^ (k - Suc 0) = z ^ k / z" by (simp add: field_simps) (metis Suc_pred \k \ 0\ neq0_conv power_Suc) have ff1: "(ff i has_field_derivative ff (Suc i) z) (at z within ball w (d/2))" if "z \ ball w (d/2)" "i \ 1" for i z proof - have "z \ path_image \" using \x \ path_image \\ d that ball_divide_subset_numeral by blast then have xz[simp]: "x \ z" using \x \ path_image \\ by blast then have neq: "x * x + z * z \ x * (z * 2)" by (blast intro: dest!: sum_sqs_eq) with xz have "\v. v \ 0 \ (x * x + z * z) * v \ (x * (z * 2) * v)" by auto then have neqq: "\v. v \ 0 \ x * (x * v) + z * (z * v) \ x * (z * (2 * v))" by (simp add: algebra_simps) show ?thesis using \i \ 1\ apply (simp add: ff_def dist_norm Nat.le_Suc_eq km1, safe) apply (rule derivative_eq_intros | simp add: km1 | simp add: field_simps neq neqq)+ done qed { fix a::real and b::real assume ab: "a > 0" "b > 0" then have "k * (1 + real k) * (1 / a) \ k * (1 + real k) * (4 / b) \ b \ 4 * a" by (subst mult_le_cancel_left_pos) (use \k \ 0\ in \auto simp: divide_simps\) with ab have "real k * (1 + real k) / a \ (real k * 4 + real k * real k * 4) / b \ b \ 4 * a" by (simp add: field_simps) } note canc = this have ff2: "cmod (ff (Suc 1) v) \ real (k * (k + 1)) / (d/2) ^ (k + 2)" if "v \ ball w (d/2)" for v proof - have lessd: "\z. cmod (\ z - v) < d/2 \ cmod (w - \ z) < d" by (metis that norm_minus_commute norm_triangle_half_r dist_norm mem_ball) have "d/2 \ cmod (x - v)" using d x that using lessd d x by (auto simp add: dist_norm path_image_def ball_def not_less [symmetric] del: divide_const_simps) then have "d \ cmod (x - v) * 2" by (simp add: field_split_simps) then have dpow_le: "d ^ (k+2) \ (cmod (x - v) * 2) ^ (k+2)" using \0 < d\ order_less_imp_le power_mono by blast have "x \ v" using that using \x \ path_image \\ ball_divide_subset_numeral d by fastforce then show ?thesis using \d > 0\ apply (simp add: ff_def norm_mult norm_divide norm_power dist_norm canc) using dpow_le apply (simp add: algebra_simps field_split_simps mult_less_0_iff) done qed have ub: "u \ ball w (d/2)" using uwd by (simp add: dist_commute dist_norm) have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) \ (real k * 4 + real k * real k * 4) * (cmod (u - w) * cmod (u - w)) / (d * (d * (d/2) ^ k))" using complex_Taylor [OF _ ff1 ff2 _ ub, of w, simplified] by (simp add: ff_def \0 < d\) then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) \ (cmod (u - w) * real k) * (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" by (simp add: field_simps) then have "cmod (inverse (x - u) ^ k - (inverse (x - w) ^ k + of_nat k * (u - w) / ((x - w) * (x - w) ^ k))) / (cmod (u - w) * real k) \ (1 + real k) * cmod (u - w) / (d/2) ^ (k+2)" using \k \ 0\ \u \ w\ by (simp add: mult_ac zero_less_mult_iff pos_divide_le_eq) also have "\ < e" using uw_less \0 < d\ by (simp add: mult_ac divide_simps) finally have e: "cmod (inverse (x-u)^k - (inverse (x-w)^k + of_nat k * (u-w) / ((x-w) * (x-w)^k))) / cmod ((u - w) * real k) < e" by (simp add: norm_mult) have "x \ u" using uwd \0 < d\ x d by (force simp: dist_norm ball_def norm_minus_commute) show ?thesis apply (rule le_less_trans [OF _ e]) using \k \ 0\ \x \ u\ \u \ w\ apply (simp add: field_simps norm_divide [symmetric]) done qed show ?thesis unfolding eventually_at apply (rule_tac x = "min (d/2) ((e*(d/2)^(k + 2))/(Suc k))" in exI) apply (force simp: \d > 0\ dist_norm that simp del: power_Suc intro: *) done qed have 2: "uniform_limit (path_image \) (\n x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k) (\x. f' x / (x - w) ^ Suc k) (at w)" unfolding uniform_limit_iff dist_norm proof clarify fix e::real assume "0 < e" have *: "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) < e" if ec: "cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k) < e / C" and x: "0 \ x" "x \ 1" for u x proof (cases "(f' (\ x)) = 0") case True then show ?thesis by (simp add: \0 < e\) next case False have "cmod (f' (\ x) * (inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - f' (\ x) / ((\ x - w) * (\ x - w) ^ k)) = cmod (f' (\ x) * ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k))" by (simp add: field_simps) also have "\ = cmod (f' (\ x)) * cmod ((inverse (\ x - u) ^ k - inverse (\ x - w) ^ k) / ((u - w) * k) - inverse (\ x - w) * inverse (\ x - w) ^ k)" by (simp add: norm_mult) also have "\ < cmod (f' (\ x)) * (e/C)" using False mult_strict_left_mono [OF ec] by force also have "\ \ e" using C by (metis False \0 < e\ frac_le less_eq_real_def mult.commute pos_le_divide_eq x zero_less_norm_iff) finally show ?thesis . qed show "\\<^sub>F n in at w. \x\path_image \. cmod (f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / of_nat k - f' x / (x - w) ^ Suc k) < e" using twom [OF divide_pos_pos [OF \0 < e\ \C > 0\]] unfolding path_image_def by (force intro: * elim: eventually_mono) qed show "(\u. f' u / (u - w) ^ (Suc k)) contour_integrable_on \" by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto have *: "(\n. contour_integral \ (\x. f' x * (inverse (x - n) ^ k - inverse (x - w) ^ k) / (n - w) / k)) \w\ contour_integral \ (\u. f' u / (u - w) ^ (Suc k))" by (rule contour_integral_uniform_limit [OF 1 2 leB \]) auto have **: "contour_integral \ (\x. f' x * (inverse (x - u) ^ k - inverse (x - w) ^ k) / ((u - w) * k)) = (f u - f w) / (u - w) / k" if "dist u w < d" for u proof - have u: "u \ s - path_image \" by (metis subsetD d dist_commute mem_ball that) show ?thesis apply (rule contour_integral_unique) apply (simp add: diff_divide_distrib algebra_simps) apply (intro has_contour_integral_diff has_contour_integral_div) using u w apply (simp_all add: field_simps int) done qed show ?thes2 apply (simp add: has_field_derivative_iff del: power_Suc) apply (rule Lim_transform_within [OF tendsto_mult_left [OF *] \0 < d\ ]) apply (simp add: \k \ 0\ **) done qed lemma Cauchy_next_derivative_circlepath: assumes contf: "continuous_on (path_image (circlepath z r)) f" and int: "\w. w \ ball z r \ ((\u. f u / (u - w)^k) has_contour_integral g w) (circlepath z r)" and k: "k \ 0" and w: "w \ ball z r" shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" (is "?thes1") and "(g has_field_derivative (k * contour_integral (circlepath z r) (\u. f u/(u - w)^(Suc k)))) (at w)" (is "?thes2") proof - have "r > 0" using w using ball_eq_empty by fastforce have wim: "w \ ball z r - path_image (circlepath z r)" using w by (auto simp: dist_norm) show ?thes1 ?thes2 by (rule Cauchy_next_derivative [OF contf _ int k open_ball valid_path_circlepath wim, where B = "2 * pi * \r\"]; auto simp: vector_derivative_circlepath norm_mult)+ qed text\ In particular, the first derivative formula.\ lemma Cauchy_derivative_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "(\u. f u/(u - w)^2) contour_integrable_on (circlepath z r)" (is "?thes1") and "(f has_field_derivative (1 / (2 * of_real pi * \) * contour_integral(circlepath z r) (\u. f u / (u - w)^2))) (at w)" (is "?thes2") proof - have [simp]: "r \ 0" using w using ball_eq_empty by fastforce have f: "continuous_on (path_image (circlepath z r)) f" by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def) have int: "\w. dist z w < r \ ((\u. f u / (u - w)) has_contour_integral (\x. 2 * of_real pi * \ * f x) w) (circlepath z r)" by (rule Cauchy_integral_circlepath [OF contf holf]) (simp add: dist_norm norm_minus_commute) show ?thes1 apply (simp add: power2_eq_square) apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1, simplified]) apply (blast intro: int) done have "((\x. 2 * of_real pi * \ * f x) has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2)) (at w)" apply (simp add: power2_eq_square) apply (rule Cauchy_next_derivative_circlepath [OF f _ _ w, where k=1 and g = "\x. 2 * of_real pi * \ * f x", simplified]) apply (blast intro: int) done then have fder: "(f has_field_derivative contour_integral (circlepath z r) (\u. f u / (u - w)^2) / (2 * of_real pi * \)) (at w)" by (rule DERIV_cdivide [where f = "\x. 2 * of_real pi * \ * f x" and c = "2 * of_real pi * \", simplified]) show ?thes2 by simp (rule fder) qed subsection\Existence of all higher derivatives\ proposition derivative_is_holomorphic: assumes "open S" and fder: "\z. z \ S \ (f has_field_derivative f' z) (at z)" shows "f' holomorphic_on S" proof - have *: "\h. (f' has_field_derivative h) (at z)" if "z \ S" for z proof - obtain r where "r > 0" and r: "cball z r \ S" using open_contains_cball \z \ S\ \open S\ by blast then have holf_cball: "f holomorphic_on cball z r" apply (simp add: holomorphic_on_def) using field_differentiable_at_within field_differentiable_def fder by blast then have "continuous_on (path_image (circlepath z r)) f" using \r > 0\ by (force elim: holomorphic_on_subset [THEN holomorphic_on_imp_continuous_on]) then have contfpi: "continuous_on (path_image (circlepath z r)) (\x. 1/(2 * of_real pi*\) * f x)" by (auto intro: continuous_intros)+ have contf_cball: "continuous_on (cball z r) f" using holf_cball by (simp add: holomorphic_on_imp_continuous_on holomorphic_on_subset) have holf_ball: "f holomorphic_on ball z r" using holf_cball using ball_subset_cball holomorphic_on_subset by blast { fix w assume w: "w \ ball z r" have intf: "(\u. f u / (u - w)\<^sup>2) contour_integrable_on circlepath z r" by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) have fder': "(f has_field_derivative 1 / (2 * of_real pi * \) * contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2)) (at w)" by (blast intro: w Cauchy_derivative_integral_circlepath [OF contf_cball holf_ball]) have f'_eq: "f' w = contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)" using fder' ball_subset_cball r w by (force intro: DERIV_unique [OF fder]) have "((\u. f u / (u - w)\<^sup>2 / (2 * of_real pi * \)) has_contour_integral contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) (circlepath z r)" by (rule has_contour_integral_div [OF has_contour_integral_integral [OF intf]]) then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral contour_integral (circlepath z r) (\u. f u / (u - w)\<^sup>2) / (2 * of_real pi * \)) (circlepath z r)" by (simp add: algebra_simps) then have "((\u. f u / (2 * of_real pi * \ * (u - w)\<^sup>2)) has_contour_integral f' w) (circlepath z r)" by (simp add: f'_eq) } note * = this show ?thesis apply (rule exI) apply (rule Cauchy_next_derivative_circlepath [OF contfpi, of 2 f', simplified]) apply (simp_all add: \0 < r\ * dist_norm) done qed show ?thesis by (simp add: holomorphic_on_open [OF \open S\] *) qed lemma holomorphic_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv f) holomorphic_on S" by (metis DERIV_deriv_iff_field_differentiable at_within_open derivative_is_holomorphic holomorphic_on_def) lemma analytic_deriv [analytic_intros]: "f analytic_on S \ (deriv f) analytic_on S" using analytic_on_holomorphic holomorphic_deriv by auto lemma holomorphic_higher_deriv [holomorphic_intros]: "\f holomorphic_on S; open S\ \ (deriv ^^ n) f holomorphic_on S" by (induction n) (auto simp: holomorphic_deriv) lemma analytic_higher_deriv [analytic_intros]: "f analytic_on S \ (deriv ^^ n) f analytic_on S" unfolding analytic_on_def using holomorphic_higher_deriv by blast lemma has_field_derivative_higher_deriv: "\f holomorphic_on S; open S; x \ S\ \ ((deriv ^^ n) f has_field_derivative (deriv ^^ (Suc n)) f x) (at x)" by (metis (no_types, hide_lams) DERIV_deriv_iff_field_differentiable at_within_open comp_apply funpow.simps(2) holomorphic_higher_deriv holomorphic_on_def) lemma valid_path_compose_holomorphic: assumes "valid_path g" and holo:"f holomorphic_on S" and "open S" "path_image g \ S" shows "valid_path (f \ g)" proof (rule valid_path_compose[OF \valid_path g\]) fix x assume "x \ path_image g" then show "f field_differentiable at x" using analytic_on_imp_differentiable_at analytic_on_open assms holo by blast next have "deriv f holomorphic_on S" using holomorphic_deriv holo \open S\ by auto then show "continuous_on (path_image g) (deriv f)" using assms(4) holomorphic_on_imp_continuous_on holomorphic_on_subset by auto qed subsection\Morera's theorem\ lemma Morera_local_triangle_ball: assumes "\z. z \ S \ \e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ (\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" shows "f analytic_on S" proof - { fix z assume "z \ S" with assms obtain e a where "0 < e" and z: "z \ ball a e" and contf: "continuous_on (ball a e) f" and 0: "\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" by fastforce have az: "dist a z < e" using mem_ball z by blast have sb_ball: "ball z (e - dist a z) \ ball a e" by (simp add: dist_commute ball_subset_ball_iff) have "\e>0. f holomorphic_on ball z e" proof (intro exI conjI) have sub_ball: "\y. dist a y < e \ closed_segment a y \ ball a e" by (meson \0 < e\ centre_in_ball convex_ball convex_contains_segment mem_ball) show "f holomorphic_on ball z (e - dist a z)" apply (rule holomorphic_on_subset [OF _ sb_ball]) apply (rule derivative_is_holomorphic[OF open_ball]) apply (rule triangle_contour_integrals_starlike_primitive [OF contf _ open_ball, of a]) apply (simp_all add: 0 \0 < e\ sub_ball) done qed (simp add: az) } then show ?thesis by (simp add: analytic_on_def) qed lemma Morera_local_triangle: assumes "\z. z \ S \ \t. open t \ z \ t \ continuous_on t f \ (\a b c. convex hull {a,b,c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" shows "f analytic_on S" proof - { fix z assume "z \ S" with assms obtain t where "open t" and z: "z \ t" and contf: "continuous_on t f" and 0: "\a b c. convex hull {a,b,c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" by force then obtain e where "e>0" and e: "ball z e \ t" using open_contains_ball by blast have [simp]: "continuous_on (ball z e) f" using contf using continuous_on_subset e by blast have eq0: "\b c. closed_segment b c \ ball z e \ contour_integral (linepath z b) f + contour_integral (linepath b c) f + contour_integral (linepath c z) f = 0" by (meson 0 z \0 < e\ centre_in_ball closed_segment_subset convex_ball dual_order.trans e starlike_convex_subset) have "\e a. 0 < e \ z \ ball a e \ continuous_on (ball a e) f \ (\b c. closed_segment b c \ ball a e \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" using \e > 0\ eq0 by force } then show ?thesis by (simp add: Morera_local_triangle_ball) qed proposition Morera_triangle: "\continuous_on S f; open S; \a b c. convex hull {a,b,c} \ S \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0\ \ f analytic_on S" using Morera_local_triangle by blast subsection\Combining theorems for higher derivatives including Leibniz rule\ lemma higher_deriv_linear [simp]: "(deriv ^^ n) (\w. c*w) = (\z. if n = 0 then c*z else if n = 1 then c else 0)" by (induction n) auto lemma higher_deriv_const [simp]: "(deriv ^^ n) (\w. c) = (\w. if n=0 then c else 0)" by (induction n) auto lemma higher_deriv_ident [simp]: "(deriv ^^ n) (\w. w) z = (if n = 0 then z else if n = 1 then 1 else 0)" apply (induction n, simp) apply (metis higher_deriv_linear lambda_one) done lemma higher_deriv_id [simp]: "(deriv ^^ n) id z = (if n = 0 then z else if n = 1 then 1 else 0)" by (simp add: id_def) lemma has_complex_derivative_funpow_1: "\(f has_field_derivative 1) (at z); f z = z\ \ (f^^n has_field_derivative 1) (at z)" apply (induction n, auto) - apply (metis DERIV_ident DERIV_transform_at id_apply zero_less_one) + 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 DERIV_transform_within_open [of "\w. -((deriv ^^ n) f w)"]) + 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 DERIV_transform_within_open [of "\w. (deriv ^^ n) f w + (deriv ^^ n) g w"]) + 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 DERIV_transform_within_open + apply (rule has_field_derivative_transform_within_open [of "\w. (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f w * (deriv ^^ (n - i)) g w)"]) apply (simp add: algebra_simps) apply (rule DERIV_cong [OF DERIV_sum]) apply (rule DERIV_cmult) apply (auto intro: DERIV_mult * sumeq \open S\ Suc.prems Suc.IH [symmetric]) done then show ?case unfolding funpow.simps o_apply by (simp add: DERIV_imp_deriv) qed lemma higher_deriv_transform_within_open: fixes z::complex assumes "f holomorphic_on S" "g holomorphic_on S" "open S" and z: "z \ S" and fg: "\w. w \ S \ f w = g w" shows "(deriv ^^ i) f z = (deriv ^^ i) g z" using z by (induction i arbitrary: z) (auto simp: fg intro: complex_derivative_transform_within_open holomorphic_higher_deriv assms) lemma higher_deriv_compose_linear: fixes z::complex assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \ S" and fg: "\w. w \ S \ u * w \ T" shows "(deriv ^^ n) (\w. f (u * w)) z = u^n * (deriv ^^ n) f (u * z)" using z proof (induction n arbitrary: z) case 0 then show ?case by simp next case (Suc n z) have holo0: "f holomorphic_on (*) u ` S" by (meson fg f holomorphic_on_subset image_subset_iff) have holo2: "(deriv ^^ n) f holomorphic_on (*) u ` S" by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T) have holo3: "(\z. u ^ n * (deriv ^^ n) f (u * z)) holomorphic_on S" by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros) have holo1: "(\w. f (u * w)) holomorphic_on S" apply (rule holomorphic_on_compose [where g=f, unfolded o_def]) apply (rule holo0 holomorphic_intros)+ done have "deriv ((deriv ^^ n) (\w. f (u * w))) z = deriv (\z. u^n * (deriv ^^ n) f (u*z)) z" apply (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems]) apply (rule holomorphic_higher_deriv [OF holo1 S]) apply (simp add: Suc.IH) done also have "\ = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" apply (rule deriv_cmult) apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and T=T, unfolded o_def]) apply (simp add: analytic_on_linear) apply (simp add: analytic_on_open f holomorphic_higher_deriv T) apply (blast intro: fg) done also have "\ = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" apply (subst complex_derivative_chain [where g = "(deriv ^^ n) f" and f = "(*) u", unfolded o_def]) apply (rule derivative_intros) using Suc.prems field_differentiable_def f fg has_field_derivative_higher_deriv T apply blast apply (simp add: deriv_linear) done finally show ?case by simp qed lemma higher_deriv_add_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w + g w) z = (deriv ^^ n) f z + (deriv ^^ n) g z" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_add show ?thesis by (auto simp: analytic_at_two) qed lemma higher_deriv_diff_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w - g w) z = (deriv ^^ n) f z - (deriv ^^ n) g z" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_diff show ?thesis by (auto simp: analytic_at_two) qed lemma higher_deriv_uminus_at: "f analytic_on {z} \ (deriv ^^ n) (\w. -(f w)) z = - ((deriv ^^ n) f z)" using higher_deriv_uminus by (auto simp: analytic_at) lemma higher_deriv_mult_at: assumes "f analytic_on {z}" "g analytic_on {z}" shows "(deriv ^^ n) (\w. f w * g w) z = (\i = 0..n. of_nat (n choose i) * (deriv ^^ i) f z * (deriv ^^ (n - i)) g z)" proof - have "f analytic_on {z} \ g analytic_on {z}" using assms by blast with higher_deriv_mult show ?thesis by (auto simp: analytic_at_two) qed text\ Nonexistence of isolated singularities and a stronger integral formula.\ proposition no_isolated_singularity: fixes z::complex assumes f: "continuous_on S f" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" shows "f holomorphic_on S" proof - { fix z assume "z \ S" and cdf: "\x. x \ S - K \ f field_differentiable at x" have "f field_differentiable at z" proof (cases "z \ K") case False then show ?thesis by (blast intro: cdf \z \ S\) next case True with finite_set_avoid [OF K, of z] obtain d where "d>0" and d: "\x. \x\K; x \ z\ \ d \ dist z x" by blast obtain e where "e>0" and e: "ball z e \ S" using S \z \ S\ by (force simp: open_contains_ball) have fde: "continuous_on (ball z (min d e)) f" by (metis Int_iff ball_min_Int continuous_on_subset e f subsetI) have cont: "{a,b,c} \ ball z (min d e) \ continuous_on (convex hull {a, b, c}) f" for a b c by (simp add: hull_minimal continuous_on_subset [OF fde]) have fd: "\{a,b,c} \ ball z (min d e); x \ interior (convex hull {a, b, c}) - K\ \ f field_differentiable at x" for a b c x by (metis cdf Diff_iff Int_iff ball_min_Int subsetD convex_ball e interior_mono interior_subset subset_hull) obtain g where "\w. w \ ball z (min d e) \ (g has_field_derivative f w) (at w within ball z (min d e))" apply (rule contour_integral_convex_primitive [OF convex_ball fde Cauchy_theorem_triangle_cofinite [OF _ K]]) using cont fd by auto then have "f holomorphic_on ball z (min d e)" by (metis open_ball at_within_open derivative_is_holomorphic) then show ?thesis unfolding holomorphic_on_def by (metis open_ball \0 < d\ \0 < e\ at_within_open centre_in_ball min_less_iff_conj) qed } with holf S K show ?thesis by (simp add: holomorphic_on_open open_Diff finite_imp_closed field_differentiable_def [symmetric]) qed lemma no_isolated_singularity': fixes z::complex assumes f: "\z. z \ K \ (f \ f z) (at z within S)" and holf: "f holomorphic_on (S - K)" and S: "open S" and K: "finite K" shows "f holomorphic_on S" proof (rule no_isolated_singularity[OF _ assms(2-)]) show "continuous_on S f" unfolding continuous_on_def proof fix z assume z: "z \ S" show "(f \ f z) (at z within S)" proof (cases "z \ K") case False from holf have "continuous_on (S - K) f" by (rule holomorphic_on_imp_continuous_on) with z False have "(f \ f z) (at z within (S - K))" by (simp add: continuous_on_def) also from z K S False have "at z within (S - K) = at z within S" by (subst (1 2) at_within_open) (auto intro: finite_imp_closed) finally show "(f \ f z) (at z within S)" . qed (insert assms z, simp_all) qed qed proposition Cauchy_integral_formula_convex: assumes S: "convex S" and K: "finite K" and contf: "continuous_on S f" and fcd: "(\x. x \ interior S - K \ f field_differentiable at x)" and z: "z \ interior S" and vpg: "valid_path \" and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - have *: "\x. x \ interior S \ f field_differentiable at x" unfolding holomorphic_on_open [symmetric] field_differentiable_def using no_isolated_singularity [where S = "interior S"] by (meson K contf continuous_at_imp_continuous_on continuous_on_interior fcd field_differentiable_at_within field_differentiable_def holomorphic_onI holomorphic_on_imp_differentiable_at open_interior) show ?thesis by (rule Cauchy_integral_formula_weak [OF S finite.emptyI contf]) (use * assms in auto) qed text\ Formula for higher derivatives.\ lemma Cauchy_has_contour_integral_higher_derivative_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "((\u. f u / (u - w) ^ (Suc k)) has_contour_integral ((2 * pi * \) / (fact k) * (deriv ^^ k) f w)) (circlepath z r)" using w proof (induction k arbitrary: w) case 0 then show ?case using assms by (auto simp: Cauchy_integral_circlepath dist_commute dist_norm) next case (Suc k) have [simp]: "r > 0" using w using ball_eq_empty by fastforce have f: "continuous_on (path_image (circlepath z r)) f" by (rule continuous_on_subset [OF contf]) (force simp: cball_def sphere_def less_imp_le) obtain X where X: "((\u. f u / (u - w) ^ Suc (Suc k)) has_contour_integral X) (circlepath z r)" using Cauchy_next_derivative_circlepath(1) [OF f Suc.IH _ Suc.prems] by (auto simp: contour_integrable_on_def) then have con: "contour_integral (circlepath z r) ((\u. f u / (u - w) ^ Suc (Suc k))) = X" by (rule contour_integral_unique) have "\n. ((deriv ^^ n) f has_field_derivative deriv ((deriv ^^ n) f) w) (at w)" using Suc.prems assms has_field_derivative_higher_deriv by auto then have dnf_diff: "\n. (deriv ^^ n) f field_differentiable (at w)" by (force simp: field_differentiable_def) have "deriv (\w. complex_of_real (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) w = of_nat (Suc k) * contour_integral (circlepath z r) (\u. f u / (u - w) ^ Suc (Suc k))" by (force intro!: DERIV_imp_deriv Cauchy_next_derivative_circlepath [OF f Suc.IH _ Suc.prems]) also have "\ = of_nat (Suc k) * X" by (simp only: con) finally have "deriv (\w. ((2 * pi) * \ / (fact k)) * (deriv ^^ k) f w) w = of_nat (Suc k) * X" . then have "((2 * pi) * \ / (fact k)) * deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X" by (metis deriv_cmult dnf_diff) then have "deriv (\w. (deriv ^^ k) f w) w = of_nat (Suc k) * X / ((2 * pi) * \ / (fact k))" by (simp add: field_simps) then show ?case using of_nat_eq_0_iff X by fastforce qed lemma Cauchy_higher_derivative_integral_circlepath: assumes contf: "continuous_on (cball z r) f" and holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "(\u. f u / (u - w)^(Suc k)) contour_integrable_on (circlepath z r)" (is "?thes1") and "(deriv ^^ k) f w = (fact k) / (2 * pi * \) * contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k))" (is "?thes2") proof - have *: "((\u. f u / (u - w) ^ Suc k) has_contour_integral (2 * pi) * \ / (fact k) * (deriv ^^ k) f w) (circlepath z r)" using Cauchy_has_contour_integral_higher_derivative_circlepath [OF assms] by simp show ?thes1 using * using contour_integrable_on_def by blast show ?thes2 unfolding contour_integral_unique [OF *] by (simp add: field_split_simps) qed corollary Cauchy_contour_integral_circlepath: assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" shows "contour_integral(circlepath z r) (\u. f u/(u - w)^(Suc k)) = (2 * pi * \) * (deriv ^^ k) f w / (fact k)" by (simp add: Cauchy_higher_derivative_integral_circlepath [OF assms]) lemma Cauchy_contour_integral_circlepath_2: assumes "continuous_on (cball z r) f" "f holomorphic_on ball z r" "w \ ball z r" shows "contour_integral(circlepath z r) (\u. f u/(u - w)^2) = (2 * pi * \) * deriv f w" using Cauchy_contour_integral_circlepath [OF assms, of 1] by (simp add: power2_eq_square) subsection\A holomorphic function is analytic, i.e. has local power series\ theorem holomorphic_power_series: assumes holf: "f holomorphic_on ball z r" and w: "w \ ball z r" shows "((\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" proof - \ \Replacing \<^term>\r\ and the original (weak) premises with stronger ones\ obtain r where "r > 0" and holfc: "f holomorphic_on cball z r" and w: "w \ ball z r" proof have "cball z ((r + dist w z) / 2) \ ball z r" using w by (simp add: dist_commute field_sum_of_halves subset_eq) then show "f holomorphic_on cball z ((r + dist w z) / 2)" by (rule holomorphic_on_subset [OF holf]) have "r > 0" using w by clarsimp (metis dist_norm le_less_trans norm_ge_zero) then show "0 < (r + dist w z) / 2" by simp (use zero_le_dist [of w z] in linarith) qed (use w in \auto simp: dist_commute\) then have holf: "f holomorphic_on ball z r" using ball_subset_cball holomorphic_on_subset by blast have contf: "continuous_on (cball z r) f" by (simp add: holfc holomorphic_on_imp_continuous_on) have cint: "\k. (\u. f u / (u - z) ^ Suc k) contour_integrable_on circlepath z r" by (rule Cauchy_higher_derivative_integral_circlepath [OF contf holf]) (simp add: \0 < r\) obtain B where "0 < B" and B: "\u. u \ cball z r \ norm(f u) \ B" by (metis (no_types) bounded_pos compact_cball compact_continuous_image compact_imp_bounded contf image_eqI) obtain k where k: "0 < k" "k \ r" and wz_eq: "norm(w - z) = r - k" and kle: "\u. norm(u - z) = r \ k \ norm(u - w)" proof show "\u. cmod (u - z) = r \ r - dist z w \ cmod (u - w)" by (metis add_diff_eq diff_add_cancel dist_norm norm_diff_ineq) qed (use w in \auto simp: dist_norm norm_minus_commute\) have ul: "uniform_limit (sphere z r) (\n x. (\kx. f x / (x - w)) sequentially" unfolding uniform_limit_iff dist_norm proof clarify fix e::real assume "0 < e" have rr: "0 \ (r - k) / r" "(r - k) / r < 1" using k by auto obtain n where n: "((r - k) / r) ^ n < e / B * k" using real_arch_pow_inv [of "e/B*k" "(r - k)/r"] \0 < e\ \0 < B\ k by force have "norm ((\k N" and r: "r = dist z u" for N u proof - have N: "((r - k) / r) ^ N < e / B * k" apply (rule le_less_trans [OF power_decreasing n]) using \n \ N\ k by auto have u [simp]: "(u \ z) \ (u \ w)" using \0 < r\ r w by auto have wzu_not1: "(w - z) / (u - z) \ 1" by (metis (no_types) dist_norm divide_eq_1_iff less_irrefl mem_ball norm_minus_commute r w) have "norm ((\kk = norm ((((w - z) / (u - z)) ^ N - 1) * (u - w) / (((w - z) / (u - z) - 1) * (u - z)) - 1) * norm (f u)" using \0 < B\ apply (auto simp: geometric_sum [OF wzu_not1]) apply (simp add: field_simps norm_mult [symmetric]) done also have "\ = norm ((u-z) ^ N * (w - u) - ((w - z) ^ N - (u-z) ^ N) * (u-w)) / (r ^ N * norm (u-w)) * norm (f u)" using \0 < r\ r by (simp add: divide_simps norm_mult norm_divide norm_power dist_norm norm_minus_commute) also have "\ = norm ((w - z) ^ N * (w - u)) / (r ^ N * norm (u - w)) * norm (f u)" by (simp add: algebra_simps) also have "\ = norm (w - z) ^ N * norm (f u) / r ^ N" by (simp add: norm_mult norm_power norm_minus_commute) also have "\ \ (((r - k)/r)^N) * B" using \0 < r\ w k apply (simp add: divide_simps) apply (rule mult_mono [OF power_mono]) apply (auto simp: norm_divide wz_eq norm_power dist_norm norm_minus_commute B r) done also have "\ < e * k" using \0 < B\ N by (simp add: divide_simps) also have "\ \ e * norm (u - w)" using r kle \0 < e\ by (simp add: dist_commute dist_norm) finally show ?thesis by (simp add: field_split_simps norm_divide del: power_Suc) qed with \0 < r\ show "\\<^sub>F n in sequentially. \x\sphere z r. norm ((\k\<^sub>F x in sequentially. contour_integral (circlepath z r) (\u. \kku. f u / (u - z) ^ Suc k) * (w - z) ^ k)" apply (rule eventuallyI) apply (subst contour_integral_sum, simp) using contour_integrable_lmul [OF cint, of "(w - z) ^ a" for a] apply (simp add: field_simps) apply (simp only: contour_integral_lmul cint algebra_simps) done have cic: "\u. (\y. \k0 < r\ by (force intro!: Cauchy_higher_derivative_integral_circlepath [OF contf holf]) have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) sums contour_integral (circlepath z r) (\u. f u/(u - w))" unfolding sums_def apply (intro Lim_transform_eventually [OF _ eq] contour_integral_uniform_limit_circlepath [OF eventuallyI ul] cic) using \0 < r\ apply auto done then have "(\k. contour_integral (circlepath z r) (\u. f u/(u - z)^(Suc k)) * (w - z)^k) sums (2 * of_real pi * \ * f w)" using w by (auto simp: dist_commute dist_norm contour_integral_unique [OF Cauchy_integral_circlepath_simple [OF holfc]]) then have "(\k. contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc k) * (w - z)^k / (\ * (of_real pi * 2))) sums ((2 * of_real pi * \ * f w) / (\ * (complex_of_real pi * 2)))" by (rule sums_divide) then have "(\n. (w - z) ^ n * contour_integral (circlepath z r) (\u. f u / (u - z) ^ Suc n) / (\ * (of_real pi * 2))) sums f w" by (simp add: field_simps) then show ?thesis by (simp add: field_simps \0 < r\ Cauchy_higher_derivative_integral_circlepath [OF contf holf]) qed subsection\The Liouville theorem and the Fundamental Theorem of Algebra\ text\ These weak Liouville versions don't even need the derivative formula.\ lemma Liouville_weak_0: assumes holf: "f holomorphic_on UNIV" and inf: "(f \ 0) at_infinity" shows "f z = 0" proof (rule ccontr) assume fz: "f z \ 0" with inf [unfolded Lim_at_infinity, rule_format, of "norm(f z)/2"] obtain B where B: "\x. B \ cmod x \ norm (f x) * 2 < cmod (f z)" by (auto simp: dist_norm) define R where "R = 1 + \B\ + norm z" have "R > 0" unfolding R_def proof - have "0 \ cmod z + \B\" by (metis (full_types) add_nonneg_nonneg norm_ge_zero real_norm_def) then show "0 < 1 + \B\ + cmod z" by linarith qed have *: "((\u. f u / (u - z)) has_contour_integral 2 * complex_of_real pi * \ * f z) (circlepath z R)" apply (rule Cauchy_integral_circlepath) using \R > 0\ apply (auto intro: holomorphic_on_subset [OF holf] holomorphic_on_imp_continuous_on)+ done have "cmod (x - z) = R \ cmod (f x) * 2 < cmod (f z)" for x unfolding R_def by (rule B) (use norm_triangle_ineq4 [of x z] in auto) with \R > 0\ fz show False using has_contour_integral_bound_circlepath [OF *, of "norm(f z)/2/R"] by (auto simp: less_imp_le norm_mult norm_divide field_split_simps) qed proposition Liouville_weak: assumes "f holomorphic_on UNIV" and "(f \ l) at_infinity" shows "f z = l" using Liouville_weak_0 [of "\z. f z - l"] by (simp add: assms holomorphic_on_const holomorphic_on_diff LIM_zero) proposition Liouville_weak_inverse: assumes "f holomorphic_on UNIV" and unbounded: "\B. eventually (\x. norm (f x) \ B) at_infinity" obtains z where "f z = 0" proof - { assume f: "\z. f z \ 0" have 1: "(\x. 1 / f x) holomorphic_on UNIV" by (simp add: holomorphic_on_divide holomorphic_on_const assms f) have 2: "((\x. 1 / f x) \ 0) at_infinity" apply (rule tendstoI [OF eventually_mono]) apply (rule_tac B="2/e" in unbounded) apply (simp add: dist_norm norm_divide field_split_simps mult_ac) done have False using Liouville_weak_0 [OF 1 2] f by simp } then show ?thesis using that by blast qed text\ In particular we get the Fundamental Theorem of Algebra.\ theorem fundamental_theorem_of_algebra: fixes a :: "nat \ complex" assumes "a 0 = 0 \ (\i \ {1..n}. a i \ 0)" obtains z where "(\i\n. a i * z^i) = 0" using assms proof (elim disjE bexE) assume "a 0 = 0" then show ?thesis by (auto simp: that [of 0]) next fix i assume i: "i \ {1..n}" and nz: "a i \ 0" have 1: "(\z. \i\n. a i * z^i) holomorphic_on UNIV" by (rule holomorphic_intros)+ show thesis proof (rule Liouville_weak_inverse [OF 1]) show "\\<^sub>F x in at_infinity. B \ cmod (\i\n. a i * x ^ i)" for B using i polyfun_extremal nz by force qed (use that in auto) qed subsection\Weierstrass convergence theorem\ lemma holomorphic_uniform_limit: assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (f n) holomorphic_on ball z r) F" and ulim: "uniform_limit (cball z r) f g F" and F: "\ trivial_limit F" obtains "continuous_on (cball z r) g" "g holomorphic_on ball z r" proof (cases r "0::real" rule: linorder_cases) case less then show ?thesis by (force simp: ball_empty less_imp_le continuous_on_def holomorphic_on_def intro: that) next case equal then show ?thesis by (force simp: holomorphic_on_def intro: that) next case greater have contg: "continuous_on (cball z r) g" using cont uniform_limit_theorem [OF eventually_mono ulim F] by blast have "path_image (circlepath z r) \ cball z r" using \0 < r\ by auto then have 1: "continuous_on (path_image (circlepath z r)) (\x. 1 / (2 * complex_of_real pi * \) * g x)" by (intro continuous_intros continuous_on_subset [OF contg]) have 2: "((\u. 1 / (2 * of_real pi * \) * g u / (u - w) ^ 1) has_contour_integral g w) (circlepath z r)" if w: "w \ ball z r" for w proof - define d where "d = (r - norm(w - z))" have "0 < d" "d \ r" using w by (auto simp: norm_minus_commute d_def dist_norm) have dle: "\u. cmod (z - u) = r \ d \ cmod (u - w)" unfolding d_def by (metis add_diff_eq diff_add_cancel norm_diff_ineq norm_minus_commute) have ev_int: "\\<^sub>F n in F. (\u. f n u / (u - w)) contour_integrable_on circlepath z r" apply (rule eventually_mono [OF cont]) using w apply (auto intro: Cauchy_higher_derivative_integral_circlepath [where k=0, simplified]) done have ul_less: "uniform_limit (sphere z r) (\n x. f n x / (x - w)) (\x. g x / (x - w)) F" using greater \0 < d\ apply (clarsimp simp add: uniform_limit_iff dist_norm norm_divide diff_divide_distrib [symmetric] divide_simps) apply (rule_tac e1="e * d" in eventually_mono [OF uniform_limitD [OF ulim]]) apply (force simp: dist_norm intro: dle mult_left_mono less_le_trans)+ done have g_cint: "(\u. g u/(u - w)) contour_integrable_on circlepath z r" by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) have cif_tends_cig: "((\n. contour_integral(circlepath z r) (\u. f n u / (u - w))) \ contour_integral(circlepath z r) (\u. g u/(u - w))) F" by (rule contour_integral_uniform_limit_circlepath [OF ev_int ul_less F \0 < r\]) have f_tends_cig: "((\n. 2 * of_real pi * \ * f n w) \ contour_integral (circlepath z r) (\u. g u / (u - w))) F" proof (rule Lim_transform_eventually) show "\\<^sub>F x in F. contour_integral (circlepath z r) (\u. f x u / (u - w)) = 2 * of_real pi * \ * f x w" apply (rule eventually_mono [OF cont contour_integral_unique [OF Cauchy_integral_circlepath]]) using w\0 < d\ d_def by auto qed (auto simp: cif_tends_cig) have "\e. 0 < e \ \\<^sub>F n in F. dist (f n w) (g w) < e" by (rule eventually_mono [OF uniform_limitD [OF ulim]]) (use w in auto) then have "((\n. 2 * of_real pi * \ * f n w) \ 2 * of_real pi * \ * g w) F" by (rule tendsto_mult_left [OF tendstoI]) then have "((\u. g u / (u - w)) has_contour_integral 2 * of_real pi * \ * g w) (circlepath z r)" using has_contour_integral_integral [OF g_cint] tendsto_unique [OF F f_tends_cig] w by fastforce then have "((\u. g u / (2 * of_real pi * \ * (u - w))) has_contour_integral g w) (circlepath z r)" using has_contour_integral_div [where c = "2 * of_real pi * \"] by (force simp: field_simps) then show ?thesis by (simp add: dist_norm) qed show ?thesis using Cauchy_next_derivative_circlepath(2) [OF 1 2, simplified] by (fastforce simp add: holomorphic_on_open contg intro: that) qed text\ Version showing that the limit is the limit of the derivatives.\ proposition has_complex_derivative_uniform_limit: fixes z::complex assumes cont: "eventually (\n. continuous_on (cball z r) (f n) \ (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))) F" and ulim: "uniform_limit (cball z r) f g F" and F: "\ trivial_limit F" and "0 < r" obtains g' where "continuous_on (cball z r) g" "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" proof - let ?conint = "contour_integral (circlepath z r)" have g: "continuous_on (cball z r) g" "g holomorphic_on ball z r" by (rule holomorphic_uniform_limit [OF eventually_mono [OF cont] ulim F]; auto simp: holomorphic_on_open field_differentiable_def)+ then obtain g' where g': "\x. x \ ball z r \ (g has_field_derivative g' x) (at x)" using DERIV_deriv_iff_has_field_derivative by (fastforce simp add: holomorphic_on_open) then have derg: "\x. x \ ball z r \ deriv g x = g' x" by (simp add: DERIV_imp_deriv) have tends_f'n_g': "((\n. f' n w) \ g' w) F" if w: "w \ ball z r" for w proof - have eq_f': "?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2) = (f' n w - g' w) * (2 * of_real pi * \)" if cont_fn: "continuous_on (cball z r) (f n)" and fnd: "\w. w \ ball z r \ (f n has_field_derivative f' n w) (at w)" for n proof - have hol_fn: "f n holomorphic_on ball z r" using fnd by (force simp: holomorphic_on_open) have "(f n has_field_derivative 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)) (at w)" by (rule Cauchy_derivative_integral_circlepath [OF cont_fn hol_fn w]) then have f': "f' n w = 1 / (2 * of_real pi * \) * ?conint (\u. f n u / (u - w)\<^sup>2)" using DERIV_unique [OF fnd] w by blast show ?thesis by (simp add: f' Cauchy_contour_integral_circlepath_2 [OF g w] derg [OF w] field_split_simps) qed define d where "d = (r - norm(w - z))^2" have "d > 0" using w by (simp add: dist_commute dist_norm d_def) have dle: "d \ cmod ((y - w)\<^sup>2)" if "r = cmod (z - y)" for y proof - have "w \ ball z (cmod (z - y))" using that w by fastforce then have "cmod (w - z) \ cmod (z - y)" by (simp add: dist_complex_def norm_minus_commute) moreover have "cmod (z - y) - cmod (w - z) \ cmod (y - w)" by (metis diff_add_cancel diff_add_eq_diff_diff_swap norm_minus_commute norm_triangle_ineq2) ultimately show ?thesis using that by (simp add: d_def norm_power power_mono) qed have 1: "\\<^sub>F n in F. (\x. f n x / (x - w)\<^sup>2) contour_integrable_on circlepath z r" by (force simp: holomorphic_on_open intro: w Cauchy_derivative_integral_circlepath eventually_mono [OF cont]) have 2: "uniform_limit (sphere z r) (\n x. f n x / (x - w)\<^sup>2) (\x. g x / (x - w)\<^sup>2) F" unfolding uniform_limit_iff proof clarify fix e::real assume "0 < e" with \r > 0\ show "\\<^sub>F n in F. \x\sphere z r. dist (f n x / (x - w)\<^sup>2) (g x / (x - w)\<^sup>2) < e" apply (simp add: norm_divide field_split_simps sphere_def dist_norm) apply (rule eventually_mono [OF uniform_limitD [OF ulim], of "e*d"]) apply (simp add: \0 < d\) apply (force simp: dist_norm dle intro: less_le_trans) done qed have "((\n. contour_integral (circlepath z r) (\x. f n x / (x - w)\<^sup>2)) \ contour_integral (circlepath z r) ((\x. g x / (x - w)\<^sup>2))) F" by (rule contour_integral_uniform_limit_circlepath [OF 1 2 F \0 < r\]) then have tendsto_0: "((\n. 1 / (2 * of_real pi * \) * (?conint (\x. f n x / (x - w)\<^sup>2) - ?conint (\x. g x / (x - w)\<^sup>2))) \ 0) F" using Lim_null by (force intro!: tendsto_mult_right_zero) have "((\n. f' n w - g' w) \ 0) F" apply (rule Lim_transform_eventually [OF tendsto_0]) apply (force simp: divide_simps intro: eq_f' eventually_mono [OF cont]) done then show ?thesis using Lim_null by blast qed obtain g' where "\w. w \ ball z r \ (g has_field_derivative (g' w)) (at w) \ ((\n. f' n w) \ g' w) F" by (blast intro: tends_f'n_g' g') then show ?thesis using g using that by blast qed subsection\<^marker>\tag unimportant\ \Some more simple/convenient versions for applications\ lemma holomorphic_uniform_sequence: assumes S: "open S" and hol_fn: "\n. (f n) holomorphic_on S" and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" shows "g holomorphic_on S" proof - have "\f'. (g has_field_derivative f') (at z)" if "z \ S" for z proof - obtain r where "0 < r" and r: "cball z r \ S" and ul: "uniform_limit (cball z r) f g sequentially" using ulim_g [OF \z \ S\] by blast have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ f n holomorphic_on ball z r" proof (intro eventuallyI conjI) show "continuous_on (cball z r) (f x)" for x using hol_fn holomorphic_on_imp_continuous_on holomorphic_on_subset r by blast show "f x holomorphic_on ball z r" for x by (metis hol_fn holomorphic_on_subset interior_cball interior_subset r) qed show ?thesis apply (rule holomorphic_uniform_limit [OF *]) using \0 < r\ centre_in_ball ul apply (auto simp: holomorphic_on_open) done qed with S show ?thesis by (simp add: holomorphic_on_open) qed lemma has_complex_derivative_uniform_sequence: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ ((f n) has_field_derivative f' n x) (at x)" and ulim_g: "\x. x \ S \ \d. 0 < d \ cball x d \ S \ uniform_limit (cball x d) f g sequentially" shows "\g'. \x \ S. (g has_field_derivative g' x) (at x) \ ((\n. f' n x) \ g' x) sequentially" proof - have y: "\y. (g has_field_derivative y) (at z) \ (\n. f' n z) \ y" if "z \ S" for z proof - obtain r where "0 < r" and r: "cball z r \ S" and ul: "uniform_limit (cball z r) f g sequentially" using ulim_g [OF \z \ S\] by blast have *: "\\<^sub>F n in sequentially. continuous_on (cball z r) (f n) \ (\w \ ball z r. ((f n) has_field_derivative (f' n w)) (at w))" proof (intro eventuallyI conjI ballI) show "continuous_on (cball z r) (f x)" for x by (meson S continuous_on_subset hfd holomorphic_on_imp_continuous_on holomorphic_on_open r) show "w \ ball z r \ (f x has_field_derivative f' x w) (at w)" for w x using ball_subset_cball hfd r by blast qed show ?thesis by (rule has_complex_derivative_uniform_limit [OF *, of g]) (use \0 < r\ ul in \force+\) qed show ?thesis by (rule bchoice) (blast intro: y) qed subsection\On analytic functions defined by a series\ lemma series_and_derivative_comparison: fixes S :: "complex set" assumes S: "open S" and h: "summable h" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\\<^sub>F n in sequentially. \x\S. norm (f n x) \ h n" obtains g g' where "\x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" proof - obtain g where g: "uniform_limit S (\n x. \id>0. cball x d \ S \ uniform_limit (cball x d) (\n x. \i S" for x proof - obtain d where "d>0" and d: "cball x d \ S" using open_contains_cball [of "S"] \x \ S\ S by blast show ?thesis proof (intro conjI exI) show "uniform_limit (cball x d) (\n x. \id > 0\ d in auto) qed have "\x. x \ S \ (\n. \i g x" by (metis tendsto_uniform_limitI [OF g]) moreover have "\g'. \x\S. (g has_field_derivative g' x) (at x) \ (\n. \i g' x" by (rule has_complex_derivative_uniform_sequence [OF S]) (auto intro: * hfd DERIV_sum)+ ultimately show ?thesis by (metis sums_def that) qed text\A version where we only have local uniform/comparative convergence.\ lemma series_and_derivative_comparison_local: fixes S :: "complex set" assumes S: "open S" and hfd: "\n x. x \ S \ (f n has_field_derivative f' n x) (at x)" and to_g: "\x. x \ S \ \d h. 0 < d \ summable h \ (\\<^sub>F n in sequentially. \y\ball x d \ S. norm (f n y) \ h n)" shows "\g g'. \x \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" proof - have "\y. (\n. f n z) sums (\n. f n z) \ (\n. f' n z) sums y \ ((\x. \n. f n x) has_field_derivative y) (at z)" if "z \ S" for z proof - obtain d h where "0 < d" "summable h" and le_h: "\\<^sub>F n in sequentially. \y\ball z d \ S. norm (f n y) \ h n" using to_g \z \ S\ by meson then obtain r where "r>0" and r: "ball z r \ ball z d \ S" using \z \ S\ S by (metis Int_iff open_ball centre_in_ball open_Int open_contains_ball_eq) have 1: "open (ball z d \ S)" by (simp add: open_Int S) have 2: "\n x. x \ ball z d \ S \ (f n has_field_derivative f' n x) (at x)" by (auto simp: hfd) obtain g g' where gg': "\x \ ball z d \ S. ((\n. f n x) sums g x) \ ((\n. f' n x) sums g' x) \ (g has_field_derivative g' x) (at x)" by (auto intro: le_h series_and_derivative_comparison [OF 1 \summable h\ hfd]) then have "(\n. f' n z) sums g' z" by (meson \0 < r\ centre_in_ball contra_subsetD r) moreover have "(\n. f n z) sums (\n. f n z)" using summable_sums centre_in_ball \0 < d\ \summable h\ le_h by (metis (full_types) Int_iff gg' summable_def that) moreover have "((\x. \n. f n x) has_field_derivative g' z) (at z)" - proof (rule DERIV_transform_at) + 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 DERIV_transform_at [where d="(r - norm(z - w))/2"]) + by (rule has_field_derivative_transform_within [where d="(r - norm(z - w))/2"]) (use w gg' [of w] in \(force simp: dist_norm)+\) then show ?thesis .. qed then show ?thesis by (simp add: holomorphic_on_open) qed corollary holomorphic_iff_power_series: "f holomorphic_on ball z r \ (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" apply (intro iffI ballI holomorphic_power_series, assumption+) apply (force intro: power_series_holomorphic [where a = "\n. (deriv ^^ n) f z / (fact n)"]) done lemma power_series_analytic: "(\w. w \ ball z r \ (\n. a n*(w - z)^n) sums f w) \ f analytic_on ball z r" by (force simp: analytic_on_open intro!: power_series_holomorphic) lemma analytic_iff_power_series: "f analytic_on ball z r \ (\w \ ball z r. (\n. (deriv ^^ n) f z / (fact n) * (w - z)^n) sums f w)" by (simp add: analytic_on_open holomorphic_iff_power_series) subsection\<^marker>\tag unimportant\ \Equality between holomorphic functions, on open ball then connected set\ lemma holomorphic_fun_eq_on_ball: "\f holomorphic_on ball z r; g holomorphic_on ball z r; w \ ball z r; \n. (deriv ^^ n) f z = (deriv ^^ n) g z\ \ f w = g w" apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) apply (auto simp: holomorphic_iff_power_series) done lemma holomorphic_fun_eq_0_on_ball: "\f holomorphic_on ball z r; w \ ball z r; \n. (deriv ^^ n) f z = 0\ \ f w = 0" apply (rule sums_unique2 [of "\n. (deriv ^^ n) f z / (fact n) * (w - z)^n"]) apply (auto simp: holomorphic_iff_power_series) done lemma holomorphic_fun_eq_0_on_connected: assumes holf: "f holomorphic_on S" and "open S" and cons: "connected S" and der: "\n. (deriv ^^ n) f z = 0" and "z \ S" "w \ S" shows "f w = 0" proof - have *: "ball x e \ (\n. {w \ S. (deriv ^^ n) f w = 0})" if "\u. (deriv ^^ u) f x = 0" "ball x e \ S" for x e proof - have "\x' n. dist x x' < e \ (deriv ^^ n) f x' = 0" apply (rule holomorphic_fun_eq_0_on_ball [OF holomorphic_higher_deriv]) apply (rule holomorphic_on_subset [OF holf]) using that apply simp_all by (metis funpow_add o_apply) with that show ?thesis by auto qed have 1: "openin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" apply (rule open_subset, force) using \open S\ apply (simp add: open_contains_ball Ball_def) apply (erule all_forward) using "*" by auto blast+ have 2: "closedin (top_of_set S) (\n. {w \ S. (deriv ^^ n) f w = 0})" using assms by (auto intro: continuous_closedin_preimage_constant holomorphic_on_imp_continuous_on holomorphic_higher_deriv) obtain e where "e>0" and e: "ball w e \ S" using openE [OF \open S\ \w \ S\] . then have holfb: "f holomorphic_on ball w e" using holf holomorphic_on_subset by blast have 3: "(\n. {w \ S. (deriv ^^ n) f w = 0}) = S \ f w = 0" using \e>0\ e by (force intro: holomorphic_fun_eq_0_on_ball [OF holfb]) show ?thesis using cons der \z \ S\ apply (simp add: connected_clopen) apply (drule_tac x="\n. {w \ S. (deriv ^^ n) f w = 0}" in spec) apply (auto simp: 1 2 3) done qed lemma holomorphic_fun_eq_on_connected: assumes "f holomorphic_on S" "g holomorphic_on S" and "open S" "connected S" and "\n. (deriv ^^ n) f z = (deriv ^^ n) g z" and "z \ S" "w \ S" shows "f w = g w" proof (rule holomorphic_fun_eq_0_on_connected [of "\x. f x - g x" S z, simplified]) show "(\x. f x - g x) holomorphic_on S" by (intro assms holomorphic_intros) show "\n. (deriv ^^ n) (\x. f x - g x) z = 0" using assms higher_deriv_diff by auto qed (use assms in auto) lemma holomorphic_fun_eq_const_on_connected: assumes holf: "f holomorphic_on S" and "open S" and cons: "connected S" and der: "\n. 0 < n \ (deriv ^^ n) f z = 0" and "z \ S" "w \ S" shows "f w = f z" proof (rule holomorphic_fun_eq_0_on_connected [of "\w. f w - f z" S z, simplified]) show "(\w. f w - f z) holomorphic_on S" by (intro assms holomorphic_intros) show "\n. (deriv ^^ n) (\w. f w - f z) z = 0" by (subst higher_deriv_diff) (use assms in \auto intro: holomorphic_intros\) qed (use assms in auto) subsection\<^marker>\tag unimportant\ \Some basic lemmas about poles/singularities\ lemma pole_lemma: assumes holf: "f holomorphic_on S" and a: "a \ interior S" shows "(\z. if z = a then deriv f a else (f z - f a) / (z - a)) holomorphic_on S" (is "?F holomorphic_on S") proof - have F1: "?F field_differentiable (at u within S)" if "u \ S" "u \ a" for u proof - have fcd: "f field_differentiable at u within S" using holf holomorphic_on_def by (simp add: \u \ S\) have cd: "(\z. (f z - f a) / (z - a)) field_differentiable at u within S" by (rule fcd derivative_intros | simp add: that)+ have "0 < dist a u" using that dist_nz by blast then show ?thesis by (rule field_differentiable_transform_within [OF _ _ _ cd]) (auto simp: \u \ S\) qed have F2: "?F field_differentiable at a" if "0 < e" "ball a e \ S" for e proof - have holfb: "f holomorphic_on ball a e" by (rule holomorphic_on_subset [OF holf \ball a e \ S\]) have 2: "?F holomorphic_on ball a e - {a}" apply (simp add: holomorphic_on_def flip: field_differentiable_def) using mem_ball that apply (auto intro: F1 field_differentiable_within_subset) done have "isCont (\z. if z = a then deriv f a else (f z - f a) / (z - a)) x" if "dist a x < e" for x proof (cases "x=a") case True then have "f field_differentiable at a" using holfb \0 < e\ holomorphic_on_imp_differentiable_at by auto with True show ?thesis by (auto simp: continuous_at has_field_derivative_iff simp flip: DERIV_deriv_iff_field_differentiable elim: rev_iffD1 [OF _ LIM_equal]) next case False with 2 that show ?thesis by (force simp: holomorphic_on_open open_Diff field_differentiable_def [symmetric] field_differentiable_imp_continuous_at) qed then have 1: "continuous_on (ball a e) ?F" by (clarsimp simp: continuous_on_eq_continuous_at) have "?F holomorphic_on ball a e" by (auto intro: no_isolated_singularity [OF 1 2]) with that show ?thesis by (simp add: holomorphic_on_open field_differentiable_def [symmetric] field_differentiable_at_within) qed show ?thesis proof fix x assume "x \ S" show "?F field_differentiable at x within S" proof (cases "x=a") case True then show ?thesis using a by (auto simp: mem_interior intro: field_differentiable_at_within F2) next case False with F1 \x \ S\ show ?thesis by blast qed qed qed lemma pole_theorem: assumes holg: "g holomorphic_on S" and a: "a \ interior S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) holomorphic_on S" using pole_lemma [OF holg a] by (rule holomorphic_transform) (simp add: eq field_split_simps) lemma pole_lemma_open: assumes "f holomorphic_on S" "open S" shows "(\z. if z = a then deriv f a else (f z - f a)/(z - a)) holomorphic_on S" proof (cases "a \ S") case True with assms interior_eq pole_lemma show ?thesis by fastforce next case False with assms show ?thesis apply (simp add: holomorphic_on_def field_differentiable_def [symmetric], clarify) apply (rule field_differentiable_transform_within [where f = "\z. (f z - f a)/(z - a)" and d = 1]) apply (rule derivative_intros | force)+ done qed lemma pole_theorem_open: assumes holg: "g holomorphic_on S" and S: "open S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) holomorphic_on S" using pole_lemma_open [OF holg S] by (rule holomorphic_transform) (auto simp: eq divide_simps) lemma pole_theorem_0: assumes holg: "g holomorphic_on S" and a: "a \ interior S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f holomorphic_on S" using pole_theorem [OF holg a eq] by (rule holomorphic_transform) (auto simp: eq field_split_simps) lemma pole_theorem_open_0: assumes holg: "g holomorphic_on S" and S: "open S" and eq: "\z. z \ S - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f holomorphic_on S" using pole_theorem_open [OF holg S eq] by (rule holomorphic_transform) (auto simp: eq field_split_simps) lemma pole_theorem_analytic: assumes g: "g analytic_on S" and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" (is "?F analytic_on S") unfolding analytic_on_def proof fix x assume "x \ S" with g obtain e where "0 < e" and e: "g holomorphic_on ball x e" by (auto simp add: analytic_on_def) obtain d where "0 < d" and d: "\w. w \ ball x d - {a} \ g w = (w - a) * f w" using \x \ S\ eq by blast have "?F holomorphic_on ball x (min d e)" using d e \x \ S\ by (fastforce simp: holomorphic_on_subset subset_ball intro!: pole_theorem_open) then show "\e>0. ?F holomorphic_on ball x e" using \0 < d\ \0 < e\ not_le by fastforce qed lemma pole_theorem_analytic_0: assumes g: "g analytic_on S" and eq: "\z. z \ S \ \d. 0 < d \ (\w \ ball z d - {a}. g w = (w - a) * f w)" and [simp]: "f a = deriv g a" "g a = 0" shows "f analytic_on S" proof - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" by auto show ?thesis using pole_theorem_analytic [OF g eq] by simp qed lemma pole_theorem_analytic_open_superset: assumes g: "g analytic_on S" and "S \ T" "open T" and eq: "\z. z \ T - {a} \ g z = (z - a) * f z" shows "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" proof (rule pole_theorem_analytic [OF g]) fix z assume "z \ S" then obtain e where "0 < e" and e: "ball z e \ T" using assms openE by blast then show "\d>0. \w\ball z d - {a}. g w = (w - a) * f w" using eq by auto qed lemma pole_theorem_analytic_open_superset_0: assumes g: "g analytic_on S" "S \ T" "open T" "\z. z \ T - {a} \ g z = (z - a) * f z" and [simp]: "f a = deriv g a" "g a = 0" shows "f analytic_on S" proof - have [simp]: "(\z. if z = a then deriv g a else f z - g a / (z - a)) = f" by auto have "(\z. if z = a then deriv g a else f z - g a/(z - a)) analytic_on S" by (rule pole_theorem_analytic_open_superset [OF g]) then show ?thesis by simp qed subsection\General, homology form of Cauchy's theorem\ text\Proof is based on Dixon's, as presented in Lang's "Complex Analysis" book (page 147).\ lemma contour_integral_continuous_on_linepath_2D: assumes "open U" and cont_dw: "\w. w \ U \ F w contour_integrable_on (linepath a b)" and cond_uu: "continuous_on (U \ U) (\(x,y). F x y)" and abu: "closed_segment a b \ U" shows "continuous_on U (\w. contour_integral (linepath a b) (F w))" proof - have *: "\d>0. \x'\U. dist x' w < d \ dist (contour_integral (linepath a b) (F x')) (contour_integral (linepath a b) (F w)) \ \" if "w \ U" "0 < \" "a \ b" for w \ proof - obtain \ where "\>0" and \: "cball w \ \ U" using open_contains_cball \open U\ \w \ U\ by force let ?TZ = "cball w \ \ closed_segment a b" have "uniformly_continuous_on ?TZ (\(x,y). F x y)" proof (rule compact_uniformly_continuous) show "continuous_on ?TZ (\(x,y). F x y)" by (rule continuous_on_subset[OF cond_uu]) (use SigmaE \ abu in blast) show "compact ?TZ" by (simp add: compact_Times) qed then obtain \ where "\>0" and \: "\x x'. \x\?TZ; x'\?TZ; dist x' x < \\ \ dist ((\(x,y). F x y) x') ((\(x,y). F x y) x) < \/norm(b - a)" apply (rule uniformly_continuous_onE [where e = "\/norm(b - a)"]) using \0 < \\ \a \ b\ by auto have \: "\norm (w - x1) \ \; x2 \ closed_segment a b; norm (w - x1') \ \; x2' \ closed_segment a b; norm ((x1', x2') - (x1, x2)) < \\ \ norm (F x1' x2' - F x1 x2) \ \ / cmod (b - a)" for x1 x2 x1' x2' using \ [of "(x1,x2)" "(x1',x2')"] by (force simp: dist_norm) have le_ee: "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \" if "x' \ U" "cmod (x' - w) < \" "cmod (x' - w) < \" for x' proof - have "(\x. F x' x - F w x) contour_integrable_on linepath a b" by (simp add: \w \ U\ cont_dw contour_integrable_diff that) then have "cmod (contour_integral (linepath a b) (\x. F x' x - F w x)) \ \/norm(b - a) * norm(b - a)" apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_integral _ \]) using \0 < \\ \0 < \\ that apply (auto simp: norm_minus_commute) done also have "\ = \" using \a \ b\ by simp finally show ?thesis . qed show ?thesis apply (rule_tac x="min \ \" in exI) using \0 < \\ \0 < \\ apply (auto simp: dist_norm contour_integral_diff [OF cont_dw cont_dw, symmetric] \w \ U\ intro: le_ee) done qed show ?thesis proof (cases "a=b") case True then show ?thesis by simp next case False show ?thesis by (rule continuous_onI) (use False in \auto intro: *\) qed qed text\This version has \<^term>\polynomial_function \\ as an additional assumption.\ lemma Cauchy_integral_formula_global_weak: assumes "open U" and holf: "f holomorphic_on U" and z: "z \ U" and \: "polynomial_function \" and pasz: "path_image \ \ U - {z}" and loop: "pathfinish \ = pathstart \" and zero: "\w. w \ U \ winding_number \ w = 0" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - obtain \' where pf\': "polynomial_function \'" and \': "\x. (\ has_vector_derivative (\' x)) (at x)" using has_vector_derivative_polynomial_function [OF \] by blast then have "bounded(path_image \')" by (simp add: path_image_def compact_imp_bounded compact_continuous_image continuous_on_polymonial_function) then obtain B where "B>0" and B: "\x. x \ path_image \' \ norm x \ B" using bounded_pos by force define d where [abs_def]: "d z w = (if w = z then deriv f z else (f w - f z)/(w - z))" for z w define v where "v = {w. w \ path_image \ \ winding_number \ w = 0}" have "path \" "valid_path \" using \ by (auto simp: path_polynomial_function valid_path_polynomial_function) then have ov: "open v" by (simp add: v_def open_winding_number_levelsets loop) have uv_Un: "U \ v = UNIV" using pasz zero by (auto simp: v_def) have conf: "continuous_on U f" by (metis holf holomorphic_on_imp_continuous_on) have hol_d: "(d y) holomorphic_on U" if "y \ U" for y proof - have *: "(\c. if c = y then deriv f y else (f c - f y) / (c - y)) holomorphic_on U" by (simp add: holf pole_lemma_open \open U\) then have "isCont (\x. if x = y then deriv f y else (f x - f y) / (x - y)) y" using at_within_open field_differentiable_imp_continuous_at holomorphic_on_def that \open U\ by fastforce then have "continuous_on U (d y)" apply (simp add: d_def continuous_on_eq_continuous_at \open U\, clarify) using * holomorphic_on_def by (meson field_differentiable_within_open field_differentiable_imp_continuous_at \open U\) moreover have "d y holomorphic_on U - {y}" proof - have "\w. w \ U - {y} \ (\w. if w = y then deriv f y else (f w - f y) / (w - y)) field_differentiable at w" apply (rule_tac d="dist w y" and f = "\w. (f w - f y)/(w - y)" in field_differentiable_transform_within) apply (auto simp: dist_pos_lt dist_commute intro!: derivative_intros) using \open U\ holf holomorphic_on_imp_differentiable_at by blast then show ?thesis unfolding field_differentiable_def by (simp add: d_def holomorphic_on_open \open U\ open_delete) qed ultimately show ?thesis by (rule no_isolated_singularity) (auto simp: \open U\) qed have cint_fxy: "(\x. (f x - f y) / (x - y)) contour_integrable_on \" if "y \ path_image \" for y proof (rule contour_integrable_holomorphic_simple [where S = "U-{y}"]) show "(\x. (f x - f y) / (x - y)) holomorphic_on U - {y}" by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) show "path_image \ \ U - {y}" using pasz that by blast qed (auto simp: \open U\ open_delete \valid_path \\) define h where "h z = (if z \ U then contour_integral \ (d z) else contour_integral \ (\w. f w/(w - z)))" for z have U: "((d z) has_contour_integral h z) \" if "z \ U" for z proof - have "d z holomorphic_on U" by (simp add: hol_d that) with that show ?thesis apply (simp add: h_def) by (meson Diff_subset \open U\ \valid_path \\ contour_integrable_holomorphic_simple has_contour_integral_integral pasz subset_trans) qed have V: "((\w. f w / (w - z)) has_contour_integral h z) \" if z: "z \ v" for z proof - have 0: "0 = (f z) * 2 * of_real (2 * pi) * \ * winding_number \ z" using v_def z by auto then have "((\x. 1 / (x - z)) has_contour_integral 0) \" using z v_def has_contour_integral_winding_number [OF \valid_path \\] by fastforce then have "((\x. f z * (1 / (x - z))) has_contour_integral 0) \" using has_contour_integral_lmul by fastforce then have "((\x. f z / (x - z)) has_contour_integral 0) \" by (simp add: field_split_simps) moreover have "((\x. (f x - f z) / (x - z)) has_contour_integral contour_integral \ (d z)) \" using z apply (auto simp: v_def) apply (metis (no_types, lifting) contour_integrable_eq d_def has_contour_integral_eq has_contour_integral_integral cint_fxy) done ultimately have *: "((\x. f z / (x - z) + (f x - f z) / (x - z)) has_contour_integral (0 + contour_integral \ (d z))) \" by (rule has_contour_integral_add) have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (d z)) \" if "z \ U" using * by (auto simp: divide_simps has_contour_integral_eq) moreover have "((\w. f w / (w - z)) has_contour_integral contour_integral \ (\w. f w / (w - z))) \" if "z \ U" apply (rule has_contour_integral_integral [OF contour_integrable_holomorphic_simple [where S=U]]) using U pasz \valid_path \\ that apply (auto intro: holomorphic_on_imp_continuous_on hol_d) apply (rule continuous_intros conf holomorphic_intros holf assms | force)+ done ultimately show ?thesis using z by (simp add: h_def) qed have znot: "z \ path_image \" using pasz by blast obtain d0 where "d0>0" and d0: "\x y. x \ path_image \ \ y \ - U \ d0 \ dist x y" using separate_compact_closed [of "path_image \" "-U"] pasz \open U\ by (fastforce simp add: \path \\ compact_path_image) obtain dd where "0 < dd" and dd: "{y + k | y k. y \ path_image \ \ k \ ball 0 dd} \ U" apply (rule that [of "d0/2"]) using \0 < d0\ apply (auto simp: dist_norm dest: d0) done have "\x x'. \x \ path_image \; dist x x' * 2 < dd\ \ \y k. x' = y + k \ y \ path_image \ \ dist 0 k * 2 \ dd" apply (rule_tac x=x in exI) apply (rule_tac x="x'-x" in exI) apply (force simp: dist_norm) done then have 1: "path_image \ \ interior {y + k |y k. y \ path_image \ \ k \ cball 0 (dd / 2)}" apply (clarsimp simp add: mem_interior) using \0 < dd\ apply (rule_tac x="dd/2" in exI, auto) done obtain T where "compact T" and subt: "path_image \ \ interior T" and T: "T \ U" apply (rule that [OF _ 1]) apply (fastforce simp add: \valid_path \\ compact_valid_path_image intro!: compact_sums) apply (rule order_trans [OF _ dd]) using \0 < dd\ by fastforce obtain L where "L>0" and L: "\f B. \f holomorphic_on interior T; \z. z\interior T \ cmod (f z) \ B\ \ cmod (contour_integral \ f) \ L * B" using contour_integral_bound_exists [OF open_interior \valid_path \\ subt] by blast have "bounded(f ` T)" by (meson \compact T\ compact_continuous_image compact_imp_bounded conf continuous_on_subset T) then obtain D where "D>0" and D: "\x. x \ T \ norm (f x) \ D" by (auto simp: bounded_pos) obtain C where "C>0" and C: "\x. x \ T \ norm x \ C" using \compact T\ bounded_pos compact_imp_bounded by force have "dist (h y) 0 \ e" if "0 < e" and le: "D * L / e + C \ cmod y" for e y proof - have "D * L / e > 0" using \D>0\ \L>0\ \e>0\ by simp with le have ybig: "norm y > C" by force with C have "y \ T" by force then have ynot: "y \ path_image \" using subt interior_subset by blast have [simp]: "winding_number \ y = 0" apply (rule winding_number_zero_outside [of _ "cball 0 C"]) using ybig interior_subset subt apply (force simp: loop \path \\ dist_norm intro!: C)+ done have [simp]: "h y = contour_integral \ (\w. f w/(w - y))" by (rule contour_integral_unique [symmetric]) (simp add: v_def ynot V) have holint: "(\w. f w / (w - y)) holomorphic_on interior T" apply (rule holomorphic_on_divide) using holf holomorphic_on_subset interior_subset T apply blast apply (rule holomorphic_intros)+ using \y \ T\ interior_subset by auto have leD: "cmod (f z / (z - y)) \ D * (e / L / D)" if z: "z \ interior T" for z proof - have "D * L / e + cmod z \ cmod y" using le C [of z] z using interior_subset by force then have DL2: "D * L / e \ cmod (z - y)" using norm_triangle_ineq2 [of y z] by (simp add: norm_minus_commute) have "cmod (f z / (z - y)) = cmod (f z) * inverse (cmod (z - y))" by (simp add: norm_mult norm_inverse Fields.field_class.field_divide_inverse) also have "\ \ D * (e / L / D)" apply (rule mult_mono) using that D interior_subset apply blast using \L>0\ \e>0\ \D>0\ DL2 apply (auto simp: norm_divide field_split_simps algebra_simps) done finally show ?thesis . qed have "dist (h y) 0 = cmod (contour_integral \ (\w. f w / (w - y)))" by (simp add: dist_norm) also have "\ \ L * (D * (e / L / D))" by (rule L [OF holint leD]) also have "\ = e" using \L>0\ \0 < D\ by auto finally show ?thesis . qed then have "(h \ 0) at_infinity" by (meson Lim_at_infinityI) moreover have "h holomorphic_on UNIV" proof - have con_ff: "continuous (at (x,z)) (\(x,y). (f y - f x) / (y - x))" if "x \ U" "z \ U" "x \ z" for x z using that conf apply (simp add: split_def continuous_on_eq_continuous_at \open U\) apply (simp | rule continuous_intros continuous_within_compose2 [where g=f])+ done have con_fstsnd: "continuous_on UNIV (\x. (fst x - snd x) ::complex)" by (rule continuous_intros)+ have open_uu_Id: "open (U \ U - Id)" apply (rule open_Diff) apply (simp add: open_Times \open U\) using continuous_closed_preimage_constant [OF con_fstsnd closed_UNIV, of 0] apply (auto simp: Id_fstsnd_eq algebra_simps) done have con_derf: "continuous (at z) (deriv f)" if "z \ U" for z apply (rule continuous_on_interior [of U]) apply (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open U\) by (simp add: interior_open that \open U\) have tendsto_f': "((\(x,y). if y = x then deriv f (x) else (f (y) - f (x)) / (y - x)) \ deriv f x) (at (x, x) within U \ U)" if "x \ U" for x proof (rule Lim_withinI) fix e::real assume "0 < e" obtain k1 where "k1>0" and k1: "\x'. norm (x' - x) \ k1 \ norm (deriv f x' - deriv f x) < e" using \0 < e\ continuous_within_E [OF con_derf [OF \x \ U\]] by (metis UNIV_I dist_norm) obtain k2 where "k2>0" and k2: "ball x k2 \ U" by (blast intro: openE [OF \open U\] \x \ U\) have neq: "norm ((f z' - f x') / (z' - x') - deriv f x) \ e" if "z' \ x'" and less_k1: "norm (x'-x, z'-x) < k1" and less_k2: "norm (x'-x, z'-x) < k2" for x' z' proof - have cs_less: "w \ closed_segment x' z' \ cmod (w - x) \ norm (x'-x, z'-x)" for w apply (drule segment_furthest_le [where y=x]) by (metis (no_types) dist_commute dist_norm norm_fst_le norm_snd_le order_trans) have derf_le: "w \ closed_segment x' z' \ z' \ x' \ cmod (deriv f w - deriv f x) \ e" for w by (blast intro: cs_less less_k1 k1 [unfolded divide_const_simps dist_norm] less_imp_le le_less_trans) have f_has_der: "\x. x \ U \ (f has_field_derivative deriv f x) (at x within U)" by (metis DERIV_deriv_iff_field_differentiable at_within_open holf holomorphic_on_def \open U\) have "closed_segment x' z' \ U" by (rule order_trans [OF _ k2]) (simp add: cs_less le_less_trans [OF _ less_k2] dist_complex_def norm_minus_commute subset_iff) then have cint_derf: "(deriv f has_contour_integral f z' - f x') (linepath x' z')" using contour_integral_primitive [OF f_has_der valid_path_linepath] pasz by simp then have *: "((\x. deriv f x / (z' - x')) has_contour_integral (f z' - f x') / (z' - x')) (linepath x' z')" by (rule has_contour_integral_div) have "norm ((f z' - f x') / (z' - x') - deriv f x) \ e/norm(z' - x') * norm(z' - x')" apply (rule has_contour_integral_bound_linepath [OF has_contour_integral_diff [OF *]]) using has_contour_integral_div [where c = "z' - x'", OF has_contour_integral_const_linepath [of "deriv f x" z' x']] \e > 0\ \z' \ x'\ apply (auto simp: norm_divide divide_simps derf_le) done also have "\ \ e" using \0 < e\ by simp finally show ?thesis . qed show "\d>0. \xa\U \ U. 0 < dist xa (x, x) \ dist xa (x, x) < d \ dist (case xa of (x, y) \ if y = x then deriv f x else (f y - f x) / (y - x)) (deriv f x) \ e" apply (rule_tac x="min k1 k2" in exI) using \k1>0\ \k2>0\ \e>0\ apply (force simp: dist_norm neq intro: dual_order.strict_trans2 k1 less_imp_le norm_fst_le) done qed have con_pa_f: "continuous_on (path_image \) f" by (meson holf holomorphic_on_imp_continuous_on holomorphic_on_subset interior_subset subt T) have le_B: "\T. T \ {0..1} \ cmod (vector_derivative \ (at T)) \ B" apply (rule B) using \' using path_image_def vector_derivative_at by fastforce have f_has_cint: "\w. w \ v - path_image \ \ ((\u. f u / (u - w) ^ 1) has_contour_integral h w) \" by (simp add: V) have cond_uu: "continuous_on (U \ U) (\(x,y). d x y)" apply (simp add: continuous_on_eq_continuous_within d_def continuous_within tendsto_f') apply (simp add: tendsto_within_open_NO_MATCH open_Times \open U\, clarify) apply (rule Lim_transform_within_open [OF _ open_uu_Id, where f = "(\(x,y). (f y - f x) / (y - x))"]) using con_ff apply (auto simp: continuous_within) done have hol_dw: "(\z. d z w) holomorphic_on U" if "w \ U" for w proof - have "continuous_on U ((\(x,y). d x y) \ (\z. (w,z)))" by (rule continuous_on_compose continuous_intros continuous_on_subset [OF cond_uu] | force intro: that)+ then have *: "continuous_on U (\z. if w = z then deriv f z else (f w - f z) / (w - z))" by (rule rev_iffD1 [OF _ continuous_on_cong [OF refl]]) (simp add: d_def field_simps) have **: "\x. \x \ U; x \ w\ \ (\z. if w = z then deriv f z else (f w - f z) / (w - z)) field_differentiable at x" apply (rule_tac f = "\x. (f w - f x)/(w - x)" and d = "dist x w" in field_differentiable_transform_within) apply (rule \open U\ derivative_intros holomorphic_on_imp_differentiable_at [OF holf] | force simp: dist_commute)+ done show ?thesis unfolding d_def apply (rule no_isolated_singularity [OF * _ \open U\, where K = "{w}"]) apply (auto simp: field_differentiable_def [symmetric] holomorphic_on_open open_Diff \open U\ **) done qed { fix a b assume abu: "closed_segment a b \ U" then have "\w. w \ U \ (\z. d z w) contour_integrable_on (linepath a b)" by (metis hol_dw continuous_on_subset contour_integrable_continuous_linepath holomorphic_on_imp_continuous_on) then have cont_cint_d: "continuous_on U (\w. contour_integral (linepath a b) (\z. d z w))" apply (rule contour_integral_continuous_on_linepath_2D [OF \open U\ _ _ abu]) apply (auto intro: continuous_on_swap_args cond_uu) done have cont_cint_d\: "continuous_on {0..1} ((\w. contour_integral (linepath a b) (\z. d z w)) \ \)" proof (rule continuous_on_compose) show "continuous_on {0..1} \" using \path \\ path_def by blast show "continuous_on (\ ` {0..1}) (\w. contour_integral (linepath a b) (\z. d z w))" using pasz unfolding path_image_def by (auto intro!: continuous_on_subset [OF cont_cint_d]) qed have cint_cint: "(\w. contour_integral (linepath a b) (\z. d z w)) contour_integrable_on \" apply (simp add: contour_integrable_on) apply (rule integrable_continuous_real) apply (rule continuous_on_mult [OF cont_cint_d\ [unfolded o_def]]) using pf\' by (simp add: continuous_on_polymonial_function vector_derivative_at [OF \']) have "contour_integral (linepath a b) h = contour_integral (linepath a b) (\z. contour_integral \ (d z))" using abu by (force simp: h_def intro: contour_integral_eq) also have "\ = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" apply (rule contour_integral_swap) apply (rule continuous_on_subset [OF cond_uu]) using abu pasz \valid_path \\ apply (auto intro!: continuous_intros) by (metis \' continuous_on_eq path_def path_polynomial_function pf\' vector_derivative_at) finally have cint_h_eq: "contour_integral (linepath a b) h = contour_integral \ (\w. contour_integral (linepath a b) (\z. d z w))" . note cint_cint cint_h_eq } note cint_h = this have conthu: "continuous_on U h" proof (simp add: continuous_on_sequentially, clarify) fix a x assume x: "x \ U" and au: "\n. a n \ U" and ax: "a \ x" then have A1: "\\<^sub>F n in sequentially. d (a n) contour_integrable_on \" by (meson U contour_integrable_on_def eventuallyI) obtain dd where "dd>0" and dd: "cball x dd \ U" using open_contains_cball \open U\ x by force have A2: "uniform_limit (path_image \) (\n. d (a n)) (d x) sequentially" unfolding uniform_limit_iff dist_norm proof clarify fix ee::real assume "0 < ee" show "\\<^sub>F n in sequentially. \\\path_image \. cmod (d (a n) \ - d x \) < ee" proof - let ?ddpa = "{(w,z) |w z. w \ cball x dd \ z \ path_image \}" have "uniformly_continuous_on ?ddpa (\(x,y). d x y)" apply (rule compact_uniformly_continuous [OF continuous_on_subset[OF cond_uu]]) using dd pasz \valid_path \\ apply (auto simp: compact_Times compact_valid_path_image simp del: mem_cball) done then obtain kk where "kk>0" and kk: "\x x'. \x \ ?ddpa; x' \ ?ddpa; dist x' x < kk\ \ dist ((\(x,y). d x y) x') ((\(x,y). d x y) x) < ee" by (rule uniformly_continuous_onE [where e = ee]) (use \0 < ee\ in auto) have kk: "\norm (w - x) \ dd; z \ path_image \; norm ((w, z) - (x, z)) < kk\ \ norm (d w z - d x z) < ee" for w z using \dd>0\ kk [of "(x,z)" "(w,z)"] by (force simp: norm_minus_commute dist_norm) show ?thesis using ax unfolding lim_sequentially eventually_sequentially apply (drule_tac x="min dd kk" in spec) using \dd > 0\ \kk > 0\ apply (fastforce simp: kk dist_norm) done qed qed have "(\n. contour_integral \ (d (a n))) \ contour_integral \ (d x)" by (rule contour_integral_uniform_limit [OF A1 A2 le_B]) (auto simp: \valid_path \\) then have tendsto_hx: "(\n. contour_integral \ (d (a n))) \ h x" by (simp add: h_def x) then show "(h \ a) \ h x" by (simp add: h_def x au o_def) qed show ?thesis proof (simp add: holomorphic_on_open field_differentiable_def [symmetric], clarify) fix z0 consider "z0 \ v" | "z0 \ U" using uv_Un by blast then show "h field_differentiable at z0" proof cases assume "z0 \ v" then show ?thesis using Cauchy_next_derivative [OF con_pa_f le_B f_has_cint _ ov] V f_has_cint \valid_path \\ by (auto simp: field_differentiable_def v_def) next assume "z0 \ U" then obtain e where "e>0" and e: "ball z0 e \ U" by (blast intro: openE [OF \open U\]) have *: "contour_integral (linepath a b) h + contour_integral (linepath b c) h + contour_integral (linepath c a) h = 0" if abc_subset: "convex hull {a, b, c} \ ball z0 e" for a b c proof - have *: "\x1 x2 z. z \ U \ closed_segment x1 x2 \ U \ (\w. d w z) contour_integrable_on linepath x1 x2" using hol_dw holomorphic_on_imp_continuous_on \open U\ by (auto intro!: contour_integrable_holomorphic_simple) have abc: "closed_segment a b \ U" "closed_segment b c \ U" "closed_segment c a \ U" using that e segments_subset_convex_hull by fastforce+ have eq0: "\w. w \ U \ contour_integral (linepath a b +++ linepath b c +++ linepath c a) (\z. d z w) = 0" apply (rule contour_integral_unique [OF Cauchy_theorem_triangle]) apply (rule holomorphic_on_subset [OF hol_dw]) using e abc_subset by auto have "contour_integral \ (\x. contour_integral (linepath a b) (\z. d z x) + (contour_integral (linepath b c) (\z. d z x) + contour_integral (linepath c a) (\z. d z x))) = 0" apply (rule contour_integral_eq_0) using abc pasz U apply (subst contour_integral_join [symmetric], auto intro: eq0 *)+ done then show ?thesis by (simp add: cint_h abc contour_integrable_add contour_integral_add [symmetric] add_ac) qed show ?thesis using e \e > 0\ by (auto intro!: holomorphic_on_imp_differentiable_at [OF _ open_ball] analytic_imp_holomorphic Morera_triangle continuous_on_subset [OF conthu] *) qed qed qed ultimately have [simp]: "h z = 0" for z by (meson Liouville_weak) have "((\w. 1 / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z) \" by (rule has_contour_integral_winding_number [OF \valid_path \\ znot]) then have "((\w. f z * (1 / (w - z))) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" by (metis mult.commute has_contour_integral_lmul) then have 1: "((\w. f z / (w - z)) has_contour_integral complex_of_real (2 * pi) * \ * winding_number \ z * f z) \" by (simp add: field_split_simps) moreover have 2: "((\w. (f w - f z) / (w - z)) has_contour_integral 0) \" using U [OF z] pasz d_def by (force elim: has_contour_integral_eq [where g = "\w. (f w - f z)/(w - z)"]) show ?thesis using has_contour_integral_add [OF 1 2] by (simp add: diff_divide_distrib) qed theorem Cauchy_integral_formula_global: assumes S: "open S" and holf: "f holomorphic_on S" and z: "z \ S" and vpg: "valid_path \" and pasz: "path_image \ \ S - {z}" and loop: "pathfinish \ = pathstart \" and zero: "\w. w \ S \ winding_number \ w = 0" shows "((\w. f w / (w - z)) has_contour_integral (2*pi * \ * winding_number \ z * f z)) \" proof - have "path \" using vpg by (blast intro: valid_path_imp_path) have hols: "(\w. f w / (w - z)) holomorphic_on S - {z}" "(\w. 1 / (w - z)) holomorphic_on S - {z}" by (rule holomorphic_intros holomorphic_on_subset [OF holf] | force)+ then have cint_fw: "(\w. f w / (w - z)) contour_integrable_on \" by (meson contour_integrable_holomorphic_simple holomorphic_on_imp_continuous_on open_delete S vpg pasz) obtain d where "d>0" and d: "\g h. \valid_path g; valid_path h; \t\{0..1}. cmod (g t - \ t) < d \ cmod (h t - \ t) < d; pathstart h = pathstart g \ pathfinish h = pathfinish g\ \ path_image h \ S - {z} \ (\f. f holomorphic_on S - {z} \ contour_integral h f = contour_integral g f)" using contour_integral_nearby_ends [OF _ \path \\ pasz] S by (simp add: open_Diff) metis obtain p where polyp: "polynomial_function p" and ps: "pathstart p = pathstart \" and pf: "pathfinish p = pathfinish \" and led: "\t\{0..1}. cmod (p t - \ t) < d" using path_approx_polynomial_function [OF \path \\ \d > 0\] by blast then have ploop: "pathfinish p = pathstart p" using loop by auto have vpp: "valid_path p" using polyp valid_path_polynomial_function by blast have [simp]: "z \ path_image \" using pasz by blast have paps: "path_image p \ S - {z}" and cint_eq: "(\f. f holomorphic_on S - {z} \ contour_integral p f = contour_integral \ f)" using pf ps led d [OF vpg vpp] \d > 0\ by auto have wn_eq: "winding_number p z = winding_number \ z" using vpp paps by (simp add: subset_Diff_insert vpg valid_path_polynomial_function winding_number_valid_path cint_eq hols) have "winding_number p w = winding_number \ w" if "w \ S" for w proof - have hol: "(\v. 1 / (v - w)) holomorphic_on S - {z}" using that by (force intro: holomorphic_intros holomorphic_on_subset [OF holf]) have "w \ path_image p" "w \ path_image \" using paps pasz that by auto then show ?thesis using vpp vpg by (simp add: subset_Diff_insert valid_path_polynomial_function winding_number_valid_path cint_eq [OF hol]) qed then have wn0: "\w. w \ S \ winding_number p w = 0" by (simp add: zero) show ?thesis using Cauchy_integral_formula_global_weak [OF S holf z polyp paps ploop wn0] hols by (metis wn_eq cint_eq has_contour_integral_eqpath cint_fw cint_eq) qed theorem Cauchy_theorem_global: assumes S: "open S" and holf: "f holomorphic_on S" and vpg: "valid_path \" and loop: "pathfinish \ = pathstart \" and pas: "path_image \ \ S" and zero: "\w. w \ S \ winding_number \ w = 0" shows "(f has_contour_integral 0) \" proof - obtain z where "z \ S" and znot: "z \ path_image \" proof - have "compact (path_image \)" using compact_valid_path_image vpg by blast then have "path_image \ \ S" by (metis (no_types) compact_open path_image_nonempty S) with pas show ?thesis by (blast intro: that) qed then have pasz: "path_image \ \ S - {z}" using pas by blast have hol: "(\w. (w - z) * f w) holomorphic_on S" by (rule holomorphic_intros holf)+ show ?thesis using Cauchy_integral_formula_global [OF S hol \z \ S\ vpg pasz loop zero] by (auto simp: znot elim!: has_contour_integral_eq) qed corollary Cauchy_theorem_global_outside: assumes "open S" "f holomorphic_on S" "valid_path \" "pathfinish \ = pathstart \" "path_image \ \ S" "\w. w \ S \ w \ outside(path_image \)" shows "(f has_contour_integral 0) \" by (metis Cauchy_theorem_global assms winding_number_zero_in_outside valid_path_imp_path) lemma simply_connected_imp_winding_number_zero: assumes "simply_connected S" "path g" "path_image g \ S" "pathfinish g = pathstart g" "z \ S" shows "winding_number g z = 0" proof - have hom: "homotopic_loops S g (linepath (pathstart g) (pathstart g))" by (meson assms homotopic_paths_imp_homotopic_loops pathfinish_linepath simply_connected_eq_contractible_path) then have "homotopic_paths (- {z}) g (linepath (pathstart g) (pathstart g))" by (meson \z \ S\ homotopic_loops_imp_homotopic_paths_null homotopic_paths_subset subset_Compl_singleton) then have "winding_number g z = winding_number(linepath (pathstart g) (pathstart g)) z" by (rule winding_number_homotopic_paths) also have "\ = 0" using assms by (force intro: winding_number_trivial) finally show ?thesis . qed lemma Cauchy_theorem_simply_connected: assumes "open S" "simply_connected S" "f holomorphic_on S" "valid_path g" "path_image g \ S" "pathfinish g = pathstart g" shows "(f has_contour_integral 0) g" using assms apply (simp add: simply_connected_eq_contractible_path) apply (auto intro!: Cauchy_theorem_null_homotopic [where a = "pathstart g"] homotopic_paths_imp_homotopic_loops) using valid_path_imp_path by blast proposition\<^marker>\tag unimportant\ holomorphic_logarithm_exists: assumes A: "convex A" "open A" and f: "f holomorphic_on A" "\x. x \ A \ f x \ 0" and z0: "z0 \ A" obtains g where "g holomorphic_on A" and "\x. x \ A \ exp (g x) = f x" proof - note f' = holomorphic_derivI [OF f(1) A(2)] obtain g where g: "\x. x \ A \ (g has_field_derivative deriv f x / f x) (at x)" proof (rule holomorphic_convex_primitive' [OF A]) show "(\x. deriv f x / f x) holomorphic_on A" by (intro holomorphic_intros f A) qed (auto simp: A at_within_open[of _ A]) define h where "h = (\x. -g z0 + ln (f z0) + g x)" from g and A have g_holo: "g holomorphic_on A" by (auto simp: holomorphic_on_def at_within_open[of _ A] field_differentiable_def) hence h_holo: "h holomorphic_on A" by (auto simp: h_def intro!: holomorphic_intros) have "\c. \x\A. f x / exp (h x) - 1 = c" proof (rule has_field_derivative_zero_constant, goal_cases) case (2 x) note [simp] = at_within_open[OF _ \open A\] from 2 and z0 and f show ?case by (auto simp: h_def exp_diff field_simps intro!: derivative_eq_intros g f') qed fact+ then obtain c where c: "\x. x \ A \ f x / exp (h x) - 1 = c" by blast from c[OF z0] and z0 and f have "c = 0" by (simp add: h_def) with c have "\x. x \ A \ exp (h x) = f x" by simp from that[OF h_holo this] show ?thesis . qed end diff --git a/src/HOL/Analysis/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,1157 +1,1144 @@ (* 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\ theory Complex_Analysis_Basics imports Derivative "HOL-Library.Nonpos_Ints" begin (* TODO FIXME: A lot of the things in here have nothing to do with complex analysis *) 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 lambda_zero: "(\h::'a::mult_zero. 0) = (*) 0" by auto lemma lambda_one: "(\x::'a::monoid_mult. x) = (*) 1" by auto lemma uniformly_continuous_on_cmul_right [continuous_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "uniformly_continuous_on s f \ uniformly_continuous_on s (\x. f x * c)" using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . lemma uniformly_continuous_on_cmul_left[continuous_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_algebra" assumes "uniformly_continuous_on s f" shows "uniformly_continuous_on s (\x. c * f x)" by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right) -lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm" - by (rule continuous_norm [OF continuous_ident]) - lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm" by (intro continuous_on_id continuous_on_norm) -lemma DERIV_zero_unique: - assumes "convex S" - and d0: "\x. x\S \ (f has_field_derivative 0) (at x within S)" - and "a \ S" - and "x \ S" - shows "f x = f a" - by (rule has_derivative_zero_unique [OF assms(1) _ assms(4,3)]) - (metis d0 has_field_derivative_imp_has_derivative lambda_zero) - -lemma DERIV_zero_connected_unique: - assumes "connected S" - and "open S" - and d0: "\x. x\S \ DERIV f x :> 0" - and "a \ S" - and "x \ S" - shows "f x = f a" - by (rule has_derivative_zero_unique_connected [OF assms(2,1) _ assms(5,4)]) - (metis has_field_derivative_def lambda_zero d0) - -lemma DERIV_transform_within: - assumes "(f has_field_derivative f') (at a within S)" - and "0 < d" "a \ S" - and "\x. x\S \ dist x a < d \ f x = g x" - shows "(g has_field_derivative f') (at a within S)" - using assms unfolding has_field_derivative_def - by (blast intro: has_derivative_transform_within) - -lemma DERIV_transform_within_open: - assumes "DERIV f a :> f'" - and "open S" "a \ S" - and "\x. x\S \ f x = g x" - shows "DERIV g a :> f'" - using assms unfolding has_field_derivative_def -by (metis has_derivative_transform_within_open) - -lemma DERIV_transform_at: - assumes "DERIV f a :> f'" - and "0 < d" - and "\x. dist x a < d \ f x = g x" - shows "DERIV g a :> f'" - by (blast intro: assms DERIV_transform_within) - -(*generalising DERIV_isconst_all, which requires type real (using the ordering)*) -lemma DERIV_zero_UNIV_unique: - "(\x. DERIV f x :> 0) \ f x = f a" - by (metis DERIV_zero_unique UNIV_I convex_UNIV) - 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 DERIV_transform_within_open at_within_open) + (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\ DERIV_transform_within_open [OF df S] DERIV_const DERIV_unique) +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\<^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\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) 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 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 subsection\<^marker>\tag unimportant\\Bound theorem\ 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 subsection\<^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 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 subsection\<^marker>\tag unimportant\ \Polynomal function extremal theorem, from HOL Light\ lemma polyfun_extremal_lemma: (*COMPLEX_POLYFUN_EXTREMAL_LEMMA in HOL Light*) fixes c :: "nat \ 'a::real_normed_div_algebra" assumes "0 < e" shows "\M. \z. M \ norm(z) \ norm (\i\n. c(i) * z^i) \ e * norm(z) ^ (Suc n)" proof (induct n) case 0 with assms show ?case apply (rule_tac x="norm (c 0) / e" in exI) apply (auto simp: field_simps) done next case (Suc n) obtain M where M: "\z. M \ norm z \ norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using Suc assms by blast show ?case proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" and "1 + norm (c (Suc n)) / e \ norm z" then have z2: "e + norm (c (Suc n)) \ e * norm z" using assms by (simp add: field_simps) have "norm (\i\n. c i * z^i) \ e * norm z ^ Suc n" using M [OF z1] by simp then have "norm (\i\n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by simp then have "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" by (blast intro: norm_triangle_le elim: ) also have "... \ (e + norm (c (Suc n))) * norm z ^ Suc n" by (simp add: norm_power norm_mult algebra_simps) also have "... \ (e * norm z) * norm z ^ Suc n" by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) finally show "norm ((\i\n. c i * z^i) + c (Suc n) * z ^ Suc n) \ e * norm z ^ Suc (Suc n)" by simp qed qed lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) fixes c :: "nat \ 'a::real_normed_div_algebra" assumes k: "c k \ 0" "1\k" and kn: "k\n" shows "eventually (\z. norm (\i\n. c(i) * z^i) \ B) at_infinity" using kn proof (induction n) case 0 then show ?case using k by simp next case (Suc m) let ?even = ?case show ?even proof (cases "c (Suc m) = 0") case True then show ?even using Suc k by auto (metis antisym_conv less_eq_Suc_le not_le) next case False then obtain M where M: "\z. M \ norm z \ norm (\i\m. c i * z^i) \ norm (c (Suc m)) / 2 * norm z ^ Suc m" using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc by auto have "\b. \z. b \ norm z \ B \ norm (\i\Suc m. c i * z^i)" proof (rule exI [where x="max M (max 1 (\B\ / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) fix z::'a assume z1: "M \ norm z" "1 \ norm z" and "\B\ * 2 / norm (c (Suc m)) \ norm z" then have z2: "\B\ \ norm (c (Suc m)) * norm z / 2" using False by (simp add: field_simps) have nz: "norm z \ norm z ^ Suc m" by (metis \1 \ norm z\ One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) have *: "\y x. norm (c (Suc m)) * norm z / 2 \ norm y - norm x \ B \ norm (x + y)" by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) have "norm z * norm (c (Suc m)) + 2 * norm (\i\m. c i * z^i) \ norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" using M [of z] Suc z1 by auto also have "... \ 2 * (norm (c (Suc m)) * norm z ^ Suc m)" using nz by (simp add: mult_mono del: power_Suc) finally show "B \ norm ((\i\m. c i * z^i) + c (Suc m) * z ^ Suc m)" using Suc.IH apply (auto simp: eventually_at_infinity) apply (rule *) apply (simp add: field_simps norm_mult norm_power) done qed then show ?even by (simp add: eventually_at_infinity) qed qed end diff --git a/src/HOL/Analysis/Complex_Transcendental.thy b/src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy +++ b/src/HOL/Analysis/Complex_Transcendental.thy @@ -1,4056 +1,4056 @@ section \Complex Transcendental Functions\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ theory Complex_Transcendental imports Complex_Analysis_Basics Summation_Tests "HOL-Library.Periodic_Fun" begin subsection\Möbius transformations\ (* TODO: Figure out what to do with Möbius transformations *) definition\<^marker>\tag important\ "moebius a b c d = (\z. (a*z+b) / (c*z+d :: 'a :: field))" theorem moebius_inverse: assumes "a * d \ b * c" "c * z + d \ 0" shows "moebius d (-b) (-c) a (moebius a b c d z) = z" proof - from assms have "(-c) * moebius a b c d z + a \ 0" unfolding moebius_def by (simp add: field_simps) with assms show ?thesis unfolding moebius_def by (simp add: moebius_def divide_simps) (simp add: algebra_simps)? qed lemma moebius_inverse': assumes "a * d \ b * c" "c * z - a \ 0" shows "moebius a b c d (moebius d (-b) (-c) a z) = z" using assms moebius_inverse[of d a "-b" "-c" z] by (auto simp: algebra_simps) lemma cmod_add_real_less: assumes "Im z \ 0" "r\0" shows "cmod (z + r) < cmod z + \r\" proof (cases z) case (Complex x y) have "r * x / \r\ < sqrt (x*x + y*y)" apply (rule real_less_rsqrt) using assms apply (simp add: Complex power2_eq_square) using not_real_square_gt_zero by blast then show ?thesis using assms Complex apply (simp add: cmod_def) apply (rule power2_less_imp_less, auto) apply (simp add: power2_eq_square field_simps) done qed lemma cmod_diff_real_less: "Im z \ 0 \ x\0 \ cmod (z - x) < cmod z + \x\" using cmod_add_real_less [of z "-x"] by simp lemma cmod_square_less_1_plus: assumes "Im z = 0 \ \Re z\ < 1" shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)" proof (cases "Im z = 0 \ Re z = 0") case True with assms abs_square_less_1 show ?thesis by (force simp add: Re_power2 Im_power2 cmod_def) next case False with cmod_diff_real_less [of "1 - z\<^sup>2" "1"] show ?thesis by (simp add: norm_power Im_power2) qed subsection\<^marker>\tag unimportant\\The Exponential Function\ lemma norm_exp_i_times [simp]: "norm (exp(\ * of_real y)) = 1" by simp lemma norm_exp_imaginary: "norm(exp z) = 1 \ Re z = 0" by simp lemma field_differentiable_within_exp: "exp field_differentiable (at z within s)" using DERIV_exp field_differentiable_at_within field_differentiable_def by blast lemma continuous_within_exp: fixes z::"'a::{real_normed_field,banach}" shows "continuous (at z within s) exp" by (simp add: continuous_at_imp_continuous_within) lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s" by (simp add: field_differentiable_within_exp holomorphic_on_def) lemma holomorphic_on_exp' [holomorphic_intros]: "f holomorphic_on s \ (\x. exp (f x)) holomorphic_on s" using holomorphic_on_compose[OF _ holomorphic_on_exp] by (simp add: o_def) subsection\Euler and de Moivre formulas\ text\The sine series times \<^term>\i\\ lemma sin_i_eq: "(\n. (\ * sin_coeff n) * z^n) sums (\ * sin z)" proof - have "(\n. \ * sin_coeff n *\<^sub>R z^n) sums (\ * sin z)" using sin_converges sums_mult by blast then show ?thesis by (simp add: scaleR_conv_of_real field_simps) qed theorem exp_Euler: "exp(\ * z) = cos(z) + \ * sin(z)" proof - have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) = (\n. (\ * z) ^ n /\<^sub>R (fact n))" proof fix n show "(cos_coeff n + \ * sin_coeff n) * z^n = (\ * z) ^ n /\<^sub>R (fact n)" by (auto simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE) qed also have "... sums (exp (\ * z))" by (rule exp_converges) finally have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (exp (\ * z))" . moreover have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (cos z + \ * sin z)" using sums_add [OF cos_converges [of z] sin_i_eq [of z]] by (simp add: field_simps scaleR_conv_of_real) ultimately show ?thesis using sums_unique2 by blast qed corollary\<^marker>\tag unimportant\ exp_minus_Euler: "exp(-(\ * z)) = cos(z) - \ * sin(z)" using exp_Euler [of "-z"] by simp lemma sin_exp_eq: "sin z = (exp(\ * z) - exp(-(\ * z))) / (2*\)" by (simp add: exp_Euler exp_minus_Euler) lemma sin_exp_eq': "sin z = \ * (exp(-(\ * z)) - exp(\ * z)) / 2" by (simp add: exp_Euler exp_minus_Euler) lemma cos_exp_eq: "cos z = (exp(\ * z) + exp(-(\ * z))) / 2" by (simp add: exp_Euler exp_minus_Euler) theorem Euler: "exp(z) = of_real(exp(Re z)) * (of_real(cos(Im z)) + \ * of_real(sin(Im z)))" by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real Complex_eq) lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Re_divide Im_exp) lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Im_divide Re_exp) lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: cos_exp_eq field_simps Re_divide Re_exp) lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2" by (simp add: cos_exp_eq field_simps Im_divide Im_exp) lemma Re_sin_pos: "0 < Re z \ Re z < pi \ Re (sin z) > 0" by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero) lemma Im_sin_nonneg: "Re z = 0 \ 0 \ Im z \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) lemma Im_sin_nonneg2: "Re z = pi \ Im z \ 0 \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) subsection\<^marker>\tag unimportant\\Relationships between real and complex trigonometric and hyperbolic functions\ lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x" by (simp add: sin_of_real) lemma real_cos_eq [simp]: "Re(cos(of_real x)) = cos x" by (simp add: cos_of_real) lemma DeMoivre: "(cos z + \ * sin z) ^ n = cos(n * z) + \ * sin(n * z)" by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute) lemma exp_cnj: "cnj (exp z) = exp (cnj z)" proof - have "(\n. cnj (z ^ n /\<^sub>R (fact n))) = (\n. (cnj z)^n /\<^sub>R (fact n))" by auto also have "... sums (exp (cnj z))" by (rule exp_converges) finally have "(\n. cnj (z ^ n /\<^sub>R (fact n))) sums (exp (cnj z))" . moreover have "(\n. cnj (z ^ n /\<^sub>R (fact n))) sums (cnj (exp z))" by (metis exp_converges sums_cnj) ultimately show ?thesis using sums_unique2 by blast qed lemma cnj_sin: "cnj(sin z) = sin(cnj z)" by (simp add: sin_exp_eq exp_cnj field_simps) lemma cnj_cos: "cnj(cos z) = cos(cnj z)" by (simp add: cos_exp_eq exp_cnj field_simps) lemma field_differentiable_at_sin: "sin field_differentiable at z" using DERIV_sin field_differentiable_def by blast lemma field_differentiable_within_sin: "sin field_differentiable (at z within S)" by (simp add: field_differentiable_at_sin field_differentiable_at_within) lemma field_differentiable_at_cos: "cos field_differentiable at z" using DERIV_cos field_differentiable_def by blast lemma field_differentiable_within_cos: "cos field_differentiable (at z within S)" by (simp add: field_differentiable_at_cos field_differentiable_at_within) lemma holomorphic_on_sin: "sin holomorphic_on S" by (simp add: field_differentiable_within_sin holomorphic_on_def) lemma holomorphic_on_cos: "cos holomorphic_on S" by (simp add: field_differentiable_within_cos holomorphic_on_def) lemma holomorphic_on_sin' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. sin (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def) lemma holomorphic_on_cos' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. cos (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def) subsection\<^marker>\tag unimportant\\More on the Polar Representation of Complex Numbers\ lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)" by (simp add: Complex_eq exp_add exp_Euler exp_of_real sin_of_real cos_of_real) lemma exp_eq_1: "exp z = 1 \ Re(z) = 0 \ (\n::int. Im(z) = of_int (2 * n) * pi)" (is "?lhs = ?rhs") proof assume "exp z = 1" then have "Re z = 0" by (metis exp_eq_one_iff norm_exp_eq_Re norm_one) with \?lhs\ show ?rhs by (metis Re_exp complex_Re_of_int cos_one_2pi_int exp_zero mult.commute mult_numeral_1 numeral_One of_int_mult of_int_numeral) next assume ?rhs then show ?lhs using Im_exp Re_exp complex_eq_iff by (simp add: cos_one_2pi_int cos_one_sin_zero mult.commute) qed lemma exp_eq: "exp w = exp z \ (\n::int. w = z + (of_int (2 * n) * pi) * \)" (is "?lhs = ?rhs") proof - have "exp w = exp z \ exp (w-z) = 1" by (simp add: exp_diff) also have "... \ (Re w = Re z \ (\n::int. Im w - Im z = of_int (2 * n) * pi))" by (simp add: exp_eq_1) also have "... \ ?rhs" by (auto simp: algebra_simps intro!: complex_eqI) finally show ?thesis . qed lemma exp_complex_eqI: "\Im w - Im z\ < 2*pi \ exp w = exp z \ w = z" by (auto simp: exp_eq abs_mult) lemma exp_integer_2pi: assumes "n \ \" shows "exp((2 * n * pi) * \) = 1" proof - have "exp((2 * n * pi) * \) = exp 0" using assms unfolding Ints_def exp_eq by auto also have "... = 1" by simp finally show ?thesis . qed lemma exp_plus_2pin [simp]: "exp (z + \ * (of_int n * (of_real pi * 2))) = exp z" by (simp add: exp_eq) lemma exp_integer_2pi_plus1: assumes "n \ \" shows "exp(((2 * n + 1) * pi) * \) = - 1" proof - from assms obtain n' where [simp]: "n = of_int n'" by (auto simp: Ints_def) have "exp(((2 * n + 1) * pi) * \) = exp (pi * \)" using assms by (subst exp_eq) (auto intro!: exI[of _ n'] simp: algebra_simps) also have "... = - 1" by simp finally show ?thesis . qed lemma inj_on_exp_pi: fixes z::complex shows "inj_on exp (ball z pi)" proof (clarsimp simp: inj_on_def exp_eq) fix y n assume "dist z (y + 2 * of_int n * of_real pi * \) < pi" "dist z y < pi" then have "dist y (y + 2 * of_int n * of_real pi * \) < pi+pi" using dist_commute_lessI dist_triangle_less_add by blast then have "norm (2 * of_int n * of_real pi * \) < 2*pi" by (simp add: dist_norm) then show "n = 0" by (auto simp: norm_mult) qed lemma cmod_add_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) + r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\1 - \2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs") proof - have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)" by (rule complex_norm_square) also have "\ = (?z1 * cnj ?z1 + ?z2 * cnj ?z2) + (?z1 * cnj ?z2 + cnj ?z1 * ?z2)" by (simp add: algebra_simps) also have "\ = (norm ?z1)\<^sup>2 + (norm ?z2)\<^sup>2 + 2 * Re (?z1 * cnj ?z2)" unfolding complex_norm_square [symmetric] cnj_add_mult_eq_Re by simp also have "\ = ?rhs" by (simp add: norm_mult) (simp add: exp_Euler complex_is_Real_iff [THEN iffD1] cos_diff algebra_simps) finally show ?thesis using of_real_eq_iff by blast qed lemma cmod_diff_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) - r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\1 - \2)" (is "(cmod (?z1 - ?z2))\<^sup>2 = ?rhs") proof - have "exp (\ * (\2 + pi)) = - exp (\ * \2)" by (simp add: exp_Euler cos_plus_pi sin_plus_pi) then have "(cmod (?z1 - ?z2))\<^sup>2 = cmod (?z1 + r2 * exp (\ * (\2 + pi))) ^2" by simp also have "\ = r1\<^sup>2 + r2\<^sup>2 + 2*r1*r2*cos (\1 - (\2 + pi))" using assms cmod_add_squared by blast also have "\ = ?rhs" by (simp add: add.commute diff_add_eq_diff_diff_swap) finally show ?thesis . qed lemma polar_convergence: fixes R::real assumes "\j. r j > 0" "R > 0" shows "((\j. r j * exp (\ * \ j)) \ (R * exp (\ * \))) \ (r \ R) \ (\k. (\j. \ j - of_int (k j) * (2 * pi)) \ \)" (is "(?z \ ?Z) = ?rhs") proof assume L: "?z \ ?Z" have rR: "r \ R" using tendsto_norm [OF L] assms by (auto simp: norm_mult abs_of_pos) moreover obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - have "cos (\ j - \) = ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)" for j apply (subst cmod_diff_squared) using assms by (auto simp: field_split_simps less_le) moreover have "(\j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \ ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))" by (intro L rR tendsto_intros) (use \R > 0\ in force) moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1" using \R > 0\ by (simp add: power2_eq_square field_split_simps) ultimately have "(\j. cos (\ j - \)) \ 1" by auto then show ?thesis using that cos_diff_limit_1 by blast qed ultimately show ?rhs by metis next assume R: ?rhs show "?z \ ?Z" proof (rule tendsto_mult) show "(\x. complex_of_real (r x)) \ of_real R" using R by (auto simp: tendsto_of_real_iff) obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" using R by metis then have "(\j. complex_of_real (\ j - of_int (k j) * (2 * pi))) \ of_real \" using tendsto_of_real_iff by force then have "(\j. exp (\ * of_real (\ j - of_int (k j) * (2 * pi)))) \ exp (\ * \)" using tendsto_mult [OF tendsto_const] isCont_exp isCont_tendsto_compose by blast moreover have "exp (\ * of_real (\ j - of_int (k j) * (2 * pi))) = exp (\ * \ j)" for j unfolding exp_eq by (rule_tac x="- k j" in exI) (auto simp: algebra_simps) ultimately show "(\j. exp (\ * \ j)) \ exp (\ * \)" by auto qed qed lemma sin_cos_eq_iff: "sin y = sin x \ cos y = cos x \ (\n::int. y = x + 2 * pi * n)" proof - { assume "sin y = sin x" "cos y = cos x" then have "cos (y-x) = 1" using cos_add [of y "-x"] by simp then have "\n::int. y-x = 2 * pi * n" using cos_one_2pi_int by auto } then show ?thesis apply (auto simp: sin_add cos_add) apply (metis add.commute diff_add_cancel) done qed lemma exp_i_ne_1: assumes "0 < x" "x < 2*pi" shows "exp(\ * of_real x) \ 1" proof assume "exp (\ * of_real x) = 1" then have "exp (\ * of_real x) = exp 0" by simp then obtain n where "\ * of_real x = (of_int (2 * n) * pi) * \" by (simp only: Ints_def exp_eq) auto then have "of_real x = (of_int (2 * n) * pi)" by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real) then have "x = (of_int (2 * n) * pi)" by simp then show False using assms by (cases n) (auto simp: zero_less_mult_iff mult_less_0_iff) qed lemma sin_eq_0: fixes z::complex shows "sin z = 0 \ (\n::int. z = of_real(n * pi))" by (simp add: sin_exp_eq exp_eq) lemma cos_eq_0: fixes z::complex shows "cos z = 0 \ (\n::int. z = of_real(n * pi) + of_real pi/2)" using sin_eq_0 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps) lemma cos_eq_1: fixes z::complex shows "cos z = 1 \ (\n::int. z = of_real(2 * n * pi))" proof - have "cos z = cos (2*(z/2))" by simp also have "... = 1 - 2 * sin (z/2) ^ 2" by (simp only: cos_double_sin) finally have [simp]: "cos z = 1 \ sin (z/2) = 0" by simp show ?thesis by (auto simp: sin_eq_0) qed lemma csin_eq_1: fixes z::complex shows "sin z = 1 \ (\n::int. z = of_real(2 * n * pi) + of_real pi/2)" using cos_eq_1 [of "z - of_real pi/2"] by (simp add: cos_diff algebra_simps) lemma csin_eq_minus1: fixes z::complex shows "sin z = -1 \ (\n::int. z = of_real(2 * n * pi) + 3/2*pi)" (is "_ = ?rhs") proof - have "sin z = -1 \ sin (-z) = 1" by (simp add: equation_minus_iff) also have "... \ (\n::int. -z = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\n::int. z = - of_real(2 * n * pi) - of_real pi/2)" apply (rule iff_exI) by (metis (no_types) is_num_normalize(8) minus_minus of_real_def real_vector.scale_minus_left uminus_add_conv_diff) also have "... = ?rhs" apply safe apply (rule_tac [2] x="-(x+1)" in exI) apply (rule_tac x="-(x+1)" in exI) apply (simp_all add: algebra_simps) done finally show ?thesis . qed lemma ccos_eq_minus1: fixes z::complex shows "cos z = -1 \ (\n::int. z = of_real(2 * n * pi) + pi)" using csin_eq_1 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps equation_minus_iff) lemma sin_eq_1: "sin x = 1 \ (\n::int. x = (2 * n + 1 / 2) * pi)" (is "_ = ?rhs") proof - have "sin x = 1 \ sin (complex_of_real x) = 1" by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\n::int. x = of_real(2 * n * pi) + of_real pi/2)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma sin_eq_minus1: "sin x = -1 \ (\n::int. x = (2*n + 3/2) * pi)" (is "_ = ?rhs") proof - have "sin x = -1 \ sin (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)" by (simp only: csin_eq_minus1) also have "... \ (\n::int. x = of_real(2 * n * pi) + 3/2*pi)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma cos_eq_minus1: "cos x = -1 \ (\n::int. x = (2*n + 1) * pi)" (is "_ = ?rhs") proof - have "cos x = -1 \ cos (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + pi)" by (simp only: ccos_eq_minus1) also have "... \ (\n::int. x = of_real(2 * n * pi) + pi)" by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma dist_exp_i_1: "norm(exp(\ * of_real t) - 1) = 2 * \sin(t / 2)\" apply (simp add: exp_Euler cmod_def power2_diff sin_of_real cos_of_real algebra_simps) using cos_double_sin [of "t/2"] apply (simp add: real_sqrt_mult) done lemma complex_sin_eq: fixes w :: complex shows "sin w = sin z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real((2*n + 1)*pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "sin w - sin z = 0" by (auto simp: algebra_simps) then have "sin ((w - z) / 2)*cos ((w + z) / 2) = 0" by (auto simp: sin_diff_sin) then consider "sin ((w - z) / 2) = 0" | "cos ((w + z) / 2) = 0" using mult_eq_0_iff by blast then show ?rhs proof cases case 1 then show ?thesis by (simp add: sin_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) next case 2 then show ?thesis by (simp add: cos_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) qed next assume ?rhs then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \ w = -z + of_real ((2* of_int n + 1)*pi)" using Ints_cases by blast then show ?lhs using Periodic_Fun.sin.plus_of_int [of z n] apply (auto simp: algebra_simps) by (metis (no_types, hide_lams) add_diff_cancel_left add_diff_cancel_left' add_minus_cancel mult.commute sin.plus_of_int sin_minus sin_plus_pi) qed lemma complex_cos_eq: fixes w :: complex shows "cos w = cos z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real(2*n*pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "cos w - cos z = 0" by (auto simp: algebra_simps) then have "sin ((w + z) / 2) * sin ((z - w) / 2) = 0" by (auto simp: cos_diff_cos) then consider "sin ((w + z) / 2) = 0" | "sin ((z - w) / 2) = 0" using mult_eq_0_iff by blast then show ?rhs proof cases case 1 then show ?thesis apply (simp add: sin_eq_0 algebra_simps) by (metis Ints_of_int of_real_of_int_eq) next case 2 then show ?thesis apply (clarsimp simp: sin_eq_0 algebra_simps) by (metis Ints_of_int add_minus_cancel distrib_right mult_of_int_commute mult_zero_right of_int_0 of_int_add of_real_of_int_eq) qed next assume ?rhs then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \ w = -z + of_real(2*n*pi)" using Ints_cases by (metis of_int_mult of_int_numeral) then show ?lhs using Periodic_Fun.cos.plus_of_int [of z n] apply (simp add: algebra_simps) by (metis cos.plus_of_int cos_minus minus_add_cancel mult.commute) qed lemma sin_eq: "sin x = sin y \ (\n \ \. x = y + 2*n*pi \ x = -y + (2*n + 1)*pi)" using complex_sin_eq [of x y] by (simp only: sin_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) lemma cos_eq: "cos x = cos y \ (\n \ \. x = y + 2*n*pi \ x = -y + 2*n*pi)" using complex_cos_eq [of x y] by (simp only: cos_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) lemma sinh_complex: fixes z :: complex shows "(exp z - inverse (exp z)) / 2 = -\ * sin(\ * z)" by (simp add: sin_exp_eq field_split_simps exp_minus) lemma sin_i_times: fixes z :: complex shows "sin(\ * z) = \ * ((exp z - inverse (exp z)) / 2)" using sinh_complex by auto lemma sinh_real: fixes x :: real shows "of_real((exp x - inverse (exp x)) / 2) = -\ * sin(\ * of_real x)" by (simp add: exp_of_real sin_i_times) lemma cosh_complex: fixes z :: complex shows "(exp z + inverse (exp z)) / 2 = cos(\ * z)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemma cosh_real: fixes x :: real shows "of_real((exp x + inverse (exp x)) / 2) = cos(\ * of_real x)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemmas cos_i_times = cosh_complex [symmetric] lemma norm_cos_squared: "norm(cos z) ^ 2 = cos(Re z) ^ 2 + (exp(Im z) - inverse(exp(Im z))) ^ 2 / 4" apply (cases z) apply (simp add: cos_add cmod_power2 cos_of_real sin_of_real Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib sin_squared_eq) apply (simp add: power2_eq_square algebra_simps field_split_simps) done lemma norm_sin_squared: "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4" apply (cases z) apply (simp add: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib cos_squared_eq) apply (simp add: power2_eq_square algebra_simps field_split_simps) done lemma exp_uminus_Im: "exp (- Im z) \ exp (cmod z)" using abs_Im_le_cmod linear order_trans by fastforce lemma norm_cos_le: fixes z::complex shows "norm(cos z) \ exp(norm z)" proof - have "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto with exp_uminus_Im show ?thesis apply (simp add: cos_exp_eq norm_divide) apply (rule order_trans [OF norm_triangle_ineq], simp) apply (metis add_mono exp_le_cancel_iff mult_2_right) done qed lemma norm_cos_plus1_le: fixes z::complex shows "norm(1 + cos z) \ 2 * exp(norm z)" proof - have mono: "\u w z::real. (1 \ w | 1 \ z) \ (w \ u & z \ u) \ 2 + w + z \ 4 * u" by arith have *: "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto have triangle3: "\x y z. norm(x + y + z) \ norm(x) + norm(y) + norm(z)" by (simp add: norm_add_rule_thm) have "norm(1 + cos z) = cmod (1 + (exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: cos_exp_eq) also have "... = cmod ((2 + exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: field_simps) also have "... = cmod (2 + exp (\ * z) + exp (- (\ * z))) / 2" by (simp add: norm_divide) finally show ?thesis by (metis exp_eq_one_iff exp_le_cancel_iff mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono) qed lemma sinh_conv_sin: "sinh z = -\ * sin (\*z)" by (simp add: sinh_field_def sin_i_times exp_minus) lemma cosh_conv_cos: "cosh z = cos (\*z)" by (simp add: cosh_field_def cos_i_times exp_minus) lemma tanh_conv_tan: "tanh z = -\ * tan (\*z)" by (simp add: tanh_def sinh_conv_sin cosh_conv_cos tan_def) lemma sin_conv_sinh: "sin z = -\ * sinh (\*z)" by (simp add: sinh_conv_sin) lemma cos_conv_cosh: "cos z = cosh (\*z)" by (simp add: cosh_conv_cos) lemma tan_conv_tanh: "tan z = -\ * tanh (\*z)" by (simp add: tan_def sin_conv_sinh cos_conv_cosh tanh_def) lemma sinh_complex_eq_iff: "sinh (z :: complex) = sinh w \ (\n\\. z = w - 2 * \ * of_real n * of_real pi \ z = -(2 * complex_of_real n + 1) * \ * complex_of_real pi - w)" (is "_ = ?rhs") proof - have "sinh z = sinh w \ sin (\ * z) = sin (\ * w)" by (simp add: sinh_conv_sin) also have "\ \ ?rhs" by (subst complex_sin_eq) (force simp: field_simps complex_eq_iff) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Taylor series for complex exponential, sine and cosine\ declare power_Suc [simp del] lemma Taylor_exp_field: fixes z::"'a::{banach,real_normed_field}" shows "norm (exp z - (\i\n. z ^ i / fact i)) \ exp (norm z) * (norm z ^ Suc n) / fact n" proof (rule field_Taylor[of _ n "\k. exp" "exp (norm z)" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume x: "x \ closed_segment 0 z" have "norm (exp x) \ exp (norm x)" by (rule norm_exp) also have "norm x \ norm z" using x by (auto simp: closed_segment_def intro!: mult_left_le_one_le) finally show "norm (exp x) \ exp (norm z)" by simp qed auto lemma Taylor_exp: "norm(exp z - (\k\n. z ^ k / (fact k))) \ exp\Re z\ * (norm z) ^ (Suc n) / (fact n)" proof (rule complex_Taylor [of _ n "\k. exp" "exp\Re z\" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume "x \ closed_segment 0 z" then show "Re x \ \Re z\" apply (clarsimp simp: closed_segment_def scaleR_conv_of_real) by (meson abs_ge_self abs_ge_zero linear mult_left_le_one_le mult_nonneg_nonpos order_trans) qed auto lemma assumes "0 \ u" "u \ 1" shows cmod_sin_le_exp: "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" and cmod_cos_le_exp: "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" proof - have mono: "\u w z::real. w \ u \ z \ u \ (w + z)/2 \ u" by simp have *: "(cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2 \ exp \Im z\" proof (rule mono) show "cmod (exp (\ * (u * z))) \ exp \Im z\" apply (simp add: abs_if mult_left_le_one_le assms not_less) by (meson assms(1) mult_nonneg_nonneg neg_le_0_iff_le order_trans) show "cmod (exp (- (\ * (u * z)))) \ exp \Im z\" apply (simp add: abs_if mult_left_le_one_le assms) by (meson \0 \ u\ less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans) qed have "cmod (sin (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) - exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq4) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" . have "cmod (cos (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) + exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" . qed lemma Taylor_sin: "norm(sin z - (\k\n. complex_of_real (sin_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ (Suc n) / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (sin z - (\i\n. (-1) ^ (i div 2) * (if even i then sin 0 else cos 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(k div 2) * (if even k then sin x else cos x)" "exp\Im z\" 0 z, simplified]) fix k x show "((\x. (- 1) ^ (k div 2) * (if even k then sin x else cos x)) has_field_derivative (- 1) ^ (Suc k div 2) * (if odd k then sin x else cos x)) (at x within closed_segment 0 z)" apply (auto simp: power_Suc) apply (intro derivative_eq_intros | simp)+ done next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ (Suc n div 2) * (if odd n then sin x else cos x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (sin_coeff k) * z ^ k = (-1)^(k div 2) * (if even k then sin 0 else cos 0) * z^k / of_nat (fact k)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis apply (rule order_trans [OF _ *]) apply (simp add: **) done qed lemma Taylor_cos: "norm(cos z - (\k\n. complex_of_real (cos_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ Suc n / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (cos z - (\i\n. (-1) ^ (Suc i div 2) * (if even i then cos 0 else sin 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" "exp\Im z\" 0 z, simplified]) fix k x assume "x \ closed_segment 0 z" "k \ n" show "((\x. (- 1) ^ (Suc k div 2) * (if even k then cos x else sin x)) has_field_derivative (- 1) ^ Suc (k div 2) * (if odd k then cos x else sin x)) (at x within closed_segment 0 z)" apply (auto simp: power_Suc) apply (intro derivative_eq_intros | simp)+ done next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ Suc (n div 2) * (if odd n then cos x else sin x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (cos_coeff k) * z ^ k = (-1)^(Suc k div 2) * (if even k then cos 0 else sin 0) * z^k / of_nat (fact k)" by (auto simp: cos_coeff_def elim!: evenE) show ?thesis apply (rule order_trans [OF _ *]) apply (simp add: **) done qed declare power_Suc [simp] text\32-bit Approximation to e\ lemma e_approx_32: "\exp(1) - 5837465777 / 2147483648\ \ (inverse(2 ^ 32)::real)" using Taylor_exp [of 1 14] exp_le apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral) apply (simp only: pos_le_divide_eq [symmetric]) done lemma e_less_272: "exp 1 < (272/100::real)" using e_approx_32 by (simp add: abs_if split: if_split_asm) lemma ln_272_gt_1: "ln (272/100) > (1::real)" by (metis e_less_272 exp_less_cancel_iff exp_ln_iff less_trans ln_exp) text\Apparently redundant. But many arguments involve integers.\ lemma ln3_gt_1: "ln 3 > (1::real)" by (simp add: less_trans [OF ln_272_gt_1]) subsection\The argument of a complex number (HOL Light version)\ definition\<^marker>\tag important\ is_Arg :: "[complex,real] \ bool" where "is_Arg z r \ z = of_real(norm z) * exp(\ * of_real r)" definition\<^marker>\tag important\ Arg2pi :: "complex \ real" where "Arg2pi z \ if z = 0 then 0 else THE t. 0 \ t \ t < 2*pi \ is_Arg z t" lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \ is_Arg z r" by (simp add: algebra_simps is_Arg_def) lemma is_Arg_eqI: assumes r: "is_Arg z r" and s: "is_Arg z s" and rs: "abs (r-s) < 2*pi" and "z \ 0" shows "r = s" proof - have zr: "z = (cmod z) * exp (\ * r)" and zs: "z = (cmod z) * exp (\ * s)" using r s by (auto simp: is_Arg_def) with \z \ 0\ have eq: "exp (\ * r) = exp (\ * s)" by (metis mult_eq_0_iff mult_left_cancel) have "\ * r = \ * s" by (rule exp_complex_eqI) (use rs in \auto simp: eq exp_complex_eqI\) then show ?thesis by simp qed text\This function returns the angle of a complex number from its representation in polar coordinates. Due to periodicity, its range is arbitrary. \<^term>\Arg2pi\ follows HOL Light in adopting the interval \[0,2\)\. But we have the same periodicity issue with logarithms, and it is usual to adopt the same interval for the complex logarithm and argument functions. Further on down, we shall define both functions for the interval \(-\,\]\. The present version is provided for compatibility.\ lemma Arg2pi_0 [simp]: "Arg2pi(0) = 0" by (simp add: Arg2pi_def) lemma Arg2pi_unique_lemma: assumes z: "is_Arg z t" and z': "is_Arg z t'" and t: "0 \ t" "t < 2*pi" and t': "0 \ t'" "t' < 2*pi" and nz: "z \ 0" shows "t' = t" proof - have [dest]: "\x y z::real. x\0 \ x+y < z \ y * of_real t') = of_real (cmod z) * exp (\ * of_real t)" by (metis z z' is_Arg_def) then have "exp (\ * of_real t') = exp (\ * of_real t)" by (metis nz mult_left_cancel mult_zero_left z is_Arg_def) then have "sin t' = sin t \ cos t' = cos t" apply (simp add: exp_Euler sin_of_real cos_of_real) by (metis Complex_eq complex.sel) then obtain n::int where n: "t' = t + 2 * n * pi" by (auto simp: sin_cos_eq_iff) then have "n=0" by (cases n) (use t t' in \auto simp: mult_less_0_iff algebra_simps\) then show "t' = t" by (simp add: n) qed lemma Arg2pi: "0 \ Arg2pi z \ Arg2pi z < 2*pi \ is_Arg z (Arg2pi z)" proof (cases "z=0") case True then show ?thesis by (simp add: Arg2pi_def is_Arg_def) next case False obtain t where t: "0 \ t" "t < 2*pi" and ReIm: "Re z / cmod z = cos t" "Im z / cmod z = sin t" using sincos_total_2pi [OF complex_unit_circle [OF False]] by blast have z: "is_Arg z t" unfolding is_Arg_def apply (rule complex_eqI) using t False ReIm apply (auto simp: exp_Euler sin_of_real cos_of_real field_split_simps) done show ?thesis apply (simp add: Arg2pi_def False) apply (rule theI [where a=t]) using t z False apply (auto intro: Arg2pi_unique_lemma) done qed corollary\<^marker>\tag unimportant\ shows Arg2pi_ge_0: "0 \ Arg2pi z" and Arg2pi_lt_2pi: "Arg2pi z < 2*pi" and Arg2pi_eq: "z = of_real(norm z) * exp(\ * of_real(Arg2pi z))" using Arg2pi is_Arg_def by auto lemma complex_norm_eq_1_exp: "norm z = 1 \ exp(\ * of_real (Arg2pi z)) = z" by (metis Arg2pi_eq cis_conv_exp mult.left_neutral norm_cis of_real_1) lemma Arg2pi_unique: "\of_real r * exp(\ * of_real a) = z; 0 < r; 0 \ a; a < 2*pi\ \ Arg2pi z = a" by (rule Arg2pi_unique_lemma [unfolded is_Arg_def, OF _ Arg2pi_eq]) (use Arg2pi [of z] in \auto simp: norm_mult\) lemma Arg2pi_minus: "z \ 0 \ Arg2pi (-z) = (if Arg2pi z < pi then Arg2pi z + pi else Arg2pi z - pi)" apply (rule Arg2pi_unique [of "norm z"]) apply (rule complex_eqI) using Arg2pi_ge_0 [of z] Arg2pi_eq [of z] Arg2pi_lt_2pi [of z] apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric]) apply (metis Re_rcis Im_rcis rcis_def)+ done lemma Arg2pi_times_of_real [simp]: assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z" proof (cases "z=0") case False show ?thesis by (rule Arg2pi_unique [of "r * norm z"]) (use Arg2pi False assms is_Arg_def in auto) qed auto lemma Arg2pi_times_of_real2 [simp]: "0 < r \ Arg2pi (z * of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real mult.commute) lemma Arg2pi_divide_of_real [simp]: "0 < r \ Arg2pi (z / of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg2pi_le_pi: "Arg2pi z \ pi \ 0 \ Im z" proof (cases "z=0") case False have "0 \ Im z \ 0 \ Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 \ Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z \ pi" by (simp add: Im_exp) (metis Arg2pi_ge_0 Arg2pi_lt_2pi sin_lt_zero sin_ge_zero not_le) finally show ?thesis by blast qed auto lemma Arg2pi_lt_pi: "0 < Arg2pi z \ Arg2pi z < pi \ 0 < Im z" proof (cases "z=0") case False have "0 < Im z \ 0 < Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 < Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_less_mult_iff) also have "... \ 0 < Arg2pi z \ Arg2pi z < pi" using Arg2pi_ge_0 Arg2pi_lt_2pi sin_le_zero sin_gt_zero apply (auto simp: Im_exp) using le_less apply fastforce using not_le by blast finally show ?thesis by blast qed auto lemma Arg2pi_eq_0: "Arg2pi z = 0 \ z \ \ \ 0 \ Re z" proof (cases "z=0") case False have "z \ \ \ 0 \ Re z \ z \ \ \ 0 \ Re (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... \ z \ \ \ 0 \ Re (exp (\ * complex_of_real (Arg2pi z)))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z = 0" proof - have [simp]: "Arg2pi z = 0 \ z \ \" using Arg2pi_eq [of z] by (auto simp: Reals_def) moreover have "\z \ \; 0 \ cos (Arg2pi z)\ \ Arg2pi z = 0" by (metis Arg2pi_lt_pi Arg2pi_ge_0 Arg2pi_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl) ultimately show ?thesis by (auto simp: Re_exp) qed finally show ?thesis by blast qed auto corollary\<^marker>\tag unimportant\ Arg2pi_gt_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg2pi z > 0" using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order unfolding nonneg_Reals_def by fastforce lemma Arg2pi_eq_pi: "Arg2pi z = pi \ z \ \ \ Re z < 0" using Arg2pi_le_pi [of z] Arg2pi_lt_pi [of z] Arg2pi_eq_0 [of z] Arg2pi_ge_0 [of z] by (fastforce simp: complex_is_Real_iff) lemma Arg2pi_eq_0_pi: "Arg2pi z = 0 \ Arg2pi z = pi \ z \ \" using Arg2pi_eq_0 Arg2pi_eq_pi not_le by auto lemma Arg2pi_of_real: "Arg2pi (of_real r) = (if r<0 then pi else 0)" using Arg2pi_eq_0_pi Arg2pi_eq_pi by fastforce lemma Arg2pi_real: "z \ \ \ Arg2pi z = (if 0 \ Re z then 0 else pi)" using Arg2pi_eq_0 Arg2pi_eq_0_pi by auto lemma Arg2pi_inverse: "Arg2pi(inverse z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False show ?thesis apply (rule Arg2pi_unique [of "inverse (norm z)"]) using Arg2pi_eq False Arg2pi_ge_0 [of z] Arg2pi_lt_2pi [of z] Arg2pi_eq_0 [of z] by (auto simp: Arg2pi_real in_Reals_norm exp_diff field_simps) qed auto lemma Arg2pi_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg2pi w = Arg2pi z \ (\x. 0 < x & w = of_real x * z)" using assms Arg2pi_eq [of z] Arg2pi_eq [of w] apply auto apply (rule_tac x="norm w / norm z" in exI) apply (simp add: field_split_simps) by (metis mult.commute mult.left_commute) lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \ Arg2pi z = 0" by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq) lemma Arg2pi_divide: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi(z / w) = Arg2pi z - Arg2pi w" apply (rule Arg2pi_unique [of "norm(z / w)"]) using assms Arg2pi_eq Arg2pi_ge_0 [of w] Arg2pi_lt_2pi [of z] apply (auto simp: exp_diff norm_divide field_simps) done lemma Arg2pi_le_div_sum: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi z = Arg2pi w + Arg2pi(z / w)" by (simp add: Arg2pi_divide assms) lemma Arg2pi_le_div_sum_eq: assumes "w \ 0" "z \ 0" shows "Arg2pi w \ Arg2pi z \ Arg2pi z = Arg2pi w + Arg2pi(z / w)" using assms by (auto simp: Arg2pi_ge_0 intro: Arg2pi_le_div_sum) lemma Arg2pi_diff: assumes "w \ 0" "z \ 0" shows "Arg2pi w - Arg2pi z = (if Arg2pi z \ Arg2pi w then Arg2pi(w / z) else Arg2pi(w/z) - 2*pi)" using assms Arg2pi_divide Arg2pi_inverse [of "w/z"] Arg2pi_eq_0_pi by (force simp add: Arg2pi_ge_0 Arg2pi_divide not_le split: if_split_asm) lemma Arg2pi_add: assumes "w \ 0" "z \ 0" shows "Arg2pi w + Arg2pi z = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi(w * z) else Arg2pi(w * z) + 2*pi)" using assms Arg2pi_diff [of "w*z" z] Arg2pi_le_div_sum_eq [of z "w*z"] apply (auto simp: Arg2pi_ge_0 Arg2pi_divide not_le) apply (metis Arg2pi_lt_2pi add.commute) apply (metis (no_types) Arg2pi add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less) done lemma Arg2pi_times: assumes "w \ 0" "z \ 0" shows "Arg2pi (w * z) = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi w + Arg2pi z else (Arg2pi w + Arg2pi z) - 2*pi)" using Arg2pi_add [OF assms] by auto lemma Arg2pi_cnj_eq_inverse: "z\0 \ Arg2pi (cnj z) = Arg2pi (inverse z)" apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute) by (metis of_real_power zero_less_norm_iff zero_less_power) lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False then show ?thesis by (simp add: Arg2pi_cnj_eq_inverse Arg2pi_inverse) qed auto lemma Arg2pi_exp: "0 \ Im z \ Im z < 2*pi \ Arg2pi(exp z) = Im z" by (rule Arg2pi_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma complex_split_polar: obtains r a::real where "z = complex_of_real r * (cos a + \ * sin a)" "0 \ r" "0 \ a" "a < 2*pi" using Arg2pi cis.ctr cis_conv_exp unfolding Complex_eq is_Arg_def by fastforce lemma Re_Im_le_cmod: "Im w * sin \ + Re w * cos \ \ cmod w" proof (cases w rule: complex_split_polar) case (1 r a) with sin_cos_le1 [of a \] show ?thesis apply (simp add: norm_mult cmod_unit_one) by (metis (no_types, hide_lams) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le) qed subsection\<^marker>\tag unimportant\\Analytic properties of tangent function\ lemma cnj_tan: "cnj(tan z) = tan(cnj z)" by (simp add: cnj_cos cnj_sin tan_def) lemma field_differentiable_at_tan: "cos z \ 0 \ tan field_differentiable at z" unfolding field_differentiable_def using DERIV_tan by blast lemma field_differentiable_within_tan: "cos z \ 0 \ tan field_differentiable (at z within s)" using field_differentiable_at_tan field_differentiable_at_within by blast lemma continuous_within_tan: "cos z \ 0 \ continuous (at z within s) tan" using continuous_at_imp_continuous_within isCont_tan by blast lemma continuous_on_tan [continuous_intros]: "(\z. z \ s \ cos z \ 0) \ continuous_on s tan" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_tan: "(\z. z \ s \ cos z \ 0) \ tan holomorphic_on s" by (simp add: field_differentiable_within_tan holomorphic_on_def) subsection\The principal branch of the Complex logarithm\ instantiation complex :: ln begin definition\<^marker>\tag important\ ln_complex :: "complex \ complex" where "ln_complex \ \z. THE w. exp w = z & -pi < Im(w) & Im(w) \ pi" text\NOTE: within this scope, the constant Ln is not yet available!\ lemma assumes "z \ 0" shows exp_Ln [simp]: "exp(ln z) = z" and mpi_less_Im_Ln: "-pi < Im(ln z)" and Im_Ln_le_pi: "Im(ln z) \ pi" proof - obtain \ where z: "z / (cmod z) = Complex (cos \) (sin \)" using complex_unimodular_polar [of "z / (norm z)"] assms by (auto simp: norm_divide field_split_simps) obtain \ where \: "- pi < \" "\ \ pi" "sin \ = sin \" "cos \ = cos \" using sincos_principal_value [of "\"] assms by (auto simp: norm_divide field_split_simps) have "exp(ln z) = z & -pi < Im(ln z) & Im(ln z) \ pi" unfolding ln_complex_def apply (rule theI [where a = "Complex (ln(norm z)) \"]) using z assms \ apply (auto simp: field_simps exp_complex_eqI exp_eq_polar cis.code) done then show "exp(ln z) = z" "-pi < Im(ln z)" "Im(ln z) \ pi" by auto qed lemma Ln_exp [simp]: assumes "-pi < Im(z)" "Im(z) \ pi" shows "ln(exp z) = z" apply (rule exp_complex_eqI) using assms mpi_less_Im_Ln [of "exp z"] Im_Ln_le_pi [of "exp z"] apply auto done subsection\<^marker>\tag unimportant\\Relation to Real Logarithm\ lemma Ln_of_real: assumes "0 < z" shows "ln(of_real z::complex) = of_real(ln z)" proof - have "ln(of_real (exp (ln z))::complex) = ln (exp (of_real (ln z)))" by (simp add: exp_of_real) also have "... = of_real(ln z)" using assms by (subst Ln_exp) auto finally show ?thesis using assms by simp qed corollary\<^marker>\tag unimportant\ Ln_in_Reals [simp]: "z \ \ \ Re z > 0 \ ln z \ \" by (auto simp: Ln_of_real elim: Reals_cases) corollary\<^marker>\tag unimportant\ Im_Ln_of_real [simp]: "r > 0 \ Im (ln (of_real r)) = 0" by (simp add: Ln_of_real) lemma cmod_Ln_Reals [simp]: "z \ \ \ 0 < Re z \ cmod (ln z) = norm (ln (Re z))" using Ln_of_real by force lemma Ln_Reals_eq: "\x \ \; Re x > 0\ \ ln x = of_real (ln (Re x))" using Ln_of_real by force lemma Ln_1 [simp]: "ln 1 = (0::complex)" proof - have "ln (exp 0) = (0::complex)" by (simp add: del: exp_zero) then show ?thesis by simp qed lemma Ln_eq_zero_iff [simp]: "x \ \\<^sub>\\<^sub>0 \ ln x = 0 \ x = 1" for x::complex by auto (metis exp_Ln exp_zero nonpos_Reals_zero_I) instance by intro_classes (rule ln_complex_def Ln_1) end abbreviation Ln :: "complex \ complex" where "Ln \ ln" lemma Ln_eq_iff: "w \ 0 \ z \ 0 \ (Ln w = Ln z \ w = z)" by (metis exp_Ln) lemma Ln_unique: "exp(z) = w \ -pi < Im(z) \ Im(z) \ pi \ Ln w = z" using Ln_exp by blast lemma Re_Ln [simp]: "z \ 0 \ Re(Ln z) = ln(norm z)" by (metis exp_Ln ln_exp norm_exp_eq_Re) corollary\<^marker>\tag unimportant\ ln_cmod_le: assumes z: "z \ 0" shows "ln (cmod z) \ cmod (Ln z)" using norm_exp [of "Ln z", simplified exp_Ln [OF z]] by (metis Re_Ln complex_Re_le_cmod z) proposition\<^marker>\tag unimportant\ exists_complex_root: fixes z :: complex assumes "n \ 0" obtains w where "z = w ^ n" proof (cases "z=0") case False then show ?thesis by (rule_tac w = "exp(Ln z / n)" in that) (simp add: assms exp_of_nat_mult [symmetric]) qed (use assms in auto) corollary\<^marker>\tag unimportant\ exists_complex_root_nonzero: fixes z::complex assumes "z \ 0" "n \ 0" obtains w where "w \ 0" "z = w ^ n" by (metis exists_complex_root [of n z] assms power_0_left) subsection\<^marker>\tag unimportant\\Derivative of Ln away from the branch cut\ lemma assumes "z \ \\<^sub>\\<^sub>0" shows has_field_derivative_Ln: "(Ln has_field_derivative inverse(z)) (at z)" and Im_Ln_less_pi: "Im (Ln z) < pi" proof - have znz [simp]: "z \ 0" using assms by auto then have "Im (Ln z) \ pi" by (metis (no_types) Im_exp Ln_in_Reals assms complex_nonpos_Reals_iff complex_is_Real_iff exp_Ln mult_zero_right not_less pi_neq_zero sin_pi znz) then show *: "Im (Ln z) < pi" using assms Im_Ln_le_pi by (simp add: le_neq_trans) let ?U = "{w. -pi < Im(w) \ Im(w) < pi}" have 1: "open ?U" by (simp add: open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt) have 2: "\x. x \ ?U \ (exp has_derivative blinfun_apply(Blinfun ((*) (exp x)))) (at x)" apply (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) using DERIV_exp has_field_derivative_def by blast have 3: "continuous_on ?U (\x. Blinfun ((*) (exp x)))" unfolding blinfun_mult_right.abs_eq [symmetric] by (intro continuous_intros) have 4: "Ln z \ ?U" by (auto simp: mpi_less_Im_Ln *) have 5: "Blinfun ((*) (inverse z)) o\<^sub>L Blinfun ((*) (exp (Ln z))) = id_blinfun" by (rule blinfun_eqI) (simp add: bounded_linear_mult_right bounded_linear_Blinfun_apply) obtain U' V g g' where "open U'" and sub: "U' \ ?U" and "Ln z \ U'" "open V" "z \ V" and hom: "homeomorphism U' V exp g" and g: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv ((*) (exp (g y)))" and bij: "\y. y \ V \ bij ((*) (exp (g y)))" using inverse_function_theorem [OF 1 2 3 4 5] by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) blast show "(Ln has_field_derivative inverse(z)) (at z)" unfolding has_field_derivative_def proof (rule has_derivative_transform_within_open) show g_eq_Ln: "g y = Ln y" if "y \ V" for y proof - obtain x where "y = exp x" "x \ U'" using hom homeomorphism_image1 that \y \ V\ by blast then show ?thesis using sub hom homeomorphism_apply1 by fastforce qed have "0 \ V" by (meson exp_not_eq_zero hom homeomorphism_def) then have "\y. y \ V \ g' y = inv ((*) y)" by (metis exp_Ln g' g_eq_Ln) then have g': "g' z = (\x. x/z)" by (metis (no_types, hide_lams) bij \z \ V\ bij_inv_eq_iff exp_Ln g_eq_Ln nonzero_mult_div_cancel_left znz) show "(g has_derivative (*) (inverse z)) (at z)" using g [OF \z \ V\] g' by (simp add: \z \ V\ field_class.field_divide_inverse has_derivative_imp_has_field_derivative has_field_derivative_imp_has_derivative) qed (auto simp: \z \ V\ \open V\) qed declare has_field_derivative_Ln [derivative_intros] declare has_field_derivative_Ln [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable at z" using field_differentiable_def has_field_derivative_Ln by blast lemma field_differentiable_within_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable (at z within S)" using field_differentiable_at_Ln field_differentiable_within_subset by blast lemma continuous_at_Ln: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) Ln" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Ln) lemma isCont_Ln' [simp,continuous_intros]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. Ln (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_Ln]) lemma continuous_within_Ln [continuous_intros]: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Ln" using continuous_at_Ln continuous_at_imp_continuous_within by blast lemma continuous_on_Ln [continuous_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ continuous_on S Ln" by (simp add: continuous_at_imp_continuous_on continuous_within_Ln) lemma continuous_on_Ln' [continuous_intros]: "continuous_on S f \ (\z. z \ S \ f z \ \\<^sub>\\<^sub>0) \ continuous_on S (\x. Ln (f x))" by (rule continuous_on_compose2[OF continuous_on_Ln, of "UNIV - nonpos_Reals" S f]) auto lemma holomorphic_on_Ln [holomorphic_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ Ln holomorphic_on S" by (simp add: field_differentiable_within_Ln holomorphic_on_def) lemma holomorphic_on_Ln' [holomorphic_intros]: "(\z. z \ A \ f z \ \\<^sub>\\<^sub>0) \ f holomorphic_on A \ (\z. Ln (f z)) holomorphic_on A" using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \\<^sub>\\<^sub>0"] by (auto simp: o_def) lemma tendsto_Ln [tendsto_intros]: fixes L F assumes "(f \ L) F" "L \ \\<^sub>\\<^sub>0" shows "((\x. Ln (f x)) \ Ln L) F" proof - have "nhds L \ filtermap f F" using assms(1) by (simp add: filterlim_def) moreover have "\\<^sub>F y in nhds L. y \ - \\<^sub>\\<^sub>0" using eventually_nhds_in_open[of "- \\<^sub>\\<^sub>0" L] assms by (auto simp: open_Compl) ultimately have "\\<^sub>F y in filtermap f F. y \ - \\<^sub>\\<^sub>0" by (rule filter_leD) moreover have "continuous_on (-\\<^sub>\\<^sub>0) Ln" by (rule continuous_on_Ln) auto ultimately show ?thesis using continuous_on_tendsto_compose[of "- \\<^sub>\\<^sub>0" Ln f L F] assms by (simp add: eventually_filtermap) qed lemma divide_ln_mono: fixes x y::real assumes "3 \ x" "x \ y" shows "x / ln x \ y / ln y" proof (rule exE [OF complex_mvt_line [of x y "\z. z / Ln z" "\z. 1/(Ln z) - 1/(Ln z)^2"]]; clarsimp simp add: closed_segment_Reals closed_segment_eq_real_ivl assms) show "\u. \x \ u; u \ y\ \ ((\z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)" using \3 \ x\ by (force intro!: derivative_eq_intros simp: field_simps power_eq_if) show "x / ln x \ y / ln y" if "Re (y / Ln y) - Re (x / Ln x) = (Re (1 / Ln u) - Re (1 / (Ln u)\<^sup>2)) * (y - x)" and x: "x \ u" "u \ y" for u proof - have eq: "y / ln y = (1 / ln u - 1 / (ln u)\<^sup>2) * (y - x) + x / ln x" using that \3 \ x\ by (auto simp: Ln_Reals_eq in_Reals_norm group_add_class.diff_eq_eq) show ?thesis using exp_le \3 \ x\ x by (simp add: eq) (simp add: power_eq_if divide_simps ln_ge_iff) qed qed theorem Ln_series: fixes z :: complex assumes "norm z < 1" shows "(\n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\n. ?f n * z^n) sums _") proof - let ?F = "\z. \n. ?f n * z^n" and ?F' = "\z. \n. diffs ?f n * z^n" have r: "conv_radius ?f = 1" by (intro conv_radius_ratio_limit_nonzero[of _ 1]) (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc) have "\c. \z\ball 0 1. ln (1 + z) - ?F z = c" proof (rule has_field_derivative_zero_constant) fix z :: complex assume z': "z \ ball 0 1" hence z: "norm z < 1" by simp define t :: complex where "t = of_real (1 + norm z) / 2" from z have t: "norm z < norm t" "norm t < 1" unfolding t_def by (simp_all add: field_simps norm_divide del: of_real_add) have "Re (-z) \ norm (-z)" by (rule complex_Re_le_cmod) also from z have "... < 1" by simp finally have "((\z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)" by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff) moreover have "(?F has_field_derivative ?F' z) (at z)" using t r by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all ultimately have "((\z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z)) (at z within ball 0 1)" by (intro derivative_intros) (simp_all add: at_within_open[OF z']) also have "(\n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all from sums_split_initial_segment[OF this, of 1] have "(\i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc) hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse) also have "inverse (1 + z) - inverse (1 + z) = 0" by simp finally show "((\z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" . qed simp_all then obtain c where c: "\z. z \ ball 0 1 \ ln (1 + z) - ?F z = c" by blast from c[of 0] have "c = 0" by (simp only: powser_zero) simp with c[of z] assms have "ln (1 + z) = ?F z" by simp moreover have "summable (\n. ?f n * z^n)" using assms r by (intro summable_in_conv_radius) simp_all ultimately show ?thesis by (simp add: sums_iff) qed lemma Ln_series': "cmod z < 1 \ (\n. - ((-z)^n) / of_nat n) sums ln (1 + z)" by (drule Ln_series) (simp add: power_minus') lemma ln_series': assumes "abs (x::real) < 1" shows "(\n. - ((-x)^n) / of_nat n) sums ln (1 + x)" proof - from assms have "(\n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)" by (intro Ln_series') simp_all also have "(\n. - ((-of_real x)^n) / of_nat n) = (\n. complex_of_real (- ((-x)^n) / of_nat n))" by (rule ext) simp also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))" by (subst Ln_of_real [symmetric]) simp_all finally show ?thesis by (subst (asm) sums_of_real_iff) qed lemma Ln_approx_linear: fixes z :: complex assumes "norm z < 1" shows "norm (ln (1 + z) - z) \ norm z^2 / (1 - norm z)" proof - let ?f = "\n. (-1)^Suc n / of_nat n" from assms have "(\n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp moreover have "(\n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp ultimately have "(\n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)" by (subst left_diff_distrib, intro sums_diff) simp_all from sums_split_initial_segment[OF this, of "Suc 1"] have "(\i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)" by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse) hence "(Ln (1 + z) - z) = (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)" by (simp add: sums_iff) also have A: "summable (\n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))" by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]]) (auto simp: assms field_simps intro!: always_eventually) hence "norm (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \ (\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))" by (intro summable_norm) (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc) also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \ norm ((-z)^2 * (-z)^i) * 1" for i by (intro mult_left_mono) (simp_all add: field_split_simps) hence "(\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i)) \ (\i. norm (-(z^2) * (-z)^i))" using A assms apply (simp_all only: norm_power norm_inverse norm_divide norm_mult) apply (intro suminf_le summable_mult summable_geometric) apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc) done also have "... = norm z^2 * (\i. norm z^i)" using assms by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power) also have "(\i. norm z^i) = inverse (1 - norm z)" using assms by (subst suminf_geometric) (simp_all add: divide_inverse) also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Quadrant-type results for Ln\ lemma cos_lt_zero_pi: "pi/2 < x \ x < 3*pi/2 \ cos x < 0" using cos_minus_pi cos_gt_zero_pi [of "x-pi"] by simp lemma Re_Ln_pos_lt: assumes "z \ 0" shows "\Im(Ln z)\ < pi/2 \ 0 < Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "\Im w\ < pi/2 \ 0 < Re(exp w)" using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"] apply (auto simp: Re_exp zero_less_mult_iff cos_gt_zero_pi abs_if split: if_split_asm) apply (metis cos_minus cos_pi_half divide_minus_left less_irrefl linorder_neqE_linordered_idom nonzero_mult_div_cancel_right zero_neq_numeral)+ done } then show ?thesis using assms by auto qed lemma Re_Ln_pos_le: assumes "z \ 0" shows "\Im(Ln z)\ \ pi/2 \ 0 \ Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "\Im w\ \ pi/2 \ 0 \ Re(exp w)" apply (auto simp: Re_exp zero_le_mult_iff cos_ge_zero) using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le apply (auto simp: abs_if split: if_split_asm) done } then show ?thesis using assms by auto qed lemma Im_Ln_pos_lt: assumes "z \ 0" shows "0 < Im(Ln z) \ Im(Ln z) < pi \ 0 < Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "0 < Im w \ Im w < pi \ 0 < Im(exp w)" using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"] apply (simp add: Im_exp zero_less_mult_iff) using less_linear apply fastforce done } then show ?thesis using assms by auto qed lemma Im_Ln_pos_le: assumes "z \ 0" shows "0 \ Im(Ln z) \ Im(Ln z) \ pi \ 0 \ Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "0 \ Im w \ Im w \ pi \ 0 \ Im(exp w)" using sin_ge_zero [of "- (Im w)"] sin_ge_zero [of "(Im w)"] apply (auto simp: Im_exp zero_le_mult_iff sin_ge_zero) apply (metis not_le not_less_iff_gr_or_eq pi_not_less_zero sin_eq_0_pi) done } then show ?thesis using assms by auto qed lemma Re_Ln_pos_lt_imp: "0 < Re(z) \ \Im(Ln z)\ < pi/2" by (metis Re_Ln_pos_lt less_irrefl zero_complex.simps(1)) lemma Im_Ln_pos_lt_imp: "0 < Im(z) \ 0 < Im(Ln z) \ Im(Ln z) < pi" by (metis Im_Ln_pos_lt not_le order_refl zero_complex.simps(2)) text\A reference to the set of positive real numbers\ lemma Im_Ln_eq_0: "z \ 0 \ (Im(Ln z) = 0 \ 0 < Re(z) \ Im(z) = 0)" by (metis Im_complex_of_real Im_exp Ln_in_Reals Re_Ln_pos_lt Re_Ln_pos_lt_imp Re_complex_of_real complex_is_Real_iff exp_Ln exp_of_real pi_gt_zero) lemma Im_Ln_eq_pi: "z \ 0 \ (Im(Ln z) = pi \ Re(z) < 0 \ Im(z) = 0)" by (metis Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt add.left_neutral complex_eq less_eq_real_def mult_zero_right not_less_iff_gr_or_eq pi_ge_zero pi_neq_zero rcis_zero_arg rcis_zero_mod) subsection\<^marker>\tag unimportant\\More Properties of Ln\ lemma cnj_Ln: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ \ \Im (cnj (Ln z))\ + \Im (Ln (cnj z))\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (cnj (Ln z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (Ln (cnj z))\ \ pi" by (meson abs_le_iff complex_cnj_zero_iff less_eq_real_def minus_less_iff False Im_Ln_le_pi mpi_less_Im_Ln) ultimately show ?thesis by simp qed finally show "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ < 2 * pi" by simp show "exp (cnj (Ln z)) = exp (Ln (cnj z))" by (metis False complex_cnj_zero_iff exp_Ln exp_cnj) qed qed (use assms in auto) lemma Ln_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "Ln(inverse z) = -(Ln z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (Ln (inverse z)) - Im (- Ln z)\ \ \Im (Ln (inverse z))\ + \Im (- Ln z)\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (Ln (inverse z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (- Ln z)\ \ pi" using False Im_Ln_le_pi mpi_less_Im_Ln by fastforce ultimately show ?thesis by simp qed finally show "\Im (Ln (inverse z)) - Im (- Ln z)\ < 2 * pi" by simp show "exp (Ln (inverse z)) = exp (- Ln z)" by (simp add: False exp_minus) qed qed (use assms in auto) lemma Ln_minus1 [simp]: "Ln(-1) = \ * pi" apply (rule exp_complex_eqI) using Im_Ln_le_pi [of "-1"] mpi_less_Im_Ln [of "-1"] cis_conv_exp cis_pi apply (auto simp: abs_if) done lemma Ln_ii [simp]: "Ln \ = \ * of_real pi/2" using Ln_exp [of "\ * (of_real pi/2)"] unfolding exp_Euler by simp lemma Ln_minus_ii [simp]: "Ln(-\) = - (\ * pi/2)" proof - have "Ln(-\) = Ln(inverse \)" by simp also have "... = - (Ln \)" using Ln_inverse by blast also have "... = - (\ * pi/2)" by simp finally show ?thesis . qed lemma Ln_times: assumes "w \ 0" "z \ 0" shows "Ln(w * z) = (if Im(Ln w + Ln z) \ -pi then (Ln(w) + Ln(z)) + \ * of_real(2*pi) else if Im(Ln w + Ln z) > pi then (Ln(w) + Ln(z)) - \ * of_real(2*pi) else Ln(w) + Ln(z))" using pi_ge_zero Im_Ln_le_pi [of w] Im_Ln_le_pi [of z] using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z] by (auto simp: exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique) corollary\<^marker>\tag unimportant\ Ln_times_simple: "\w \ 0; z \ 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \ pi\ \ Ln(w * z) = Ln(w) + Ln(z)" by (simp add: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_of_real: "\r > 0; z \ 0\ \ Ln(of_real r * z) = ln r + Ln(z)" using mpi_less_Im_Ln Im_Ln_le_pi by (force simp: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_Reals: "\r \ Reals; Re r > 0; z \ 0\ \ Ln(r * z) = ln (Re r) + Ln(z)" using Ln_Reals_eq Ln_times_of_real by fastforce corollary\<^marker>\tag unimportant\ Ln_divide_of_real: "\r > 0; z \ 0\ \ Ln(z / of_real r) = Ln(z) - ln r" using Ln_times_of_real [of "inverse r" z] by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric] del: of_real_inverse) corollary\<^marker>\tag unimportant\ Ln_prod: fixes f :: "'a \ complex" assumes "finite A" "\x. x \ A \ f x \ 0" shows "\n. Ln (prod f A) = (\x \ A. Ln (f x) + (of_int (n x) * (2 * pi)) * \)" using assms proof (induction A) case (insert x A) then obtain n where n: "Ln (prod f A) = (\x\A. Ln (f x) + of_real (of_int (n x) * (2 * pi)) * \)" by auto define D where "D \ Im (Ln (f x)) + Im (Ln (prod f A))" define q::int where "q \ (if D \ -pi then 1 else if D > pi then -1 else 0)" have "prod f A \ 0" "f x \ 0" by (auto simp: insert.hyps insert.prems) with insert.hyps pi_ge_zero show ?case by (rule_tac x="n(x:=q)" in exI) (force simp: Ln_times q_def D_def n intro!: sum.cong) qed auto lemma Ln_minus: assumes "z \ 0" shows "Ln(-z) = (if Im(z) \ 0 \ \(Re(z) < 0 \ Im(z) = 0) then Ln(z) + \ * pi else Ln(z) - \ * pi)" (is "_ = ?rhs") using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] by (fastforce simp: exp_add exp_diff exp_Euler intro!: Ln_unique) lemma Ln_inverse_if: assumes "z \ 0" shows "Ln (inverse z) = (if z \ \\<^sub>\\<^sub>0 then -(Ln z) + \ * 2 * complex_of_real pi else -(Ln z))" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis by (simp add: Ln_inverse) next case True then have z: "Im z = 0" "Re z < 0" using assms apply (auto simp: complex_nonpos_Reals_iff) by (metis complex_is_Real_iff le_imp_less_or_eq of_real_0 of_real_Re) have "Ln(inverse z) = Ln(- (inverse (-z)))" by simp also have "... = Ln (inverse (-z)) + \ * complex_of_real pi" using assms z apply (simp add: Ln_minus) apply (simp add: field_simps) done also have "... = - Ln (- z) + \ * complex_of_real pi" apply (subst Ln_inverse) using z by (auto simp add: complex_nonneg_Reals_iff) also have "... = - (Ln z) + \ * 2 * complex_of_real pi" by (subst Ln_minus) (use assms z in auto) finally show ?thesis by (simp add: True) qed lemma Ln_times_ii: assumes "z \ 0" shows "Ln(\ * z) = (if 0 \ Re(z) | Im(z) < 0 then Ln(z) + \ * of_real pi/2 else Ln(z) - \ * of_real(3 * pi/2))" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] Re_Ln_pos_le [of z] by (simp add: Ln_times) auto lemma Ln_of_nat [simp]: "0 < n \ Ln (of_nat n) = of_real (ln (of_nat n))" by (subst of_real_of_nat_eq[symmetric], subst Ln_of_real[symmetric]) simp_all lemma Ln_of_nat_over_of_nat: assumes "m > 0" "n > 0" shows "Ln (of_nat m / of_nat n) = of_real (ln (of_nat m) - ln (of_nat n))" proof - have "of_nat m / of_nat n = (of_real (of_nat m / of_nat n) :: complex)" by simp also from assms have "Ln ... = of_real (ln (of_nat m / of_nat n))" by (simp add: Ln_of_real[symmetric]) also from assms have "... = of_real (ln (of_nat m) - ln (of_nat n))" by (simp add: ln_div) finally show ?thesis . qed subsection\The Argument of a Complex Number\ text\Finally: it's is defined for the same interval as the complex logarithm: \(-\,\]\.\ definition\<^marker>\tag important\ Arg :: "complex \ real" where "Arg z \ (if z = 0 then 0 else Im (Ln z))" lemma Arg_of_real: "Arg (of_real r) = (if r<0 then pi else 0)" by (simp add: Im_Ln_eq_pi Arg_def) lemma mpi_less_Arg: "-pi < Arg z" and Arg_le_pi: "Arg z \ pi" by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi) lemma assumes "z \ 0" shows Arg_eq: "z = of_real(norm z) * exp(\ * Arg z)" using assms exp_Ln exp_eq_polar by (auto simp: Arg_def) lemma is_Arg_Arg: "z \ 0 \ is_Arg z (Arg z)" by (simp add: Arg_eq is_Arg_def) lemma Argument_exists: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" proof - let ?rp = "r - Arg z + pi" define k where "k \ \?rp / (2 * pi)\" have "(Arg z + of_int k * (2 * pi)) \ R" using floor_divide_lower [of "2*pi" ?rp] floor_divide_upper [of "2*pi" ?rp] by (auto simp: k_def algebra_simps R) then show ?thesis using Arg_eq \z \ 0\ is_Arg_2pi_iff is_Arg_def that by blast qed lemma Argument_exists_unique: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" "\t. \is_Arg z t; t\R\ \ s=t" proof - obtain s where s: "is_Arg z s" "s\R" using Argument_exists [OF assms] . moreover have "\t. \is_Arg z t; t\R\ \ s=t" using assms s by (auto simp: is_Arg_eqI) ultimately show thesis using that by blast qed lemma Argument_Ex1: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" shows "\!s. is_Arg z s \ s \ R" using Argument_exists_unique [OF assms] by metis lemma Arg_divide: assumes "w \ 0" "z \ 0" shows "is_Arg (z / w) (Arg z - Arg w)" using Arg_eq [of z] Arg_eq [of w] Arg_eq [of "norm(z / w)"] assms by (auto simp: is_Arg_def norm_divide field_simps exp_diff Arg_of_real) lemma Arg_unique_lemma: assumes z: "is_Arg z t" and z': "is_Arg z t'" and t: "- pi < t" "t \ pi" and t': "- pi < t'" "t' \ pi" and nz: "z \ 0" shows "t' = t" using Arg2pi_unique_lemma [of "- (inverse z)"] proof - have "pi - t' = pi - t" proof (rule Arg2pi_unique_lemma [of "- (inverse z)"]) have "- (inverse z) = - (inverse (of_real(norm z) * exp(\ * t)))" by (metis is_Arg_def z) also have "... = (cmod (- inverse z)) * exp (\ * (pi - t))" by (auto simp: field_simps exp_diff norm_divide) finally show "is_Arg (- inverse z) (pi - t)" unfolding is_Arg_def . have "- (inverse z) = - (inverse (of_real(norm z) * exp(\ * t')))" by (metis is_Arg_def z') also have "... = (cmod (- inverse z)) * exp (\ * (pi - t'))" by (auto simp: field_simps exp_diff norm_divide) finally show "is_Arg (- inverse z) (pi - t')" unfolding is_Arg_def . qed (use assms in auto) then show ?thesis by simp qed lemma complex_norm_eq_1_exp_eq: "norm z = 1 \ exp(\ * (Arg z)) = z" by (metis Arg_eq exp_not_eq_zero exp_zero mult.left_neutral norm_zero of_real_1 norm_exp_i_times) lemma Arg_unique: "\of_real r * exp(\ * a) = z; 0 < r; -pi < a; a \ pi\ \ Arg z = a" by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq]) (use mpi_less_Arg Arg_le_pi in \auto simp: norm_mult\) lemma Arg_minus: assumes "z \ 0" shows "Arg (-z) = (if Arg z \ 0 then Arg z + pi else Arg z - pi)" proof - have [simp]: "cmod z * cos (Arg z) = Re z" using assms Arg_eq [of z] by (metis Re_exp exp_Ln norm_exp_eq_Re Arg_def) have [simp]: "cmod z * sin (Arg z) = Im z" using assms Arg_eq [of z] by (metis Im_exp exp_Ln norm_exp_eq_Re Arg_def) show ?thesis apply (rule Arg_unique [of "norm z"]) apply (rule complex_eqI) using mpi_less_Arg [of z] Arg_le_pi [of z] assms apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric]) done qed lemma Arg_times_of_real [simp]: assumes "0 < r" shows "Arg (of_real r * z) = Arg z" proof (cases "z=0") case True then show ?thesis by (simp add: Arg_def) next case False with Arg_eq assms show ?thesis by (auto simp: mpi_less_Arg Arg_le_pi intro!: Arg_unique [of "r * norm z"]) qed lemma Arg_times_of_real2 [simp]: "0 < r \ Arg (z * of_real r) = Arg z" by (metis Arg_times_of_real mult.commute) lemma Arg_divide_of_real [simp]: "0 < r \ Arg (z / of_real r) = Arg z" by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg_less_0: "0 \ Arg z \ 0 \ Im z" using Im_Ln_le_pi Im_Ln_pos_le by (simp add: Arg_def) lemma Arg_eq_pi: "Arg z = pi \ Re z < 0 \ Im z = 0" by (auto simp: Arg_def Im_Ln_eq_pi) lemma Arg_lt_pi: "0 < Arg z \ Arg z < pi \ 0 < Im z" using Arg_less_0 [of z] Im_Ln_pos_lt by (auto simp: order.order_iff_strict Arg_def) lemma Arg_eq_0: "Arg z = 0 \ z \ \ \ 0 \ Re z" using complex_is_Real_iff by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left) corollary\<^marker>\tag unimportant\ Arg_ne_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg z \ 0" using assms by (auto simp: nonneg_Reals_def Arg_eq_0) lemma Arg_eq_pi_iff: "Arg z = pi \ z \ \ \ Re z < 0" proof (cases "z=0") case False then show ?thesis using Arg_eq_0 [of "-z"] Arg_eq_pi complex_is_Real_iff by blast qed (simp add: Arg_def) lemma Arg_eq_0_pi: "Arg z = 0 \ Arg z = pi \ z \ \" using Arg_eq_pi_iff Arg_eq_0 by force lemma Arg_real: "z \ \ \ Arg z = (if 0 \ Re z then 0 else pi)" using Arg_eq_0 Arg_eq_0_pi by auto lemma Arg_inverse: "Arg(inverse z) = (if z \ \ then Arg z else - Arg z)" proof (cases "z \ \") case True then show ?thesis by simp (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse) next case False then have "Arg z < pi" "z \ 0" using Arg_eq_0_pi Arg_le_pi by (auto simp: less_eq_real_def) then show ?thesis apply (simp add: False) apply (rule Arg_unique [of "inverse (norm z)"]) using False mpi_less_Arg [of z] Arg_eq [of z] apply (auto simp: exp_minus field_simps) done qed lemma Arg_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg w = Arg z \ (\x. 0 < x \ w = of_real x * z)" using assms Arg_eq [of z] Arg_eq [of w] apply auto apply (rule_tac x="norm w / norm z" in exI) apply (simp add: field_split_simps) by (metis mult.commute mult.left_commute) lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \ Arg z = 0" by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq) lemma Arg_cnj_eq_inverse: "z\0 \ Arg (cnj z) = Arg (inverse z)" apply (simp add: Arg_eq_iff field_split_simps complex_norm_square [symmetric] mult.commute) by (metis of_real_power zero_less_norm_iff zero_less_power) lemma Arg_cnj: "Arg(cnj z) = (if z \ \ then Arg z else - Arg z)" by (metis Arg_cnj_eq_inverse Arg_inverse Reals_0 complex_cnj_zero) lemma Arg_exp: "-pi < Im z \ Im z \ pi \ Arg(exp z) = Im z" by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma Ln_Arg: "z\0 \ Ln(z) = ln(norm z) + \ * Arg(z)" by (metis Arg_def Re_Ln complex_eq) lemma continuous_at_Arg: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg" proof - have [simp]: "(\z. Im (Ln z)) \z\ Arg z" using Arg_def assms continuous_at by fastforce show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\w. \w \ - \\<^sub>\\<^sub>0; w \ z\ \ Im (Ln w) = Arg w" by (metis Arg_def Compl_iff nonpos_Reals_zero_I) qed (use assms in auto) qed lemma continuous_within_Arg: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Arg" using continuous_at_Arg continuous_at_imp_continuous_within by blast subsection\The Unwinding Number and the Ln product Formula\ text\Note that in this special case the unwinding number is -1, 0 or 1. But it's always an integer.\ lemma is_Arg_exp_Im: "is_Arg (exp z) (Im z)" using exp_eq_polar is_Arg_def norm_exp_eq_Re by auto lemma is_Arg_exp_diff_2pi: assumes "is_Arg (exp z) \" shows "\k. Im z - of_int k * (2 * pi) = \" proof (intro exI is_Arg_eqI) let ?k = "\(Im z - \) / (2 * pi)\" show "is_Arg (exp z) (Im z - real_of_int ?k * (2 * pi))" by (metis diff_add_cancel is_Arg_2pi_iff is_Arg_exp_Im) show "\Im z - real_of_int ?k * (2 * pi) - \\ < 2 * pi" using floor_divide_upper [of "2*pi" "Im z - \"] floor_divide_lower [of "2*pi" "Im z - \"] by (auto simp: algebra_simps abs_if) qed (auto simp: is_Arg_exp_Im assms) lemma Arg_exp_diff_2pi: "\k. Im z - of_int k * (2 * pi) = Arg (exp z)" using is_Arg_exp_diff_2pi [OF is_Arg_Arg] by auto lemma unwinding_in_Ints: "(z - Ln(exp z)) / (of_real(2*pi) * \) \ \" using Arg_exp_diff_2pi [of z] by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI) definition\<^marker>\tag important\ unwinding :: "complex \ int" where "unwinding z \ THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \)" lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \)" using unwinding_in_Ints [of z] unfolding unwinding_def Ints_def by force lemma unwinding_2pi: "(2*pi) * \ * unwinding(z) = z - Ln(exp z)" by (simp add: unwinding) lemma Ln_times_unwinding: "w \ 0 \ z \ 0 \ Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \ * unwinding(Ln w + Ln z)" using unwinding_2pi by (simp add: exp_add) subsection\<^marker>\tag unimportant\\Relation between Ln and Arg2pi, and hence continuity of Arg2pi\ lemma Arg2pi_Ln: assumes "0 < Arg2pi z" shows "Arg2pi z = Im(Ln(-z)) + pi" proof (cases "z = 0") case True with assms show ?thesis by simp next case False then have "z / of_real(norm z) = exp(\ * of_real(Arg2pi z))" using Arg2pi [of z] by (metis is_Arg_def abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff) then have "- z / of_real(norm z) = exp (\ * (of_real (Arg2pi z) - pi))" using cis_conv_exp cis_pi by (auto simp: exp_diff algebra_simps) then have "ln (- z / of_real(norm z)) = ln (exp (\ * (of_real (Arg2pi z) - pi)))" by simp also have "... = \ * (of_real(Arg2pi z) - pi)" using Arg2pi [of z] assms pi_not_less_zero by auto finally have "Arg2pi z = Im (Ln (- z / of_real (cmod z))) + pi" by simp also have "... = Im (Ln (-z) - ln (cmod z)) + pi" by (metis diff_0_right minus_diff_eq zero_less_norm_iff Ln_divide_of_real False) also have "... = Im (Ln (-z)) + pi" by simp finally show ?thesis . qed lemma continuous_at_Arg2pi: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg2pi" proof - have *: "isCont (\z. Im (Ln (- z)) + pi) z" by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+ have [simp]: "Im x \ 0 \ Im (Ln (- x)) + pi = Arg2pi x" for x using Arg2pi_Ln by (simp add: Arg2pi_gt_0 complex_nonneg_Reals_iff) consider "Re z < 0" | "Im z \ 0" using assms using complex_nonneg_Reals_iff not_le by blast then have [simp]: "(\z. Im (Ln (- z)) + pi) \z\ Arg2pi z" using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within) show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\x. \x \ - \\<^sub>\\<^sub>0; x \ z\ \ Im (Ln (- x)) + pi = Arg2pi x" by (auto simp add: Arg2pi_Ln [OF Arg2pi_gt_0] complex_nonneg_Reals_iff) qed (use assms in auto) qed text\Relation between Arg2pi and arctangent in upper halfplane\ lemma Arg2pi_arctan_upperhalf: assumes "0 < Im z" shows "Arg2pi z = pi/2 - arctan(Re z / Im z)" proof (cases "z = 0") case False show ?thesis proof (rule Arg2pi_unique [of "norm z"]) show "(cmod z) * exp (\ * (pi / 2 - arctan (Re z / Im z))) = z" apply (auto simp: exp_Euler cos_diff sin_diff) using assms norm_complex_def [of z, symmetric] apply (simp add: sin_of_real cos_of_real sin_arctan cos_arctan field_simps real_sqrt_divide) apply (metis complex_eq) done qed (use False arctan [of "Re z / Im z"] in auto) qed (use assms in auto) lemma Arg2pi_eq_Im_Ln: assumes "0 \ Im z" "0 < Re z" shows "Arg2pi z = Im (Ln z)" proof (cases "Im z = 0") case True then show ?thesis using Arg2pi_eq_0 Ln_in_Reals assms(2) complex_is_Real_iff by auto next case False then have *: "Arg2pi z > 0" using Arg2pi_gt_0 complex_is_Real_iff by blast then have "z \ 0" by auto with * assms False show ?thesis by (subst Arg2pi_Ln) (auto simp: Ln_minus) qed lemma continuous_within_upperhalf_Arg2pi: assumes "z \ 0" shows "continuous (at z within {z. 0 \ Im z}) Arg2pi" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis using continuous_at_Arg2pi continuous_at_imp_continuous_within by auto next case True then have z: "z \ \" "0 < Re z" using assms by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0) then have [simp]: "Arg2pi z = 0" "Im (Ln z) = 0" by (auto simp: Arg2pi_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff) show ?thesis proof (clarsimp simp add: continuous_within Lim_within dist_norm) fix e::real assume "0 < e" moreover have "continuous (at z) (\x. Im (Ln x))" using z by (simp add: continuous_at_Ln complex_nonpos_Reals_iff) ultimately obtain d where d: "d>0" "\x. x \ z \ cmod (x - z) < d \ \Im (Ln x)\ < e" by (auto simp: continuous_within Lim_within dist_norm) { fix x assume "cmod (x - z) < Re z / 2" then have "\Re x - Re z\ < Re z / 2" by (metis le_less_trans abs_Re_le_cmod minus_complex.simps(1)) then have "0 < Re x" using z by linarith } then show "\d>0. \x. 0 \ Im x \ x \ z \ cmod (x - z) < d \ \Arg2pi x\ < e" apply (rule_tac x="min d (Re z / 2)" in exI) using z d apply (auto simp: Arg2pi_eq_Im_Ln) done qed qed lemma continuous_on_upperhalf_Arg2pi: "continuous_on ({z. 0 \ Im z} - {0}) Arg2pi" unfolding continuous_on_eq_continuous_within by (metis DiffE Diff_subset continuous_within_subset continuous_within_upperhalf_Arg2pi insertCI) lemma open_Arg2pi2pi_less_Int: assumes "0 \ s" "t \ 2*pi" shows "open ({y. s < Arg2pi y} \ {y. Arg2pi y < t})" proof - have 1: "continuous_on (UNIV - \\<^sub>\\<^sub>0) Arg2pi" using continuous_at_Arg2pi continuous_at_imp_continuous_within by (auto simp: continuous_on_eq_continuous_within) have 2: "open (UNIV - \\<^sub>\\<^sub>0 :: complex set)" by (simp add: open_Diff) have "open ({z. s < z} \ {z. z < t})" using open_lessThan [of t] open_greaterThan [of s] by (metis greaterThan_def lessThan_def open_Int) moreover have "{y. s < Arg2pi y} \ {y. Arg2pi y < t} \ - \\<^sub>\\<^sub>0" using assms by (auto simp: Arg2pi_real complex_nonneg_Reals_iff complex_is_Real_iff) ultimately show ?thesis using continuous_imp_open_vimage [OF 1 2, of "{z. Re z > s} \ {z. Re z < t}"] by auto qed lemma open_Arg2pi2pi_gt: "open {z. t < Arg2pi z}" proof (cases "t < 0") case True then have "{z. t < Arg2pi z} = UNIV" using Arg2pi_ge_0 less_le_trans by auto then show ?thesis by simp next case False then show ?thesis using open_Arg2pi2pi_less_Int [of t "2*pi"] Arg2pi_lt_2pi by auto qed lemma closed_Arg2pi2pi_le: "closed {z. Arg2pi z \ t}" using open_Arg2pi2pi_gt [of t] by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le) subsection\<^marker>\tag unimportant\\Complex Powers\ lemma powr_to_1 [simp]: "z powr 1 = (z::complex)" by (simp add: powr_def) lemma powr_nat: fixes n::nat and z::complex shows "z powr n = (if z = 0 then 0 else z^n)" by (simp add: exp_of_nat_mult powr_def) lemma norm_powr_real: "w \ \ \ 0 < Re w \ norm(w powr z) = exp(Re z * ln(Re w))" apply (simp add: powr_def) using Im_Ln_eq_0 complex_is_Real_iff norm_complex_def by auto lemma powr_complexpow [simp]: fixes x::complex shows "x \ 0 \ x powr (of_nat n) = x^n" by (induct n) (auto simp: ac_simps powr_add) lemma powr_complexnumeral [simp]: fixes x::complex shows "x \ 0 \ x powr (numeral n) = x ^ (numeral n)" by (metis of_nat_numeral powr_complexpow) lemma cnj_powr: assumes "Im a = 0 \ Re a \ 0" shows "cnj (a powr b) = cnj a powr cnj b" proof (cases "a = 0") case False with assms have "a \ \\<^sub>\\<^sub>0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) with False show ?thesis by (simp add: powr_def exp_cnj cnj_Ln) qed simp lemma powr_real_real: assumes "w \ \" "z \ \" "0 < Re w" shows "w powr z = exp(Re z * ln(Re w))" proof - have "w \ 0" using assms by auto with assms show ?thesis by (simp add: powr_def Ln_Reals_eq of_real_exp) qed lemma powr_of_real: fixes x::real and y::real shows "0 \ x \ of_real x powr (of_real y::complex) = of_real (x powr y)" by (simp_all add: powr_def exp_eq_polar) lemma powr_of_int: fixes z::complex and n::int assumes "z\(0::complex)" shows "z powr of_int n = (if n\0 then z^nat n else inverse (z^nat (-n)))" by (metis assms not_le of_int_of_nat powr_complexpow powr_minus) lemma powr_Reals_eq: "\x \ \; y \ \; Re x \ 0\ \ x powr y = of_real (Re x powr Re y)" by (metis of_real_Re powr_of_real) lemma norm_powr_real_mono: "\w \ \; 1 < Re w\ \ cmod(w powr z1) \ cmod(w powr z2) \ Re z1 \ Re z2" by (auto simp: powr_def algebra_simps Reals_def Ln_of_real) lemma powr_times_real: "\x \ \; y \ \; 0 \ Re x; 0 \ Re y\ \ (x * y) powr z = x powr z * y powr z" by (auto simp: Reals_def powr_def Ln_times exp_add algebra_simps less_eq_real_def Ln_of_real) lemma Re_powr_le: "r \ \\<^sub>\\<^sub>0 \ Re (r powr z) \ Re r powr Re z" by (auto simp: powr_def nonneg_Reals_def order_trans [OF complex_Re_le_cmod]) lemma fixes w::complex shows Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \" and nonneg_Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \\<^sub>\\<^sub>0" by (auto simp: nonneg_Reals_def Reals_def powr_of_real) lemma powr_neg_real_complex: shows "(- of_real x) powr a = (-1) powr (of_real (sgn x) * a) * of_real x powr (a :: complex)" proof (cases "x = 0") assume x: "x \ 0" hence "(-x) powr a = exp (a * ln (-of_real x))" by (simp add: powr_def) also from x have "ln (-of_real x) = Ln (of_real x) + of_real (sgn x) * pi * \" by (simp add: Ln_minus Ln_of_real) also from x have "exp (a * ...) = cis pi powr (of_real (sgn x) * a) * of_real x powr a" by (simp add: powr_def exp_add algebra_simps Ln_of_real cis_conv_exp) also note cis_pi finally show ?thesis by simp qed simp_all lemma has_field_derivative_powr: fixes z :: complex assumes "z \ \\<^sub>\\<^sub>0" shows "((\z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)" proof (cases "z=0") case False show ?thesis unfolding powr_def - proof (rule DERIV_transform_at) + proof (rule has_field_derivative_transform_within) show "((\z. exp (s * Ln z)) has_field_derivative s * (if z = 0 then 0 else exp ((s - 1) * Ln z))) (at z)" apply (intro derivative_eq_intros | simp add: assms)+ by (simp add: False divide_complex_def exp_diff left_diff_distrib') qed (use False in auto) qed (use assms in auto) declare has_field_derivative_powr[THEN DERIV_chain2, derivative_intros] lemma has_field_derivative_powr_of_int: fixes z :: complex assumes gderiv:"(g has_field_derivative gd) (at z within s)" and "g z\0" shows "((\z. g z powr of_int n) has_field_derivative (n * g z powr (of_int n - 1) * gd)) (at z within s)" proof - define dd where "dd = of_int n * g z powr (of_int (n - 1)) * gd" obtain e where "e>0" and e_dist:"\y\s. dist z y < e \ g y \ 0" using DERIV_continuous[OF gderiv,THEN continuous_within_avoid] \g z\0\ by auto have ?thesis when "n\0" proof - define dd' where "dd' = of_int n * g z ^ (nat n - 1) * gd" have "dd=dd'" proof (cases "n=0") case False then have "n-1 \0" using \n\0\ by auto then have "g z powr (of_int (n - 1)) = g z ^ (nat n - 1)" using powr_of_int[OF \g z\0\,of "n-1"] by (simp add: nat_diff_distrib') then show ?thesis unfolding dd_def dd'_def by simp qed (simp add:dd_def dd'_def) then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within s)" by simp also have "... \ ((\z. g z ^ nat n) has_field_derivative dd') (at z within s)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within s. g x powr of_int n = g x ^ nat n" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." unfolding dd'_def using gderiv that by (auto intro!: derivative_eq_intros) finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s)" . then show ?thesis unfolding dd_def by simp qed moreover have ?thesis when "n<0" proof - define dd' where "dd' = of_int n / g z ^ (nat (1 - n)) * gd" have "dd=dd'" proof - have "g z powr of_int (n - 1) = inverse (g z ^ nat (1-n))" using powr_of_int[OF \g z\0\,of "n-1"] that by auto then show ?thesis unfolding dd_def dd'_def by (simp add: divide_inverse) qed then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within s)" by simp also have "... \ ((\z. inverse (g z ^ nat (-n))) has_field_derivative dd') (at z within s)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within s. g x powr of_int n = inverse (g x ^ nat (- n))" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." proof - have "nat (- n) + nat (1 - n) - Suc 0 = nat (- n) + nat (- n)" by auto then show ?thesis unfolding dd'_def using gderiv that \g z\0\ by (auto intro!: derivative_eq_intros simp add:field_split_simps power_add[symmetric]) qed finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within s)" . then show ?thesis unfolding dd_def by simp qed ultimately show ?thesis by force qed lemma field_differentiable_powr_of_int: fixes z :: complex assumes gderiv:"g field_differentiable (at z within s)" and "g z\0" shows "(\z. g z powr of_int n) field_differentiable (at z within s)" using has_field_derivative_powr_of_int assms(2) field_differentiable_def gderiv by blast lemma holomorphic_on_powr_of_int [holomorphic_intros]: assumes "f holomorphic_on s" "\z\s. f z\0" shows "(\z. (f z) powr of_int n) holomorphic_on s" proof (cases "n\0") case True then have "?thesis \ (\z. (f z) ^ nat n) holomorphic_on s" apply (rule_tac holomorphic_cong) using assms(2) by (auto simp add:powr_of_int) moreover have "(\z. (f z) ^ nat n) holomorphic_on s" using assms(1) by (auto intro:holomorphic_intros) ultimately show ?thesis by auto next case False then have "?thesis \ (\z. inverse (f z) ^ nat (-n)) holomorphic_on s" apply (rule_tac holomorphic_cong) using assms(2) by (auto simp add:powr_of_int power_inverse) moreover have "(\z. inverse (f z) ^ nat (-n)) holomorphic_on s" using assms by (auto intro!:holomorphic_intros) ultimately show ?thesis by auto qed lemma has_field_derivative_powr_right [derivative_intros]: "w \ 0 \ ((\z. w powr z) has_field_derivative Ln w * w powr z) (at z)" unfolding powr_def by (intro derivative_eq_intros | simp)+ lemma field_differentiable_powr_right [derivative_intros]: fixes w::complex shows "w \ 0 \ (\z. w powr z) field_differentiable (at z)" using field_differentiable_def has_field_derivative_powr_right by blast lemma holomorphic_on_powr_right [holomorphic_intros]: assumes "f holomorphic_on s" shows "(\z. w powr (f z)) holomorphic_on s" proof (cases "w = 0") case False with assms show ?thesis unfolding holomorphic_on_def field_differentiable_def by (metis (full_types) DERIV_chain' has_field_derivative_powr_right) qed simp lemma holomorphic_on_divide_gen [holomorphic_intros]: assumes f: "f holomorphic_on s" and g: "g holomorphic_on s" and 0: "\z z'. \z \ s; z' \ s\ \ g z = 0 \ g z' = 0" shows "(\z. f z / g z) holomorphic_on s" proof (cases "\z\s. g z = 0") case True with 0 have "g z = 0" if "z \ s" for z using that by blast then show ?thesis using g holomorphic_transform by auto next case False with 0 have "g z \ 0" if "z \ s" for z using that by blast with holomorphic_on_divide show ?thesis using f g by blast qed lemma norm_powr_real_powr: "w \ \ \ 0 \ Re w \ cmod (w powr z) = Re w powr Re z" by (metis dual_order.order_iff_strict norm_powr_real norm_zero of_real_0 of_real_Re powr_def) lemma tendsto_powr_complex: fixes f g :: "_ \ complex" assumes a: "a \ \\<^sub>\\<^sub>0" assumes f: "(f \ a) F" and g: "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" proof - from a have [simp]: "a \ 0" by auto from f g a have "((\z. exp (g z * ln (f z))) \ a powr b) F" (is ?P) by (auto intro!: tendsto_intros simp: powr_def) also { have "eventually (\z. z \ 0) (nhds a)" by (intro t1_space_nhds) simp_all with f have "eventually (\z. f z \ 0) F" using filterlim_iff by blast } hence "?P \ ((\z. f z powr g z) \ a powr b) F" by (intro tendsto_cong refl) (simp_all add: powr_def mult_ac) finally show ?thesis . qed lemma tendsto_powr_complex_0: fixes f g :: "'a \ complex" assumes f: "(f \ 0) F" and g: "(g \ b) F" and b: "Re b > 0" shows "((\z. f z powr g z) \ 0) F" proof (rule tendsto_norm_zero_cancel) define h where "h = (\z. if f z = 0 then 0 else exp (Re (g z) * ln (cmod (f z)) + abs (Im (g z)) * pi))" { fix z :: 'a assume z: "f z \ 0" define c where "c = abs (Im (g z)) * pi" from mpi_less_Im_Ln[OF z] Im_Ln_le_pi[OF z] have "abs (Im (Ln (f z))) \ pi" by simp from mult_left_mono[OF this, of "abs (Im (g z))"] have "abs (Im (g z) * Im (ln (f z))) \ c" by (simp add: abs_mult c_def) hence "-Im (g z) * Im (ln (f z)) \ c" by simp hence "norm (f z powr g z) \ h z" by (simp add: powr_def field_simps h_def c_def) } hence le: "norm (f z powr g z) \ h z" for z by (cases "f x = 0") (simp_all add: h_def) have g': "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all have "((\x. norm (f x)) \ 0) (inf F (principal {z. f z \ 0}))" by (subst tendsto_norm_zero_iff, rule tendsto_mono[OF _ f]) simp_all moreover { have "filterlim (\x. norm (f x)) (principal {0<..}) (principal {z. f z \ 0})" by (auto simp: filterlim_def) hence "filterlim (\x. norm (f x)) (principal {0<..}) (inf F (principal {z. f z \ 0}))" by (rule filterlim_mono) simp_all } ultimately have norm: "filterlim (\x. norm (f x)) (at_right 0) (inf F (principal {z. f z \ 0}))" by (simp add: filterlim_inf at_within_def) have A: "LIM x inf F (principal {z. f z \ 0}). Re (g x) * -ln (cmod (f x)) :> at_top" by (rule filterlim_tendsto_pos_mult_at_top tendsto_intros g' b filterlim_compose[OF filterlim_uminus_at_top_at_bot] filterlim_compose[OF ln_at_0] norm)+ have B: "LIM x inf F (principal {z. f z \ 0}). -\Im (g x)\ * pi + -(Re (g x) * ln (cmod (f x))) :> at_top" by (rule filterlim_tendsto_add_at_top tendsto_intros g')+ (insert A, simp_all) have C: "(h \ 0) F" unfolding h_def by (intro filterlim_If tendsto_const filterlim_compose[OF exp_at_bot]) (insert B, auto simp: filterlim_uminus_at_bot algebra_simps) show "((\x. norm (f x powr g x)) \ 0) F" by (rule Lim_null_comparison[OF always_eventually C]) (insert le, auto) qed lemma tendsto_powr_complex' [tendsto_intros]: fixes f g :: "_ \ complex" assumes "a \ \\<^sub>\\<^sub>0 \ (a = 0 \ Re b > 0)" and "(f \ a) F" "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" using assms tendsto_powr_complex tendsto_powr_complex_0 by fastforce lemma tendsto_neg_powr_complex_of_real: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. complex_of_real (f x) powr s) \ 0) F" proof - have "((\x. norm (complex_of_real (f x) powr s)) \ 0) F" proof (rule Lim_transform_eventually) from assms(1) have "eventually (\x. f x \ 0) F" by (auto simp: filterlim_at_top) thus "eventually (\x. f x powr Re s = norm (of_real (f x) powr s)) F" by eventually_elim (simp add: norm_powr_real_powr) from assms show "((\x. f x powr Re s) \ 0) F" by (intro tendsto_neg_powr) qed thus ?thesis by (simp add: tendsto_norm_zero_iff) qed lemma tendsto_neg_powr_complex_of_nat: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. of_nat (f x) powr s) \ 0) F" proof - have "((\x. of_real (real (f x)) powr s) \ 0) F" using assms(2) by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real] filterlim_compose[OF _ assms(1)] filterlim_real_sequentially filterlim_ident) auto thus ?thesis by simp qed lemma continuous_powr_complex: assumes "f (netlimit F) \ \\<^sub>\\<^sub>0" "continuous F f" "continuous F g" shows "continuous F (\z. f z powr g z :: complex)" using assms unfolding continuous_def by (intro tendsto_powr_complex) simp_all lemma isCont_powr_complex [continuous_intros]: assumes "f z \ \\<^sub>\\<^sub>0" "isCont f z" "isCont g z" shows "isCont (\z. f z powr g z :: complex) z" using assms unfolding isCont_def by (intro tendsto_powr_complex) simp_all lemma continuous_on_powr_complex [continuous_intros]: assumes "A \ {z. Re (f z) \ 0 \ Im (f z) \ 0}" assumes "\z. z \ A \ f z = 0 \ Re (g z) > 0" assumes "continuous_on A f" "continuous_on A g" shows "continuous_on A (\z. f z powr g z)" unfolding continuous_on_def proof fix z assume z: "z \ A" show "((\z. f z powr g z) \ f z powr g z) (at z within A)" proof (cases "f z = 0") case False from assms(1,2) z have "Re (f z) \ 0 \ Im (f z) \ 0" "f z = 0 \ Re (g z) > 0" by auto with assms(3,4) z show ?thesis by (intro tendsto_powr_complex') (auto elim!: nonpos_Reals_cases simp: complex_eq_iff continuous_on_def) next case True with assms z show ?thesis by (auto intro!: tendsto_powr_complex_0 simp: continuous_on_def) qed qed subsection\<^marker>\tag unimportant\\Some Limits involving Logarithms\ lemma lim_Ln_over_power: fixes s::complex assumes "0 < Re s" shows "(\n. Ln (of_nat n) / of_nat n powr s) \ 0" proof (simp add: lim_sequentially dist_norm, clarify) fix e::real assume e: "0 < e" have "\xo>0. \x\xo. 0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" proof (rule_tac x="2/(e * (Re s)\<^sup>2)" in exI, safe) show "0 < 2 / (e * (Re s)\<^sup>2)" using e assms by (simp add: field_simps) next fix x::real assume x: "2 / (e * (Re s)\<^sup>2) \ x" have "2 / (e * (Re s)\<^sup>2) > 0" using e assms by simp with x have "x > 0" by linarith then have "x * 2 \ e * (x\<^sup>2 * (Re s)\<^sup>2)" using e assms x by (auto simp: power2_eq_square field_simps) also have "... < e * (2 + (x * (Re s * 2) + x\<^sup>2 * (Re s)\<^sup>2))" using e assms \x > 0\ by (auto simp: power2_eq_square field_simps add_pos_pos) finally show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" by (auto simp: algebra_simps) qed then have "\xo>0. \x\xo. x / e < 1 + (Re s * x) + (1/2) * (Re s * x)^2" using e by (simp add: field_simps) then have "\xo>0. \x\xo. x / e < exp (Re s * x)" using assms by (force intro: less_le_trans [OF _ exp_lower_Taylor_quadratic]) then obtain xo where "xo > 0" and xo: "\x. x \ xo \ x < e * exp (Re s * x)" using e by (auto simp: field_simps) have "norm (Ln (of_nat n) / of_nat n powr s) < e" if "n \ nat \exp xo\" for n using e xo [of "ln n"] that apply (auto simp: norm_divide norm_powr_real field_split_simps) apply (metis exp_less_mono exp_ln not_le of_nat_0_less_iff) done then show "\no. \n\no. norm (Ln (of_nat n) / of_nat n powr s) < e" by blast qed lemma lim_Ln_over_n: "((\n. Ln(of_nat n) / of_nat n) \ 0) sequentially" using lim_Ln_over_power [of 1] by simp lemma lim_ln_over_power: fixes s :: real assumes "0 < s" shows "((\n. ln n / (n powr s)) \ 0) sequentially" using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_ln_over_n [tendsto_intros]: "((\n. ln(real_of_nat n) / of_nat n) \ 0) sequentially" using lim_ln_over_power [of 1, THEN filterlim_sequentially_Suc [THEN iffD2]] apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) done lemma lim_log_over_n [tendsto_intros]: "(\n. log k n/n) \ 0" proof - have *: "log k n/n = (1/ln k) * (ln n / n)" for n unfolding log_def by auto have "(\n. (1/ln k) * (ln n / n)) \ (1/ln k) * 0" by (intro tendsto_intros) then show ?thesis unfolding * by auto qed lemma lim_1_over_complex_power: assumes "0 < Re s" shows "(\n. 1 / of_nat n powr s) \ 0" proof (rule Lim_null_comparison) have "\n>0. 3 \ n \ 1 \ ln (real_of_nat n)" using ln_272_gt_1 by (force intro: order_trans [of _ "ln (272/100)"]) then show "\\<^sub>F x in sequentially. cmod (1 / of_nat x powr s) \ cmod (Ln (of_nat x) / of_nat x powr s)" by (auto simp: norm_divide field_split_simps eventually_sequentially) show "(\n. cmod (Ln (of_nat n) / of_nat n powr s)) \ 0" using lim_Ln_over_power [OF assms] by (metis tendsto_norm_zero_iff) qed lemma lim_1_over_real_power: fixes s :: real assumes "0 < s" shows "((\n. 1 / (of_nat n powr s)) \ 0) sequentially" using lim_1_over_complex_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_1_over_Ln: "((\n. 1 / Ln(of_nat n)) \ 0) sequentially" proof (clarsimp simp add: lim_sequentially dist_norm norm_divide field_split_simps) fix r::real assume "0 < r" have ir: "inverse (exp (inverse r)) > 0" by simp obtain n where n: "1 < of_nat n * inverse (exp (inverse r))" using ex_less_of_nat_mult [of _ 1, OF ir] by auto then have "exp (inverse r) < of_nat n" by (simp add: field_split_simps) then have "ln (exp (inverse r)) < ln (of_nat n)" by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff) with \0 < r\ have "1 < r * ln (real_of_nat n)" by (simp add: field_simps) moreover have "n > 0" using n using neq0_conv by fastforce ultimately show "\no. \k. Ln (of_nat k) \ 0 \ no \ k \ 1 < r * cmod (Ln (of_nat k))" using n \0 < r\ by (rule_tac x=n in exI) (force simp: field_split_simps intro: less_le_trans) qed lemma lim_1_over_ln: "((\n. 1 / ln(real_of_nat n)) \ 0) sequentially" using lim_1_over_Ln [THEN filterlim_sequentially_Suc [THEN iffD2]] apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp add: lim_sequentially dist_norm) apply (simp add: Ln_Reals_eq norm_powr_real_powr norm_divide) done lemma lim_ln1_over_ln: "(\n. ln(Suc n) / ln n) \ 1" proof (rule Lim_transform_eventually) have "(\n. ln(1 + 1/n) / ln n) \ 0" proof (rule Lim_transform_bound) show "(inverse o real) \ 0" by (metis comp_def lim_inverse_n lim_explicit) show "\\<^sub>F n in sequentially. norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" proof fix n::nat assume n: "3 \ n" then have "ln 3 \ ln n" and ln0: "0 \ ln n" by auto with ln3_gt_1 have "1/ ln n \ 1" by (simp add: field_split_simps) moreover have "ln (1 + 1 / real n) \ 1/n" by (simp add: ln_add_one_self_le_self) ultimately have "ln (1 + 1 / real n) * (1 / ln n) \ (1/n) * 1" by (intro mult_mono) (use n in auto) then show "norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" by (simp add: field_simps ln0) qed qed then show "(\n. 1 + ln(1 + 1/n) / ln n) \ 1" by (metis (full_types) add.right_neutral tendsto_add_const_iff) show "\\<^sub>F k in sequentially. 1 + ln (1 + 1 / k) / ln k = ln(Suc k) / ln k" by (simp add: field_split_simps ln_div eventually_sequentiallyI [of 2]) qed lemma lim_ln_over_ln1: "(\n. ln n / ln(Suc n)) \ 1" proof - have "(\n. inverse (ln(Suc n) / ln n)) \ inverse 1" by (rule tendsto_inverse [OF lim_ln1_over_ln]) auto then show ?thesis by simp qed subsection\<^marker>\tag unimportant\\Relation between Square Root and exp/ln, hence its derivative\ lemma csqrt_exp_Ln: assumes "z \ 0" shows "csqrt z = exp(Ln(z) / 2)" proof - have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))" by (metis exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) also have "... = z" using assms exp_Ln by blast finally have "csqrt z = csqrt ((exp (Ln z / 2))\<^sup>2)" by simp also have "... = exp (Ln z / 2)" apply (subst csqrt_square) using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms apply (auto simp: Re_exp Im_exp zero_less_mult_iff zero_le_mult_iff, fastforce+) done finally show ?thesis using assms csqrt_square by simp qed lemma csqrt_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "csqrt (inverse z) = inverse (csqrt z)" proof (cases "z=0") case False then show ?thesis using assms csqrt_exp_Ln Ln_inverse exp_minus by (simp add: csqrt_exp_Ln Ln_inverse exp_minus) qed auto lemma cnj_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(csqrt z) = csqrt(cnj z)" proof (cases "z=0") case False then show ?thesis by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj) qed auto lemma has_field_derivative_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "(csqrt has_field_derivative inverse(2 * csqrt z)) (at z)" proof - have z: "z \ 0" using assms by auto then have *: "inverse z = inverse (2*z) * 2" by (simp add: field_split_simps) have [simp]: "exp (Ln z / 2) * inverse z = inverse (csqrt z)" by (simp add: z field_simps csqrt_exp_Ln [symmetric]) (metis power2_csqrt power2_eq_square) have "Im z = 0 \ 0 < Re z" using assms complex_nonpos_Reals_iff not_less by blast with z have "((\z. exp (Ln z / 2)) has_field_derivative inverse (2 * csqrt z)) (at z)" by (force intro: derivative_eq_intros * simp add: assms) then show ?thesis - proof (rule DERIV_transform_at) + proof (rule has_field_derivative_transform_within) show "\x. dist x z < cmod z \ exp (Ln x / 2) = csqrt x" by (metis csqrt_exp_Ln dist_0_norm less_irrefl) qed (use z in auto) qed lemma field_differentiable_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable at z" using field_differentiable_def has_field_derivative_csqrt by blast lemma field_differentiable_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable (at z within s)" using field_differentiable_at_csqrt field_differentiable_within_subset by blast lemma continuous_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) csqrt" by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at) corollary\<^marker>\tag unimportant\ isCont_csqrt' [simp]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. csqrt (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_csqrt]) lemma continuous_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within s) csqrt" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_csqrt) lemma continuous_on_csqrt [continuous_intros]: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ continuous_on s csqrt" by (simp add: continuous_at_imp_continuous_on continuous_within_csqrt) lemma holomorphic_on_csqrt: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ csqrt holomorphic_on s" by (simp add: field_differentiable_within_csqrt holomorphic_on_def) lemma continuous_within_closed_nontrivial: "closed s \ a \ s ==> continuous (at a within s) f" using open_Compl by (force simp add: continuous_def eventually_at_topological filterlim_iff open_Collect_neg) lemma continuous_within_csqrt_posreal: "continuous (at z within (\ \ {w. 0 \ Re(w)})) csqrt" proof (cases "z \ \\<^sub>\\<^sub>0") case True have *: "\e. \0 < e\ \ \x'\\ \ {w. 0 \ Re w}. cmod x' < e^2 \ cmod (csqrt x') < e" by (auto simp: Reals_def real_less_lsqrt) have "Im z = 0" "Re z < 0 \ z = 0" using True cnj.code complex_cnj_zero_iff by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce with * show ?thesis apply (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge]) apply (auto simp: continuous_within_eps_delta) using zero_less_power by blast next case False then show ?thesis by (blast intro: continuous_within_csqrt) qed subsection\Complex arctangent\ text\The branch cut gives standard bounds in the real case.\ definition\<^marker>\tag important\ Arctan :: "complex \ complex" where "Arctan \ \z. (\/2) * Ln((1 - \*z) / (1 + \*z))" lemma Arctan_def_moebius: "Arctan z = \/2 * Ln (moebius (-\) 1 \ 1 z)" by (simp add: Arctan_def moebius_def add_ac) lemma Ln_conv_Arctan: assumes "z \ -1" shows "Ln z = -2*\ * Arctan (moebius 1 (- 1) (- \) (- \) z)" proof - have "Arctan (moebius 1 (- 1) (- \) (- \) z) = \/2 * Ln (moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z))" by (simp add: Arctan_def_moebius) also from assms have "\ * z \ \ * (-1)" by (subst mult_left_cancel) simp hence "\ * z - -\ \ 0" by (simp add: eq_neg_iff_add_eq_0) from moebius_inverse'[OF _ this, of 1 1] have "moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z) = z" by simp finally show ?thesis by (simp add: field_simps) qed lemma Arctan_0 [simp]: "Arctan 0 = 0" by (simp add: Arctan_def) lemma Im_complex_div_lemma: "Im((1 - \*z) / (1 + \*z)) = 0 \ Re z = 0" by (auto simp: Im_complex_div_eq_0 algebra_simps) lemma Re_complex_div_lemma: "0 < Re((1 - \*z) / (1 + \*z)) \ norm z < 1" by (simp add: Re_complex_div_gt_0 algebra_simps cmod_def power2_eq_square) lemma tan_Arctan: assumes "z\<^sup>2 \ -1" shows [simp]:"tan(Arctan z) = z" proof - have "1 + \*z \ 0" by (metis assms complex_i_mult_minus i_squared minus_unique power2_eq_square power2_minus) moreover have "1 - \*z \ 0" by (metis assms complex_i_mult_minus i_squared power2_eq_square power2_minus right_minus_eq) ultimately show ?thesis by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus csqrt_exp_Ln [symmetric] divide_simps power2_eq_square [symmetric]) qed lemma Arctan_tan [simp]: assumes "\Re z\ < pi/2" shows "Arctan(tan z) = z" proof - have ge_pi2: "\n::int. \of_int (2*n + 1) * pi/2\ \ pi/2" by (case_tac n rule: int_cases) (auto simp: abs_mult) have "exp (\*z)*exp (\*z) = -1 \ exp (2*\*z) = -1" by (metis distrib_right exp_add mult_2) also have "... \ exp (2*\*z) = exp (\*pi)" using cis_conv_exp cis_pi by auto also have "... \ exp (2*\*z - \*pi) = 1" by (metis (no_types) diff_add_cancel diff_minus_eq_add exp_add exp_minus_inverse mult.commute) also have "... \ Re(\*2*z - \*pi) = 0 \ (\n::int. Im(\*2*z - \*pi) = of_int (2 * n) * pi)" by (simp add: exp_eq_1) also have "... \ Im z = 0 \ (\n::int. 2 * Re z = of_int (2*n + 1) * pi)" by (simp add: algebra_simps) also have "... \ False" using assms ge_pi2 apply (auto simp: algebra_simps) by (metis abs_mult_pos not_less of_nat_less_0_iff of_nat_numeral) finally have *: "exp (\*z)*exp (\*z) + 1 \ 0" by (auto simp: add.commute minus_unique) show ?thesis using assms * apply (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps i_times_eq_iff power2_eq_square [symmetric]) apply (rule Ln_unique) apply (auto simp: divide_simps exp_minus) apply (simp add: algebra_simps exp_double [symmetric]) done qed lemma assumes "Re z = 0 \ \Im z\ < 1" shows Re_Arctan_bounds: "\Re(Arctan z)\ < pi/2" and has_field_derivative_Arctan: "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" proof - have nz0: "1 + \*z \ 0" using assms by (metis abs_one add_diff_cancel_left' complex_i_mult_minus diff_0 i_squared imaginary_unit.simps less_asym neg_equal_iff_equal) have "z \ -\" using assms by auto then have zz: "1 + z * z \ 0" by (metis abs_one assms i_squared imaginary_unit.simps less_irrefl minus_unique square_eq_iff) have nz1: "1 - \*z \ 0" using assms by (force simp add: i_times_eq_iff) have nz2: "inverse (1 + \*z) \ 0" using assms by (metis Im_complex_div_lemma Re_complex_div_lemma cmod_eq_Im divide_complex_def less_irrefl mult_zero_right zero_complex.simps(1) zero_complex.simps(2)) have nzi: "((1 - \*z) * inverse (1 + \*z)) \ 0" using nz1 nz2 by auto have "Im ((1 - \*z) / (1 + \*z)) = 0 \ 0 < Re ((1 - \*z) / (1 + \*z))" apply (simp add: divide_complex_def) apply (simp add: divide_simps split: if_split_asm) using assms apply (auto simp: algebra_simps abs_square_less_1 [unfolded power2_eq_square]) done then have *: "((1 - \*z) / (1 + \*z)) \ \\<^sub>\\<^sub>0" by (auto simp add: complex_nonpos_Reals_iff) show "\Re(Arctan z)\ < pi/2" unfolding Arctan_def divide_complex_def using mpi_less_Im_Ln [OF nzi] apply (auto simp: abs_if intro!: Im_Ln_less_pi * [unfolded divide_complex_def]) done show "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" unfolding Arctan_def scaleR_conv_of_real apply (intro derivative_eq_intros | simp add: nz0 *)+ using nz0 nz1 zz apply (simp add: algebra_simps field_split_simps power2_eq_square) apply algebra done qed lemma field_differentiable_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable at z" using has_field_derivative_Arctan by (auto simp: field_differentiable_def) lemma field_differentiable_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable (at z within s)" using field_differentiable_at_Arctan field_differentiable_at_within by blast declare has_field_derivative_Arctan [derivative_intros] declare has_field_derivative_Arctan [THEN DERIV_chain2, derivative_intros] lemma continuous_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z) Arctan" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Arctan) lemma continuous_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z within s) Arctan" using continuous_at_Arctan continuous_at_imp_continuous_within by blast lemma continuous_on_Arctan [continuous_intros]: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ continuous_on s Arctan" by (auto simp: continuous_at_imp_continuous_on continuous_within_Arctan) lemma holomorphic_on_Arctan: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ Arctan holomorphic_on s" by (simp add: field_differentiable_within_Arctan holomorphic_on_def) theorem Arctan_series: assumes z: "norm (z :: complex) < 1" defines "g \ \n. if odd n then -\*\^n / n else 0" defines "h \ \z n. (-1)^n / of_nat (2*n+1) * (z::complex)^(2*n+1)" shows "(\n. g n * z^n) sums Arctan z" and "h z sums Arctan z" proof - define G where [abs_def]: "G z = (\n. g n * z^n)" for z have summable: "summable (\n. g n * u^n)" if "norm u < 1" for u proof (cases "u = 0") assume u: "u \ 0" have "(\n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\n. ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))" proof fix n have "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal (((2*Suc n+1) / (Suc n)) / ((2*Suc n-1) / (Suc n)))" by (simp add: h_def norm_mult norm_power norm_divide field_split_simps power2_eq_square eval_nat_numeral del: of_nat_add of_nat_Suc) also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? finally show "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n))))" . qed also have "\ \ ereal (inverse (norm u)^2) * ereal ((2 + 0) / (2 - 0))" by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all finally have "liminf (\n. ereal (cmod (h u n) / cmod (h u (Suc n)))) = inverse (norm u)^2" by (intro lim_imp_Liminf) simp_all moreover from power_strict_mono[OF that, of 2] u have "inverse (norm u)^2 > 1" by (simp add: field_split_simps) ultimately have A: "liminf (\n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp from u have "summable (h u)" by (intro summable_norm_cancel[OF ratio_test_convergence[OF _ A]]) (auto simp: h_def norm_divide norm_mult norm_power simp del: of_nat_Suc intro!: mult_pos_pos divide_pos_pos always_eventually) thus "summable (\n. g n * u^n)" by (subst summable_mono_reindex[of "\n. 2*n+1", symmetric]) (auto simp: power_mult strict_mono_def g_def h_def elim!: oddE) qed (simp add: h_def) have "\c. \u\ball 0 1. Arctan u - G u = c" proof (rule has_field_derivative_zero_constant) fix u :: complex assume "u \ ball 0 1" hence u: "norm u < 1" by (simp add: dist_0_norm) define K where "K = (norm u + 1) / 2" from u and abs_Im_le_cmod[of u] have Im_u: "\Im u\ < 1" by linarith from u have K: "0 \ K" "norm u < K" "K < 1" by (simp_all add: K_def) hence "(G has_field_derivative (\n. diffs g n * u ^ n)) (at u)" unfolding G_def by (intro termdiffs_strong[of _ "of_real K"] summable) simp_all also have "(\n. diffs g n * u^n) = (\n. if even n then (\*u)^n else 0)" by (intro ext) (simp_all del: of_nat_Suc add: g_def diffs_def power_mult_distrib) also have "suminf \ = (\n. (-(u^2))^n)" by (subst suminf_mono_reindex[of "\n. 2*n", symmetric]) (auto elim!: evenE simp: strict_mono_def power_mult power_mult_distrib) also from u have "norm u^2 < 1^2" by (intro power_strict_mono) simp_all hence "(\n. (-(u^2))^n) = inverse (1 + u^2)" by (subst suminf_geometric) (simp_all add: norm_power inverse_eq_divide) finally have "(G has_field_derivative inverse (1 + u\<^sup>2)) (at u)" . from DERIV_diff[OF has_field_derivative_Arctan this] Im_u u show "((\u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)" by (simp_all add: at_within_open[OF _ open_ball]) qed simp_all then obtain c where c: "\u. norm u < 1 \ Arctan u - G u = c" by auto from this[of 0] have "c = 0" by (simp add: G_def g_def) with c z have "Arctan z = G z" by simp with summable[OF z] show "(\n. g n * z^n) sums Arctan z" unfolding G_def by (simp add: sums_iff) thus "h z sums Arctan z" by (subst (asm) sums_mono_reindex[of "\n. 2*n+1", symmetric]) (auto elim!: oddE simp: strict_mono_def power_mult g_def h_def) qed text \A quickly-converging series for the logarithm, based on the arctangent.\ theorem ln_series_quadratic: assumes x: "x > (0::real)" shows "(\n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x" proof - define y :: complex where "y = of_real ((x-1)/(x+1))" from x have x': "complex_of_real x \ of_real (-1)" by (subst of_real_eq_iff) auto from x have "\x - 1\ < \x + 1\" by linarith hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1" by (simp add: norm_divide del: of_real_add of_real_diff) hence "norm (\ * y) < 1" unfolding y_def by (subst norm_mult) simp hence "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) sums ((-2*\) * Arctan (\*y))" by (intro Arctan_series sums_mult) simp_all also have "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) = (\n. (-2*\) * ((-1)^n * (\*y*(-y\<^sup>2)^n)/of_nat (2*n+1)))" by (intro ext) (simp_all add: power_mult power_mult_distrib) also have "\ = (\n. 2*y* ((-1) * (-y\<^sup>2))^n/of_nat (2*n+1))" by (intro ext, subst power_mult_distrib) (simp add: algebra_simps power_mult) also have "\ = (\n. 2*y^(2*n+1) / of_nat (2*n+1))" by (subst power_add, subst power_mult) (simp add: mult_ac) also have "\ = (\n. of_real (2*((x-1)/(x+1))^(2*n+1) / of_nat (2*n+1)))" by (intro ext) (simp add: y_def) also have "\ * y = (of_real x - 1) / (-\ * (of_real x + 1))" by (subst divide_divide_eq_left [symmetric]) (simp add: y_def) also have "\ = moebius 1 (-1) (-\) (-\) (of_real x)" by (simp add: moebius_def algebra_simps) also from x' have "-2*\*Arctan \ = Ln (of_real x)" by (intro Ln_conv_Arctan [symmetric]) simp_all also from x have "\ = ln x" by (rule Ln_of_real) finally show ?thesis by (subst (asm) sums_of_real_iff) qed subsection\<^marker>\tag unimportant\ \Real arctangent\ lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0" proof - have ne: "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) have "Re (Ln ((1 - \ * x) * inverse (1 + \ * x))) = 0" apply (rule norm_exp_imaginary) apply (subst exp_Ln) using ne apply (simp_all add: cmod_def complex_eq_iff) apply (auto simp: field_split_simps) apply algebra done then show ?thesis unfolding Arctan_def divide_complex_def by (simp add: complex_eq_iff) qed lemma arctan_eq_Re_Arctan: "arctan x = Re (Arctan (of_real x))" proof (rule arctan_unique) show "- (pi / 2) < Re (Arctan (complex_of_real x))" apply (simp add: Arctan_def) apply (rule Im_Ln_less_pi) apply (auto simp: Im_complex_div_lemma complex_nonpos_Reals_iff) done next have *: " (1 - \*x) / (1 + \*x) \ 0" by (simp add: field_split_simps) ( simp add: complex_eq_iff) show "Re (Arctan (complex_of_real x)) < pi / 2" using mpi_less_Im_Ln [OF *] by (simp add: Arctan_def) next have "tan (Re (Arctan (of_real x))) = Re (tan (Arctan (of_real x)))" apply (auto simp: tan_def Complex.Re_divide Re_sin Re_cos Im_sin Im_cos) apply (simp add: field_simps) by (simp add: power2_eq_square) also have "... = x" apply (subst tan_Arctan, auto) by (metis diff_0_right minus_diff_eq mult_zero_left not_le of_real_1 of_real_eq_iff of_real_minus of_real_power power2_eq_square real_minus_mult_self_le zero_less_one) finally show "tan (Re (Arctan (complex_of_real x))) = x" . qed lemma Arctan_of_real: "Arctan (of_real x) = of_real (arctan x)" unfolding arctan_eq_Re_Arctan divide_complex_def by (simp add: complex_eq_iff) lemma Arctan_in_Reals [simp]: "z \ \ \ Arctan z \ \" by (metis Reals_cases Reals_of_real Arctan_of_real) declare arctan_one [simp] lemma arctan_less_pi4_pos: "x < 1 \ arctan x < pi/4" by (metis arctan_less_iff arctan_one) lemma arctan_less_pi4_neg: "-1 < x \ -(pi/4) < arctan x" by (metis arctan_less_iff arctan_minus arctan_one) lemma arctan_less_pi4: "\x\ < 1 \ \arctan x\ < pi/4" by (metis abs_less_iff arctan_less_pi4_pos arctan_minus) lemma arctan_le_pi4: "\x\ \ 1 \ \arctan x\ \ pi/4" by (metis abs_le_iff arctan_le_iff arctan_minus arctan_one) lemma abs_arctan: "\arctan x\ = arctan \x\" by (simp add: abs_if arctan_minus) lemma arctan_add_raw: assumes "\arctan x + arctan y\ < pi/2" shows "arctan x + arctan y = arctan((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) show 12: "- (pi / 2) < arctan x + arctan y" "arctan x + arctan y < pi / 2" using assms by linarith+ show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" using cos_gt_zero_pi [OF 12] by (simp add: arctan tan_add) qed lemma arctan_inverse: assumes "0 < x" shows "arctan(inverse x) = pi/2 - arctan x" proof - have "arctan(inverse x) = arctan(inverse(tan(arctan x)))" by (simp add: arctan) also have "... = arctan (tan (pi / 2 - arctan x))" by (simp add: tan_cot) also have "... = pi/2 - arctan x" proof - have "0 < pi - arctan x" using arctan_ubound [of x] pi_gt_zero by linarith with assms show ?thesis by (simp add: Transcendental.arctan_tan) qed finally show ?thesis . qed lemma arctan_add_small: assumes "\x * y\ < 1" shows "(arctan x + arctan y = arctan((x + y) / (1 - x * y)))" proof (cases "x = 0 \ y = 0") case True then show ?thesis by auto next case False then have *: "\arctan x\ < pi / 2 - \arctan y\" using assms apply (auto simp add: abs_arctan arctan_inverse [symmetric] arctan_less_iff) apply (simp add: field_split_simps abs_mult) done show ?thesis apply (rule arctan_add_raw) using * by linarith qed lemma abs_arctan_le: fixes x::real shows "\arctan x\ \ \x\" proof - have 1: "\x. x \ \ \ cmod (inverse (1 + x\<^sup>2)) \ 1" by (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square) have "cmod (Arctan w - Arctan z) \ 1 * cmod (w-z)" if "w \ \" "z \ \" for w z apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1]) apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan]) using 1 that apply (auto simp: Reals_def) done then have "cmod (Arctan (of_real x) - Arctan 0) \ 1 * cmod (of_real x -0)" using Reals_0 Reals_of_real by blast then show ?thesis by (simp add: Arctan_of_real) qed lemma arctan_le_self: "0 \ x \ arctan x \ x" by (metis abs_arctan_le abs_of_nonneg zero_le_arctan_iff) lemma abs_tan_ge: "\x\ < pi/2 \ \x\ \ \tan x\" by (metis abs_arctan_le abs_less_iff arctan_tan minus_less_iff) lemma arctan_bounds: assumes "0 \ x" "x < 1" shows arctan_lower_bound: "(\k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \ arctan x" (is "(\k<_. (- 1)^ k * ?a k) \ _") and arctan_upper_bound: "arctan x \ (\k<2 * n + 1. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" proof - have tendsto_zero: "?a \ 0" proof (rule tendsto_eq_rhs) show "(\k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \ 0 * 0" using assms by (intro tendsto_mult real_tendsto_divide_at_top) (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top) qed simp have nonneg: "0 \ ?a n" for n by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms) have le: "?a (Suc n) \ ?a n" for n by (rule mult_mono[OF _ power_decreasing]) (auto simp: field_split_simps assms less_imp_le) from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n] summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n] assms show "(\k<2*n. (- 1)^ k * ?a k) \ arctan x" "arctan x \ (\k<2 * n + 1. (- 1)^ k * ?a k)" by (auto simp: arctan_series) qed subsection\<^marker>\tag unimportant\ \Bounds on pi using real arctangent\ lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)" using machin by simp lemma pi_approx: "3.141592653588 \ pi" "pi \ 3.1415926535899" unfolding pi_machin using arctan_bounds[of "1/5" 4] arctan_bounds[of "1/239" 4] by (simp_all add: eval_nat_numeral) lemma pi_gt3: "pi > 3" using pi_approx by simp subsection\Inverse Sine\ definition\<^marker>\tag important\ Arcsin :: "complex \ complex" where "Arcsin \ \z. -\ * Ln(\ * z + csqrt(1 - z\<^sup>2))" lemma Arcsin_body_lemma: "\ * z + csqrt(1 - z\<^sup>2) \ 0" using power2_csqrt [of "1 - z\<^sup>2"] apply auto by (metis complex_i_mult_minus diff_add_cancel diff_minus_eq_add diff_self mult.assoc mult.left_commute numeral_One power2_csqrt power2_eq_square zero_neq_numeral) lemma Arcsin_range_lemma: "\Re z\ < 1 \ 0 < Re(\ * z + csqrt(1 - z\<^sup>2))" using Complex.cmod_power2 [of z, symmetric] by (simp add: real_less_rsqrt algebra_simps Re_power2 cmod_square_less_1_plus) lemma Re_Arcsin: "Re(Arcsin z) = Im (Ln (\ * z + csqrt(1 - z\<^sup>2)))" by (simp add: Arcsin_def) lemma Im_Arcsin: "Im(Arcsin z) = - ln (cmod (\ * z + csqrt (1 - z\<^sup>2)))" by (simp add: Arcsin_def Arcsin_body_lemma) lemma one_minus_z2_notin_nonpos_Reals: assumes "(Im z = 0 \ \Re z\ < 1)" shows "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" using assms apply (auto simp: complex_nonpos_Reals_iff Re_power2 Im_power2) using power2_less_0 [of "Im z"] apply force using abs_square_less_1 not_le by blast lemma isCont_Arcsin_lemma: assumes le0: "Re (\ * z + csqrt (1 - z\<^sup>2)) \ 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp: cmod_def abs_square_less_1 [symmetric]) next case False have leim: "(cmod (1 - z\<^sup>2) + (1 - Re (z\<^sup>2))) / 2 \ (Im z)\<^sup>2" using le0 sqrt_le_D by fastforce have neq: "(cmod z)\<^sup>2 \ 1 + cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 = 1 + sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2" using leim cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1] by (simp add: norm_power Re_power2 norm_minus_commute [of 1]) ultimately show False by (simp add: Re_power2 Im_power2 cmod_power2) qed lemma isCont_Arcsin: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arcsin z" proof - have 1: "\ * z + csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff) have 2: "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" by (simp add: one_minus_z2_notin_nonpos_Reals assms) show ?thesis using assms unfolding Arcsin_def by (intro isCont_Ln' isCont_csqrt' continuous_intros 1 2) qed lemma isCont_Arcsin' [simp]: shows "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arcsin (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arcsin]) lemma sin_Arcsin [simp]: "sin(Arcsin z) = z" proof - have "\*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \ (\*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ (\*z) + csqrt (1 - z\<^sup>2) = 0" by (metis Arcsin_body_lemma distrib_right no_zero_divisors zero_neq_numeral) ultimately show ?thesis apply (simp add: sin_exp_eq Arcsin_def Arcsin_body_lemma exp_minus divide_simps) apply (simp add: algebra_simps) apply (simp add: power2_eq_square [symmetric] algebra_simps) done qed lemma Re_eq_pihalf_lemma: "\Re z\ = pi/2 \ Im z = 0 \ Re ((exp (\*z) + inverse (exp (\*z))) / 2) = 0 \ 0 \ Im ((exp (\*z) + inverse (exp (\*z))) / 2)" apply (simp add: cos_i_times [symmetric] Re_cos Im_cos abs_if del: eq_divide_eq_numeral1) by (metis cos_minus cos_pi_half) lemma Re_less_pihalf_lemma: assumes "\Re z\ < pi / 2" shows "0 < Re ((exp (\*z) + inverse (exp (\*z))) / 2)" proof - have "0 < cos (Re z)" using assms using cos_gt_zero_pi by auto then show ?thesis by (simp add: cos_i_times [symmetric] Re_cos Im_cos add_pos_pos) qed lemma Arcsin_sin: assumes "\Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)" shows "Arcsin(sin z) = z" proof - have "Arcsin(sin z) = - (\ * Ln (csqrt (1 - (\ * (exp (\*z) - inverse (exp (\*z))))\<^sup>2 / 4) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: sin_exp_eq Arcsin_def exp_minus power_divide) also have "... = - (\ * Ln (csqrt (((exp (\*z) + inverse (exp (\*z)))/2)\<^sup>2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: field_simps power2_eq_square) also have "... = - (\ * Ln (((exp (\*z) + inverse (exp (\*z)))/2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" apply (subst csqrt_square) using assms Re_eq_pihalf_lemma Re_less_pihalf_lemma apply auto done also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms by (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm) finally show ?thesis . qed lemma Arcsin_unique: "\sin z = w; \Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)\ \ Arcsin w = z" by (metis Arcsin_sin) lemma Arcsin_0 [simp]: "Arcsin 0 = 0" by (metis Arcsin_sin norm_zero pi_half_gt_zero real_norm_def sin_zero zero_complex.simps(1)) lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2" by (metis Arcsin_sin Im_complex_of_real Re_complex_of_real numeral_One of_real_numeral pi_half_ge_zero real_sqrt_abs real_sqrt_pow2 real_sqrt_power sin_of_real sin_pi_half) lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)" by (metis Arcsin_1 Arcsin_sin Im_complex_of_real Re_complex_of_real abs_of_nonneg of_real_minus pi_half_ge_zero power2_minus real_sqrt_abs sin_Arcsin sin_minus) lemma has_field_derivative_Arcsin: assumes "Im z = 0 \ \Re z\ < 1" shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)" proof - have "(sin (Arcsin z))\<^sup>2 \ 1" using assms one_minus_z2_notin_nonpos_Reals by force then have "cos (Arcsin z) \ 0" by (metis diff_0_right power_zero_numeral sin_squared_eq) then show ?thesis by (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) (auto intro: isCont_Arcsin assms) qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable at z" using field_differentiable_def has_field_derivative_Arcsin by blast lemma field_differentiable_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable (at z within s)" using field_differentiable_at_Arcsin field_differentiable_within_subset by blast lemma continuous_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arcsin" using continuous_at_imp_continuous_within isCont_Arcsin by blast lemma continuous_on_Arcsin [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arcsin" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arcsin: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arcsin holomorphic_on s" by (simp add: field_differentiable_within_Arcsin holomorphic_on_def) subsection\Inverse Cosine\ definition\<^marker>\tag important\ Arccos :: "complex \ complex" where "Arccos \ \z. -\ * Ln(z + \ * csqrt(1 - z\<^sup>2))" lemma Arccos_range_lemma: "\Re z\ < 1 \ 0 < Im(z + \ * csqrt(1 - z\<^sup>2))" using Arcsin_range_lemma [of "-z"] by simp lemma Arccos_body_lemma: "z + \ * csqrt(1 - z\<^sup>2) \ 0" using Arcsin_body_lemma [of z] by (metis Arcsin_body_lemma complex_i_mult_minus diff_minus_eq_add power2_minus right_minus_eq) lemma Re_Arccos: "Re(Arccos z) = Im (Ln (z + \ * csqrt(1 - z\<^sup>2)))" by (simp add: Arccos_def) lemma Im_Arccos: "Im(Arccos z) = - ln (cmod (z + \ * csqrt (1 - z\<^sup>2)))" by (simp add: Arccos_def Arccos_body_lemma) text\A very tricky argument to find!\ lemma isCont_Arccos_lemma: assumes eq0: "Im (z + \ * csqrt (1 - z\<^sup>2)) = 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp add: cmod_def abs_square_less_1 [symmetric]) next case False have Imz: "Im z = - sqrt ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" using eq0 abs_Re_le_cmod [of "1-z\<^sup>2"] by (simp add: Re_power2 algebra_simps) have "(cmod z)\<^sup>2 - 1 \ cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 - 1 = sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have "(Im z)\<^sup>2 = ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" apply (subst Imz) using abs_Re_le_cmod [of "1-z\<^sup>2"] apply (simp add: Re_power2) done ultimately show False by (simp add: cmod_power2) qed lemma isCont_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arccos z" proof - have "z + \ * csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms) with assms show ?thesis apply (simp add: Arccos_def) apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+ apply (simp_all add: one_minus_z2_notin_nonpos_Reals assms) done qed lemma isCont_Arccos' [simp]: shows "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arccos (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arccos]) lemma cos_Arccos [simp]: "cos(Arccos z) = z" proof - have "z*2 + \ * (2 * csqrt (1 - z\<^sup>2)) = 0 \ z*2 + \ * csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ z + \ * csqrt (1 - z\<^sup>2) = 0" by (metis distrib_right mult_eq_0_iff zero_neq_numeral) ultimately show ?thesis apply (simp add: cos_exp_eq Arccos_def Arccos_body_lemma exp_minus field_simps) apply (simp add: power2_eq_square [symmetric]) done qed lemma Arccos_cos: assumes "0 < Re z & Re z < pi \ Re z = 0 & 0 \ Im z \ Re z = pi & Im z \ 0" shows "Arccos(cos z) = z" proof - have *: "((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z))) = sin z" by (simp add: sin_exp_eq exp_minus field_simps power2_eq_square) have "1 - (exp (\ * z) + inverse (exp (\ * z)))\<^sup>2 / 4 = ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2" by (simp add: field_simps power2_eq_square) then have "Arccos(cos z) = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * csqrt (((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2)))" by (simp add: cos_exp_eq Arccos_def exp_minus power_divide) also have "... = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))))" apply (subst csqrt_square) using assms Re_sin_pos [of z] Im_sin_nonneg [of z] Im_sin_nonneg2 [of z] apply (auto simp: * Re_sin Im_sin) done also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms apply (subst Complex_Transcendental.Ln_exp, auto) done finally show ?thesis . qed lemma Arccos_unique: "\cos z = w; 0 < Re z \ Re z < pi \ Re z = 0 \ 0 \ Im z \ Re z = pi \ Im z \ 0\ \ Arccos w = z" using Arccos_cos by blast lemma Arccos_0 [simp]: "Arccos 0 = pi/2" by (rule Arccos_unique) auto lemma Arccos_1 [simp]: "Arccos 1 = 0" by (rule Arccos_unique) auto lemma Arccos_minus1: "Arccos(-1) = pi" by (rule Arccos_unique) auto lemma has_field_derivative_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "(Arccos has_field_derivative - inverse(sin(Arccos z))) (at z)" proof - have "x\<^sup>2 \ -1" for x::real by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x__]^2)))))") with assms have "(cos (Arccos z))\<^sup>2 \ 1" by (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1) then have "- sin (Arccos z) \ 0" by (metis cos_squared_eq diff_0_right mult_zero_left neg_0_equal_iff_equal power2_eq_square) then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)" by (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]]) (auto intro: isCont_Arccos assms) then show ?thesis by simp qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable at z" using field_differentiable_def has_field_derivative_Arccos by blast lemma field_differentiable_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable (at z within s)" using field_differentiable_at_Arccos field_differentiable_within_subset by blast lemma continuous_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arccos" using continuous_at_imp_continuous_within isCont_Arccos by blast lemma continuous_on_Arccos [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arccos" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arccos: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arccos holomorphic_on s" by (simp add: field_differentiable_within_Arccos holomorphic_on_def) subsection\<^marker>\tag unimportant\\Upper and Lower Bounds for Inverse Sine and Cosine\ lemma Arcsin_bounds: "\Re z\ < 1 \ \Re(Arcsin z)\ < pi/2" unfolding Re_Arcsin by (blast intro: Re_Ln_pos_lt_imp Arcsin_range_lemma) lemma Arccos_bounds: "\Re z\ < 1 \ 0 < Re(Arccos z) \ Re(Arccos z) < pi" unfolding Re_Arccos by (blast intro!: Im_Ln_pos_lt_imp Arccos_range_lemma) lemma Re_Arccos_bounds: "-pi < Re(Arccos z) \ Re(Arccos z) \ pi" unfolding Re_Arccos by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arccos_body_lemma) lemma Re_Arccos_bound: "\Re(Arccos z)\ \ pi" by (meson Re_Arccos_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma Im_Arccos_bound: "\Im (Arccos w)\ \ cmod w" proof - have "(Im (Arccos w))\<^sup>2 \ (cmod (cos (Arccos w)))\<^sup>2 - (cos (Re (Arccos w)))\<^sup>2" using norm_cos_squared [of "Arccos w"] real_le_abs_sinh [of "Im (Arccos w)"] apply (simp only: abs_le_square_iff) apply (simp add: field_split_simps) done also have "... \ (cmod w)\<^sup>2" by (auto simp: cmod_power2) finally show ?thesis using abs_le_square_iff by force qed lemma Re_Arcsin_bounds: "-pi < Re(Arcsin z) & Re(Arcsin z) \ pi" unfolding Re_Arcsin by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arcsin_body_lemma) lemma Re_Arcsin_bound: "\Re(Arcsin z)\ \ pi" by (meson Re_Arcsin_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma norm_Arccos_bounded: fixes w :: complex shows "norm (Arccos w) \ pi + norm w" proof - have Re: "(Re (Arccos w))\<^sup>2 \ pi\<^sup>2" "(Im (Arccos w))\<^sup>2 \ (cmod w)\<^sup>2" using Re_Arccos_bound [of w] Im_Arccos_bound [of w] abs_le_square_iff by force+ have "Arccos w \ Arccos w \ pi\<^sup>2 + (cmod w)\<^sup>2" using Re by (simp add: dot_square_norm cmod_power2 [of "Arccos w"]) then have "cmod (Arccos w) \ pi + cmod (cos (Arccos w))" apply (simp add: norm_le_square) by (metis dot_square_norm norm_ge_zero norm_le_square pi_ge_zero triangle_lemma) then show "cmod (Arccos w) \ pi + cmod w" by auto qed subsection\<^marker>\tag unimportant\\Interrelations between Arcsin and Arccos\ lemma cos_Arcsin_nonzero: assumes "z\<^sup>2 \ 1" shows "cos(Arcsin z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)" by (simp add: power_mult_distrib algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ z\<^sup>2 - 1" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (z\<^sup>2 - 1)\<^sup>2" by simp then have "z\<^sup>2 * (z\<^sup>2 - 1) = (z\<^sup>2 - 1)*(z\<^sup>2 - 1)" using eq power2_eq_square by auto then show False using assms by simp qed then have "1 + \ * z * (csqrt (1 - z * z)) \ z\<^sup>2" by (metis add_minus_cancel power2_eq_square uminus_add_conv_diff) then have "2*(1 + \ * z * (csqrt (1 - z * z))) \ 2*z\<^sup>2" (*FIXME cancel_numeral_factor*) by (metis mult_cancel_left zero_neq_numeral) then have "(\ * z + csqrt (1 - z\<^sup>2))\<^sup>2 \ -1" using assms apply (auto simp: power2_sum) apply (simp add: power2_eq_square algebra_simps) done then show ?thesis apply (simp add: cos_exp_eq Arcsin_def exp_minus) apply (simp add: divide_simps Arcsin_body_lemma) apply (metis add.commute minus_unique power2_eq_square) done qed lemma sin_Arccos_nonzero: assumes "z\<^sup>2 \ 1" shows "sin(Arccos z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)" by (simp add: power_mult_distrib algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ 1 - z\<^sup>2" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (1 - z\<^sup>2)\<^sup>2" by simp then have "-(z\<^sup>2) * (1 - z\<^sup>2) = (1 - z\<^sup>2)*(1 - z\<^sup>2)" using eq power2_eq_square by auto then have "-(z\<^sup>2) = (1 - z\<^sup>2)" using assms by (metis add.commute add.right_neutral diff_add_cancel mult_right_cancel) then show False using assms by simp qed then have "z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2)) \ 1" by (simp add: algebra_simps) then have "2*(z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2))) \ 2*1" by (metis mult_cancel_left2 zero_neq_numeral) (*FIXME cancel_numeral_factor*) then have "(z + \ * csqrt (1 - z\<^sup>2))\<^sup>2 \ 1" using assms apply (auto simp: Power.comm_semiring_1_class.power2_sum power_mult_distrib) apply (simp add: power2_eq_square algebra_simps) done then show ?thesis apply (simp add: sin_exp_eq Arccos_def exp_minus) apply (simp add: divide_simps Arccos_body_lemma) apply (simp add: power2_eq_square) done qed lemma cos_sin_csqrt: assumes "0 < cos(Re z) \ cos(Re z) = 0 \ Im z * sin(Re z) \ 0" shows "cos z = csqrt(1 - (sin z)\<^sup>2)" apply (rule csqrt_unique [THEN sym]) apply (simp add: cos_squared_eq) using assms apply (auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff) done lemma sin_cos_csqrt: assumes "0 < sin(Re z) \ sin(Re z) = 0 \ 0 \ Im z * cos(Re z)" shows "sin z = csqrt(1 - (cos z)\<^sup>2)" apply (rule csqrt_unique [THEN sym]) apply (simp add: sin_squared_eq) using assms apply (auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff) done lemma Arcsin_Arccos_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arcsin z = Arccos(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma Arccos_Arcsin_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arccos z = Arcsin(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma sin_Arccos: "0 < Re z | Re z = 0 & 0 \ Im z \ sin(Arccos z) = csqrt(1 - z\<^sup>2)" by (simp add: Arccos_Arcsin_csqrt_pos) lemma cos_Arcsin: "0 < Re z | Re z = 0 & 0 \ Im z \ cos(Arcsin z) = csqrt(1 - z\<^sup>2)" by (simp add: Arcsin_Arccos_csqrt_pos) subsection\<^marker>\tag unimportant\\Relationship with Arcsin on the Real Numbers\ lemma Im_Arcsin_of_real: assumes "\x\ \ 1" shows "Im (Arcsin (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arcsin exp_minus) qed corollary\<^marker>\tag unimportant\ Arcsin_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arcsin z \ \" by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arcsin_eq_Re_Arcsin: assumes "\x\ \ 1" shows "arcsin x = Re (Arcsin (of_real x))" unfolding arcsin_def proof (rule the_equality, safe) show "- (pi / 2) \ Re (Arcsin (complex_of_real x))" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "Re (Arcsin (complex_of_real x)) \ pi / 2" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "sin (Re (Arcsin (complex_of_real x))) = x" using Re_sin [of "Arcsin (of_real x)"] Arcsin_body_lemma [of "of_real x"] by (simp add: Im_Arcsin_of_real assms) next fix x' assume "- (pi / 2) \ x'" "x' \ pi / 2" "x = sin x'" then show "x' = Re (Arcsin (complex_of_real (sin x')))" apply (simp add: sin_of_real [symmetric]) apply (subst Arcsin_sin) apply (auto simp: ) done qed lemma of_real_arcsin: "\x\ \ 1 \ of_real(arcsin x) = Arcsin(of_real x)" by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Relationship with Arccos on the Real Numbers\ lemma Im_Arccos_of_real: assumes "\x\ \ 1" shows "Im (Arccos (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arccos exp_minus) qed corollary\<^marker>\tag unimportant\ Arccos_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arccos z \ \" by (metis Im_Arccos_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arccos_eq_Re_Arccos: assumes "\x\ \ 1" shows "arccos x = Re (Arccos (of_real x))" unfolding arccos_def proof (rule the_equality, safe) show "0 \ Re (Arccos (complex_of_real x))" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "Re (Arccos (complex_of_real x)) \ pi" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "cos (Re (Arccos (complex_of_real x))) = x" using Re_cos [of "Arccos (of_real x)"] Arccos_body_lemma [of "of_real x"] by (simp add: Im_Arccos_of_real assms) next fix x' assume "0 \ x'" "x' \ pi" "x = cos x'" then show "x' = Re (Arccos (complex_of_real (cos x')))" apply (simp add: cos_of_real [symmetric]) apply (subst Arccos_cos) apply (auto simp: ) done qed lemma of_real_arccos: "\x\ \ 1 \ of_real(arccos x) = Arccos(of_real x)" by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Some interrelationships among the real inverse trig functions\ lemma arccos_arctan: assumes "-1 < x" "x < 1" shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))" proof - have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0" proof (rule sin_eq_0_pi) show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)" using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi" using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0" using assms by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan power2_eq_square square_eq_1_iff) qed then show ?thesis by simp qed lemma arcsin_plus_arccos: assumes "-1 \ x" "x \ 1" shows "arcsin x + arccos x = pi/2" proof - have "arcsin x = pi/2 - arccos x" apply (rule sin_inj_pi) using assms arcsin [OF assms] arccos [OF assms] apply (auto simp: algebra_simps sin_diff) done then show ?thesis by (simp add: algebra_simps) qed lemma arcsin_arccos_eq: "-1 \ x \ x \ 1 \ arcsin x = pi/2 - arccos x" using arcsin_plus_arccos by force lemma arccos_arcsin_eq: "-1 \ x \ x \ 1 \ arccos x = pi/2 - arcsin x" using arcsin_plus_arccos by force lemma arcsin_arctan: "-1 < x \ x < 1 \ arcsin x = arctan(x / sqrt(1 - x\<^sup>2))" by (simp add: arccos_arctan arcsin_arccos_eq) lemma csqrt_1_diff_eq: "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg) lemma arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arcsin_Arccos_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arcsin_arccos_sqrt_neg: "-1 \ x \ x \ 0 \ arcsin x = -arccos(sqrt(1 - x\<^sup>2))" using arcsin_arccos_sqrt_pos [of "-x"] by (simp add: arcsin_minus) lemma arccos_arcsin_sqrt_pos: "0 \ x \ x \ 1 \ arccos x = arcsin(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arccos_Arcsin_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arccos_arcsin_sqrt_neg: "-1 \ x \ x \ 0 \ arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))" using arccos_arcsin_sqrt_pos [of "-x"] by (simp add: arccos_minus) subsection\<^marker>\tag unimportant\\Continuity results for arcsin and arccos\ lemma continuous_on_Arcsin_real [continuous_intros]: "continuous_on {w \ \. \Re w\ \ 1} Arcsin" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arcsin (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arcsin (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arcsin_eq_Re_Arcsin) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arcsin [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arcsin_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arcsin" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arcsin_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma continuous_on_Arccos_real: "continuous_on {w \ \. \Re w\ \ 1} Arccos" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arccos (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arccos (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arccos_eq_Re_Arccos) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arccos [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arccos_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arccos" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arccos_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma sinh_ln_complex: "x \ 0 \ sinh (ln x :: complex) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus scaleR_conv_of_real exp_of_real) lemma cosh_ln_complex: "x \ 0 \ cosh (ln x :: complex) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus scaleR_conv_of_real) lemma tanh_ln_complex: "x \ 0 \ tanh (ln x :: complex) = (x ^ 2 - 1) / (x ^ 2 + 1)" by (simp add: tanh_def sinh_ln_complex cosh_ln_complex divide_simps power2_eq_square) subsection\Roots of unity\ theorem complex_root_unity: fixes j::nat assumes "n \ 0" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n)^n = 1" proof - have *: "of_nat j * (complex_of_real pi * 2) = complex_of_real (2 * real j * pi)" by (simp add: of_real_numeral) then show ?thesis apply (simp add: exp_of_nat_mult [symmetric] mult_ac exp_Euler) apply (simp only: * cos_of_real sin_of_real) apply (simp add: ) done qed lemma complex_root_unity_eq: fixes j::nat and k::nat assumes "1 \ n" shows "(exp(2 * of_real pi * \ * of_nat j / of_nat n) = exp(2 * of_real pi * \ * of_nat k / of_nat n) \ j mod n = k mod n)" proof - have "(\z::int. \ * (of_nat j * (of_real pi * 2)) = \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ (\z::int. of_nat j * (\ * (of_real pi * 2)) = (of_nat k + of_nat n * of_int z) * (\ * (of_real pi * 2)))" by (simp add: algebra_simps) also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * (of_int z :: complex))" by simp also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * z)" apply (rule HOL.iff_exI) apply (auto simp: ) using of_int_eq_iff apply fastforce by (metis of_int_add of_int_mult of_int_of_nat_eq) also have "... \ int j mod int n = int k mod int n" by (auto simp: mod_eq_dvd_iff dvd_def algebra_simps) also have "... \ j mod n = k mod n" by (metis of_nat_eq_iff zmod_int) finally have "(\z. \ * (of_nat j * (of_real pi * 2)) = \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ j mod n = k mod n" . note * = this show ?thesis using assms by (simp add: exp_eq field_split_simps mult_ac of_real_numeral *) qed corollary bij_betw_roots_unity: "bij_betw (\j. exp(2 * of_real pi * \ * of_nat j / of_nat n)) {.. * of_nat j / of_nat n) | j. j < n}" by (auto simp: bij_betw_def inj_on_def complex_root_unity_eq) lemma complex_root_unity_eq_1: fixes j::nat and k::nat assumes "1 \ n" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n) = 1 \ n dvd j" proof - have "1 = exp(2 * of_real pi * \ * (of_nat n / of_nat n))" using assms by simp then have "exp(2 * of_real pi * \ * (of_nat j / of_nat n)) = 1 \ j mod n = n mod n" using complex_root_unity_eq [of n j n] assms by simp then show ?thesis by auto qed lemma finite_complex_roots_unity_explicit: "finite {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n}" by simp lemma card_complex_roots_unity_explicit: "card {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n} = n" by (simp add: Finite_Set.bij_betw_same_card [OF bij_betw_roots_unity, symmetric]) lemma complex_roots_unity: assumes "1 \ n" shows "{z::complex. z^n = 1} = {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n}" apply (rule Finite_Set.card_seteq [symmetric]) using assms apply (auto simp: card_complex_roots_unity_explicit finite_roots_unity complex_root_unity card_roots_unity) done lemma card_complex_roots_unity: "1 \ n \ card {z::complex. z^n = 1} = n" by (simp add: card_complex_roots_unity_explicit complex_roots_unity) lemma complex_not_root_unity: "1 \ n \ \u::complex. norm u = 1 \ u^n \ 1" apply (rule_tac x="exp (of_real pi * \ * of_real (1 / n))" in exI) apply (auto simp: Re_complex_div_eq_0 exp_of_nat_mult [symmetric] mult_ac exp_Euler) done end diff --git a/src/HOL/Analysis/Great_Picard.thy b/src/HOL/Analysis/Great_Picard.thy --- a/src/HOL/Analysis/Great_Picard.thy +++ b/src/HOL/Analysis/Great_Picard.thy @@ -1,1848 +1,1848 @@ 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 DERIV_transform_within_open [where S = "ball z0 r"]]) + 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/Riemann_Mapping.thy b/src/HOL/Analysis/Riemann_Mapping.thy --- a/src/HOL/Analysis/Riemann_Mapping.thy +++ b/src/HOL/Analysis/Riemann_Mapping.thy @@ -1,1489 +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 DERIV_transform_within_open [OF DERIV_const]]) + 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 DERIV_transform_within_open [OF DERIV_chain]) + 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/Analysis/Topology_Euclidean_Space.thy b/src/HOL/Analysis/Topology_Euclidean_Space.thy --- a/src/HOL/Analysis/Topology_Euclidean_Space.thy +++ b/src/HOL/Analysis/Topology_Euclidean_Space.thy @@ -1,2497 +1,2509 @@ (* Author: L C Paulson, University of Cambridge Author: Amine Chaieb, University of Cambridge Author: Robert Himmelmann, TU Muenchen Author: Brian Huffman, Portland State University *) chapter \Vector Analysis\ theory Topology_Euclidean_Space imports Elementary_Normed_Spaces Linear_Algebra Norm_Arith begin section \Elementary Topology in Euclidean Space\ lemma euclidean_dist_l2: fixes x y :: "'a :: euclidean_space" shows "dist x y = L2_set (\i. dist (x \ i) (y \ i)) Basis" unfolding dist_norm norm_eq_sqrt_inner L2_set_def by (subst euclidean_inner) (simp add: power2_eq_square inner_diff_left) lemma norm_nth_le: "norm (x \ i) \ norm x" if "i \ Basis" proof - have "(x \ i)\<^sup>2 = (\i\{i}. (x \ i)\<^sup>2)" by simp also have "\ \ (\i\Basis. (x \ i)\<^sup>2)" by (intro sum_mono2) (auto simp: that) finally show ?thesis unfolding norm_conv_dist euclidean_dist_l2[of x] L2_set_def by (auto intro!: real_le_rsqrt) qed subsection \Continuity of the representation WRT an orthogonal basis\ lemma orthogonal_Basis: "pairwise orthogonal Basis" by (simp add: inner_not_same_Basis orthogonal_def pairwise_def) lemma representation_bound: fixes B :: "'N::real_inner set" assumes "finite B" "independent B" "b \ B" and orth: "pairwise orthogonal B" obtains m where "m > 0" "\x. x \ span B \ \representation B x b\ \ m * norm x" proof fix x assume x: "x \ span B" have "b \ 0" using \independent B\ \b \ B\ dependent_zero by blast have [simp]: "b \ b' = (if b' = b then (norm b)\<^sup>2 else 0)" if "b \ B" "b' \ B" for b b' using orth by (simp add: orthogonal_def pairwise_def norm_eq_sqrt_inner that) have "norm x = norm (\b\B. representation B x b *\<^sub>R b)" using real_vector.sum_representation_eq [OF \independent B\ x \finite B\ order_refl] by simp also have "\ = sqrt ((\b\B. representation B x b *\<^sub>R b) \ (\b\B. representation B x b *\<^sub>R b))" by (simp add: norm_eq_sqrt_inner) also have "\ = sqrt (\b\B. (representation B x b *\<^sub>R b) \ (representation B x b *\<^sub>R b))" using \finite B\ by (simp add: inner_sum_left inner_sum_right if_distrib [of "\x. _ * x"] cong: if_cong sum.cong_simp) also have "\ = sqrt (\b\B. (norm (representation B x b *\<^sub>R b))\<^sup>2)" by (simp add: mult.commute mult.left_commute power2_eq_square) also have "\ = sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" by (simp add: norm_mult power_mult_distrib) finally have "norm x = sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" . moreover have "sqrt ((representation B x b)\<^sup>2 * (norm b)\<^sup>2) \ sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" using \b \ B\ \finite B\ by (auto intro: member_le_sum) then have "\representation B x b\ \ (1 / norm b) * sqrt (\b\B. (representation B x b)\<^sup>2 * (norm b)\<^sup>2)" using \b \ 0\ by (simp add: field_split_simps real_sqrt_mult del: real_sqrt_le_iff) ultimately show "\representation B x b\ \ (1 / norm b) * norm x" by simp next show "0 < 1 / norm b" using \independent B\ \b \ B\ dependent_zero by auto qed lemma continuous_on_representation: fixes B :: "'N::euclidean_space set" assumes "finite B" "independent B" "b \ B" "pairwise orthogonal B" shows "continuous_on (span B) (\x. representation B x b)" proof show "\d>0. \x'\span B. dist x' x < d \ dist (representation B x' b) (representation B x b) \ e" if "e > 0" "x \ span B" for x e proof - obtain m where "m > 0" and m: "\x. x \ span B \ \representation B x b\ \ m * norm x" using assms representation_bound by blast show ?thesis unfolding dist_norm proof (intro exI conjI ballI impI) show "e/m > 0" by (simp add: \e > 0\ \m > 0\) show "norm (representation B x' b - representation B x b) \ e" if x': "x' \ span B" and less: "norm (x'-x) < e/m" for x' proof - have "\representation B (x'-x) b\ \ m * norm (x'-x)" using m [of "x'-x"] \x \ span B\ span_diff x' by blast also have "\ < e" by (metis \m > 0\ less mult.commute pos_less_divide_eq) finally have "\representation B (x'-x) b\ \ e" by simp then show ?thesis by (simp add: \x \ span B\ \independent B\ representation_diff x') qed qed qed qed subsection\<^marker>\tag unimportant\\Balls in Euclidean Space\ lemma cball_subset_cball_iff: fixes a :: "'a :: euclidean_space" shows "cball a r \ cball a' r' \ dist a a' + r \ r' \ r < 0" (is "?lhs \ ?rhs") proof assume ?lhs then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next case False then have [simp]: "r \ 0" by simp have "norm (a - a') + r \ r'" proof (cases "a = a'") case True then show ?thesis using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] by (force simp: SOME_Basis dist_norm) next case False have "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = norm (a' - a - (r / norm (a - a')) *\<^sub>R (a - a'))" by (simp add: algebra_simps) also have "... = norm ((-1 - (r / norm (a - a'))) *\<^sub>R (a - a'))" by (simp add: algebra_simps) also from \a \ a'\ have "... = \- norm (a - a') - r\" by simp (simp add: field_simps) finally have [simp]: "norm (a' - (a + (r / norm (a - a')) *\<^sub>R (a - a'))) = \norm (a - a') + r\" by linarith from \a \ a'\ show ?thesis using subsetD [where c = "a' + (1 + r / norm(a - a')) *\<^sub>R (a - a')", OF \?lhs\] by (simp add: dist_norm scaleR_add_left) qed then show ?rhs by (simp add: dist_norm) qed qed metric lemma cball_subset_ball_iff: "cball a r \ ball a' r' \ dist a a' + r < r' \ r < 0" (is "?lhs \ ?rhs") for a :: "'a::euclidean_space" proof assume ?lhs then show ?rhs proof (cases "r < 0") case True then show ?rhs by simp next case False then have [simp]: "r \ 0" by simp have "norm (a - a') + r < r'" proof (cases "a = a'") case True then show ?thesis using subsetD [where c = "a + r *\<^sub>R (SOME i. i \ Basis)", OF \?lhs\] subsetD [where c = a, OF \?lhs\] by (force simp: SOME_Basis dist_norm) next case False have False if "norm (a - a') + r \ r'" proof - from that have "\r' - norm (a - a')\ \ r" by (simp split: abs_split) (metis \0 \ r\ \?lhs\ centre_in_cball dist_commute dist_norm less_asym mem_ball subset_eq) then show ?thesis using subsetD [where c = "a + (r' / norm(a - a') - 1) *\<^sub>R (a - a')", OF \?lhs\] \a \ a'\ apply (simp add: dist_norm) apply (simp add: scaleR_left_diff_distrib) apply (simp add: field_simps) done qed then show ?thesis by force qed then show ?rhs by (simp add: dist_norm) qed next assume ?rhs then show ?lhs by metric qed lemma ball_subset_cball_iff: "ball a r \ cball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") for a :: "'a::euclidean_space" proof (cases "r \ 0") case True then show ?thesis by metric next case False show ?thesis proof assume ?lhs then have "(cball a r \ cball a' r')" by (metis False closed_cball closure_ball closure_closed closure_mono not_less) with False show ?rhs by (fastforce iff: cball_subset_cball_iff) next assume ?rhs with False show ?lhs by metric qed qed lemma ball_subset_ball_iff: fixes a :: "'a :: euclidean_space" shows "ball a r \ ball a' r' \ dist a a' + r \ r' \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True then show ?thesis by metric next case False show ?thesis proof assume ?lhs then have "0 < r'" using False by metric then have "(cball a r \ cball a' r')" by (metis False\?lhs\ closure_ball closure_mono not_less) then show ?rhs using False cball_subset_cball_iff by fastforce qed metric qed lemma ball_eq_ball_iff: fixes x :: "'a :: euclidean_space" shows "ball x d = ball y e \ d \ 0 \ e \ 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (cases "d \ 0 \ e \ 0") case True with \?lhs\ show ?rhs by safe (simp_all only: ball_eq_empty [of y e, symmetric] ball_eq_empty [of x d, symmetric]) next case False with \?lhs\ show ?rhs apply (auto simp: set_eq_subset ball_subset_ball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next assume ?rhs then show ?lhs by (auto simp: set_eq_subset ball_subset_ball_iff) qed lemma cball_eq_cball_iff: fixes x :: "'a :: euclidean_space" shows "cball x d = cball y e \ d < 0 \ e < 0 \ x=y \ d=e" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs proof (cases "d < 0 \ e < 0") case True with \?lhs\ show ?rhs by safe (simp_all only: cball_eq_empty [of y e, symmetric] cball_eq_empty [of x d, symmetric]) next case False with \?lhs\ show ?rhs apply (auto simp: set_eq_subset cball_subset_cball_iff dist_norm norm_minus_commute algebra_simps) apply (metis add_le_same_cancel1 le_add_same_cancel1 norm_ge_zero norm_pths(2) order_trans) apply (metis add_increasing2 add_le_imp_le_right eq_iff norm_ge_zero) done qed next assume ?rhs then show ?lhs by (auto simp: set_eq_subset cball_subset_cball_iff) qed lemma ball_eq_cball_iff: fixes x :: "'a :: euclidean_space" shows "ball x d = cball y e \ d \ 0 \ e < 0" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs apply (auto simp: set_eq_subset ball_subset_cball_iff cball_subset_ball_iff algebra_simps) apply (metis add_increasing2 add_le_cancel_right add_less_same_cancel1 dist_not_less_zero less_le_trans zero_le_dist) apply (metis add_less_same_cancel1 dist_not_less_zero less_le_trans not_le) using \?lhs\ ball_eq_empty cball_eq_empty apply blast+ done next assume ?rhs then show ?lhs by auto qed lemma cball_eq_ball_iff: fixes x :: "'a :: euclidean_space" shows "cball x d = ball y e \ d < 0 \ e \ 0" using ball_eq_cball_iff by blast lemma finite_ball_avoid: fixes S :: "'a :: euclidean_space set" assumes "open S" "finite X" "p \ S" shows "\e>0. \w\ball p e. w\S \ (w\p \ w\X)" proof - obtain e1 where "0 < e1" and e1_b:"ball p e1 \ S" using open_contains_ball_eq[OF \open S\] assms by auto obtain e2 where "0 < e2" and "\x\X. x \ p \ e2 \ dist p x" using finite_set_avoid[OF \finite X\,of p] by auto hence "\w\ball p (min e1 e2). w\S \ (w\p \ w\X)" using e1_b by auto thus "\e>0. \w\ball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ \e1>0\ apply (rule_tac x="min e1 e2" in exI) by auto qed lemma finite_cball_avoid: fixes S :: "'a :: euclidean_space set" assumes "open S" "finite X" "p \ S" shows "\e>0. \w\cball p e. w\S \ (w\p \ w\X)" proof - obtain e1 where "e1>0" and e1: "\w\ball p e1. w\S \ (w\p \ w\X)" using finite_ball_avoid[OF assms] by auto define e2 where "e2 \ e1/2" have "e2>0" and "e2 < e1" unfolding e2_def using \e1>0\ by auto then have "cball p e2 \ ball p e1" by (subst cball_subset_ball_iff,auto) then show "\e>0. \w\cball p e. w \ S \ (w \ p \ w \ X)" using \e2>0\ e1 by auto qed lemma dim_cball: assumes "e > 0" shows "dim (cball (0 :: 'n::euclidean_space) e) = DIM('n)" proof - { fix x :: "'n::euclidean_space" define y where "y = (e / norm x) *\<^sub>R x" then have "y \ cball 0 e" using assms by auto moreover have *: "x = (norm x / e) *\<^sub>R y" using y_def assms by simp moreover from * have "x = (norm x/e) *\<^sub>R y" by auto ultimately have "x \ span (cball 0 e)" using span_scale[of y "cball 0 e" "norm x/e"] span_superset[of "cball 0 e"] by (simp add: span_base) } then have "span (cball 0 e) = (UNIV :: 'n::euclidean_space set)" by auto then show ?thesis using dim_span[of "cball (0 :: 'n::euclidean_space) e"] by (auto simp: dim_UNIV) qed subsection \Boxes\ abbreviation One :: "'a::euclidean_space" where "One \ \Basis" lemma One_non_0: assumes "One = (0::'a::euclidean_space)" shows False proof - have "dependent (Basis :: 'a set)" apply (simp add: dependent_finite) apply (rule_tac x="\i. 1" in exI) using SOME_Basis apply (auto simp: assms) done with independent_Basis show False by force qed corollary One_neq_0[iff]: "One \ 0" by (metis One_non_0) corollary Zero_neq_One[iff]: "0 \ One" by (metis One_non_0) definition\<^marker>\tag important\ (in euclidean_space) eucl_less (infix " (\i\Basis. a \ i < b \ i)" definition\<^marker>\tag important\ box_eucl_less: "box a b = {x. a x \tag important\ "cbox a b = {x. \i\Basis. a \ i \ x \ i \ x \ i \ b \ i}" lemma box_def: "box a b = {x. \i\Basis. a \ i < x \ i \ x \ i < b \ i}" and in_box_eucl_less: "x \ box a b \ a x box a b \ (\i\Basis. a \ i < x \ i \ x \ i < b \ i)" "x \ cbox a b \ (\i\Basis. a \ i \ x \ i \ x \ i \ b \ i)" by (auto simp: box_eucl_less eucl_less_def cbox_def) lemma cbox_Pair_eq: "cbox (a, c) (b, d) = cbox a b \ cbox c d" by (force simp: cbox_def Basis_prod_def) lemma cbox_Pair_iff [iff]: "(x, y) \ cbox (a, c) (b, d) \ x \ cbox a b \ y \ cbox c d" by (force simp: cbox_Pair_eq) lemma cbox_Complex_eq: "cbox (Complex a c) (Complex b d) = (\(x,y). Complex x y) ` (cbox a b \ cbox c d)" apply (auto simp: cbox_def Basis_complex_def) apply (rule_tac x = "(Re x, Im x)" in image_eqI) using complex_eq by auto lemma cbox_Pair_eq_0: "cbox (a, c) (b, d) = {} \ cbox a b = {} \ cbox c d = {}" by (force simp: cbox_Pair_eq) lemma swap_cbox_Pair [simp]: "prod.swap ` cbox (c, a) (d, b) = cbox (a,c) (b,d)" by auto lemma mem_box_real[simp]: "(x::real) \ box a b \ a < x \ x < b" "(x::real) \ cbox a b \ a \ x \ x \ b" by (auto simp: mem_box) lemma box_real[simp]: fixes a b:: real shows "box a b = {a <..< b}" "cbox a b = {a .. b}" by auto lemma box_Int_box: fixes a :: "'a::euclidean_space" shows "box a b \ box c d = box (\i\Basis. max (a\i) (c\i) *\<^sub>R i) (\i\Basis. min (b\i) (d\i) *\<^sub>R i)" unfolding set_eq_iff and Int_iff and mem_box by auto lemma rational_boxes: fixes x :: "'a::euclidean_space" assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \) \ x \ box a b \ box a b \ ball x e" proof - define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" then have e: "e' > 0" using assms by (auto simp: DIM_positive) have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i - e'" "x \ i"] e show "?th i" by auto qed from choice[OF this] obtain a where a: "\xa. a xa \ \ \ a xa < x \ xa \ x \ xa - a xa < e'" .. have "\i. \y. y \ \ \ x \ i < y \ y - x \ i < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i" "x \ i + e'"] e show "?th i" by auto qed from choice[OF this] obtain b where b: "\xa. b xa \ \ \ x \ xa < b xa \ b xa - x \ xa < e'" .. let ?a = "\i\Basis. a i *\<^sub>R i" and ?b = "\i\Basis. b i *\<^sub>R i" show ?thesis proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) fix y :: 'a assume *: "y \ box ?a ?b" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) also have "\ < sqrt (\(i::'a)\Basis. e^2 / real (DIM('a)))" proof (rule real_sqrt_less_mono, rule sum_strict_mono) fix i :: "'a" assume i: "i \ Basis" have "a i < y\i \ y\i < b i" using * i by (auto simp: box_def) moreover have "a i < x\i" "x\i - a i < e'" using a by auto moreover have "x\i < b i" "b i - x\i < e'" using b by auto ultimately have "\x\i - y\i\ < 2 * e'" by auto then have "dist (x \ i) (y \ i) < e/sqrt (real (DIM('a)))" unfolding e'_def by (auto simp: dist_real_def) then have "(dist (x \ i) (y \ i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" by (rule power_strict_mono) auto then show "(dist (x \ i) (y \ i))\<^sup>2 < e\<^sup>2 / real DIM('a)" by (simp add: power_divide) qed auto also have "\ = e" using \0 < e\ by simp finally show "y \ ball x e" by (auto simp: ball_def) qed (insert a b, auto simp: box_def) qed lemma open_UNION_box: fixes M :: "'a::euclidean_space set" assumes "open M" defines "a' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" defines "b' \ \f :: 'a \ real \ real. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" defines "I \ {f\Basis \\<^sub>E \ \ \. box (a' f) (b' f) \ M}" shows "M = (\f\I. box (a' f) (b' f))" proof - have "x \ (\f\I. box (a' f) (b' f))" if "x \ M" for x proof - obtain e where e: "e > 0" "ball x e \ M" using openE[OF \open M\ \x \ M\] by auto moreover obtain a b where ab: "x \ box a b" "\i \ Basis. a \ i \ \" "\i\Basis. b \ i \ \" "box a b \ ball x e" using rational_boxes[OF e(1)] by metis ultimately show ?thesis by (intro UN_I[of "\i\Basis. (a \ i, b \ i)"]) (auto simp: euclidean_representation I_def a'_def b'_def) qed then show ?thesis by (auto simp: I_def) qed corollary open_countable_Union_open_box: fixes S :: "'a :: euclidean_space set" assumes "open S" obtains \ where "countable \" "\ \ Pow S" "\X. X \ \ \ \a b. X = box a b" "\\ = S" proof - let ?a = "\f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" let ?b = "\f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" let ?I = "{f\Basis \\<^sub>E \ \ \. box (?a f) (?b f) \ S}" let ?\ = "(\f. box (?a f) (?b f)) ` ?I" show ?thesis proof have "countable ?I" by (simp add: countable_PiE countable_rat) then show "countable ?\" by blast show "\?\ = S" using open_UNION_box [OF assms] by metis qed auto qed lemma rational_cboxes: fixes x :: "'a::euclidean_space" assumes "e > 0" shows "\a b. (\i\Basis. a \ i \ \ \ b \ i \ \) \ x \ cbox a b \ cbox a b \ ball x e" proof - define e' where "e' = e / (2 * sqrt (real (DIM ('a))))" then have e: "e' > 0" using assms by auto have "\i. \y. y \ \ \ y < x \ i \ x \ i - y < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i - e'" "x \ i"] e show "?th i" by auto qed from choice[OF this] obtain a where a: "\u. a u \ \ \ a u < x \ u \ x \ u - a u < e'" .. have "\i. \y. y \ \ \ x \ i < y \ y - x \ i < e'" (is "\i. ?th i") proof fix i from Rats_dense_in_real[of "x \ i" "x \ i + e'"] e show "?th i" by auto qed from choice[OF this] obtain b where b: "\u. b u \ \ \ x \ u < b u \ b u - x \ u < e'" .. let ?a = "\i\Basis. a i *\<^sub>R i" and ?b = "\i\Basis. b i *\<^sub>R i" show ?thesis proof (rule exI[of _ ?a], rule exI[of _ ?b], safe) fix y :: 'a assume *: "y \ cbox ?a ?b" have "dist x y = sqrt (\i\Basis. (dist (x \ i) (y \ i))\<^sup>2)" unfolding L2_set_def[symmetric] by (rule euclidean_dist_l2) also have "\ < sqrt (\(i::'a)\Basis. e^2 / real (DIM('a)))" proof (rule real_sqrt_less_mono, rule sum_strict_mono) fix i :: "'a" assume i: "i \ Basis" have "a i \ y\i \ y\i \ b i" using * i by (auto simp: cbox_def) moreover have "a i < x\i" "x\i - a i < e'" using a by auto moreover have "x\i < b i" "b i - x\i < e'" using b by auto ultimately have "\x\i - y\i\ < 2 * e'" by auto then have "dist (x \ i) (y \ i) < e/sqrt (real (DIM('a)))" unfolding e'_def by (auto simp: dist_real_def) then have "(dist (x \ i) (y \ i))\<^sup>2 < (e/sqrt (real (DIM('a))))\<^sup>2" by (rule power_strict_mono) auto then show "(dist (x \ i) (y \ i))\<^sup>2 < e\<^sup>2 / real DIM('a)" by (simp add: power_divide) qed auto also have "\ = e" using \0 < e\ by simp finally show "y \ ball x e" by (auto simp: ball_def) next show "x \ cbox (\i\Basis. a i *\<^sub>R i) (\i\Basis. b i *\<^sub>R i)" using a b less_imp_le by (auto simp: cbox_def) qed (use a b cbox_def in auto) qed lemma open_UNION_cbox: fixes M :: "'a::euclidean_space set" assumes "open M" defines "a' \ \f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" defines "b' \ \f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" defines "I \ {f\Basis \\<^sub>E \ \ \. cbox (a' f) (b' f) \ M}" shows "M = (\f\I. cbox (a' f) (b' f))" proof - have "x \ (\f\I. cbox (a' f) (b' f))" if "x \ M" for x proof - obtain e where e: "e > 0" "ball x e \ M" using openE[OF \open M\ \x \ M\] by auto moreover obtain a b where ab: "x \ cbox a b" "\i \ Basis. a \ i \ \" "\i \ Basis. b \ i \ \" "cbox a b \ ball x e" using rational_cboxes[OF e(1)] by metis ultimately show ?thesis by (intro UN_I[of "\i\Basis. (a \ i, b \ i)"]) (auto simp: euclidean_representation I_def a'_def b'_def) qed then show ?thesis by (auto simp: I_def) qed corollary open_countable_Union_open_cbox: fixes S :: "'a :: euclidean_space set" assumes "open S" obtains \ where "countable \" "\ \ Pow S" "\X. X \ \ \ \a b. X = cbox a b" "\\ = S" proof - let ?a = "\f. (\(i::'a)\Basis. fst (f i) *\<^sub>R i)" let ?b = "\f. (\(i::'a)\Basis. snd (f i) *\<^sub>R i)" let ?I = "{f\Basis \\<^sub>E \ \ \. cbox (?a f) (?b f) \ S}" let ?\ = "(\f. cbox (?a f) (?b f)) ` ?I" show ?thesis proof have "countable ?I" by (simp add: countable_PiE countable_rat) then show "countable ?\" by blast show "\?\ = S" using open_UNION_cbox [OF assms] by metis qed auto qed lemma box_eq_empty: fixes a :: "'a::euclidean_space" shows "(box a b = {} \ (\i\Basis. b\i \ a\i))" (is ?th1) and "(cbox a b = {} \ (\i\Basis. b\i < a\i))" (is ?th2) proof - { fix i x assume i: "i\Basis" and as:"b\i \ a\i" and x:"x\box a b" then have "a \ i < x \ i \ x \ i < b \ i" unfolding mem_box by (auto simp: box_def) then have "a\i < b\i" by auto then have False using as by auto } moreover { assume as: "\i\Basis. \ (b\i \ a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i :: 'a assume i: "i \ Basis" have "a\i < b\i" using as[THEN bspec[where x=i]] i by auto then have "a\i < ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i < b\i" by (auto simp: inner_add_left) } then have "box a b \ {}" using mem_box(1)[of "?x" a b] by auto } ultimately show ?th1 by blast { fix i x assume i: "i \ Basis" and as:"b\i < a\i" and x:"x\cbox a b" then have "a \ i \ x \ i \ x \ i \ b \ i" unfolding mem_box by auto then have "a\i \ b\i" by auto then have False using as by auto } moreover { assume as:"\i\Basis. \ (b\i < a\i)" let ?x = "(1/2) *\<^sub>R (a + b)" { fix i :: 'a assume i:"i \ Basis" have "a\i \ b\i" using as[THEN bspec[where x=i]] i by auto then have "a\i \ ((1/2) *\<^sub>R (a+b)) \ i" "((1/2) *\<^sub>R (a+b)) \ i \ b\i" by (auto simp: inner_add_left) } then have "cbox a b \ {}" using mem_box(2)[of "?x" a b] by auto } ultimately show ?th2 by blast qed lemma box_ne_empty: fixes a :: "'a::euclidean_space" shows "cbox a b \ {} \ (\i\Basis. a\i \ b\i)" and "box a b \ {} \ (\i\Basis. a\i < b\i)" unfolding box_eq_empty[of a b] by fastforce+ lemma fixes a :: "'a::euclidean_space" shows cbox_sing [simp]: "cbox a a = {a}" and box_sing [simp]: "box a a = {}" unfolding set_eq_iff mem_box eq_iff [symmetric] by (auto intro!: euclidean_eqI[where 'a='a]) (metis all_not_in_conv nonempty_Basis) lemma subset_box_imp: fixes a :: "'a::euclidean_space" shows "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ cbox c d \ cbox a b" and "(\i\Basis. a\i < c\i \ d\i < b\i) \ cbox c d \ box a b" and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ cbox a b" and "(\i\Basis. a\i \ c\i \ d\i \ b\i) \ box c d \ box a b" unfolding subset_eq[unfolded Ball_def] unfolding mem_box by (best intro: order_trans less_le_trans le_less_trans less_imp_le)+ lemma box_subset_cbox: fixes a :: "'a::euclidean_space" shows "box a b \ cbox a b" unfolding subset_eq [unfolded Ball_def] mem_box by (fast intro: less_imp_le) lemma subset_box: fixes a :: "'a::euclidean_space" shows "cbox c d \ cbox a b \ (\i\Basis. c\i \ d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th1) and "cbox c d \ box a b \ (\i\Basis. c\i \ d\i) \ (\i\Basis. a\i < c\i \ d\i < b\i)" (is ?th2) and "box c d \ cbox a b \ (\i\Basis. c\i < d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th3) and "box c d \ box a b \ (\i\Basis. c\i < d\i) \ (\i\Basis. a\i \ c\i \ d\i \ b\i)" (is ?th4) proof - let ?lesscd = "\i\Basis. c\i < d\i" let ?lerhs = "\i\Basis. a\i \ c\i \ d\i \ b\i" show ?th1 ?th2 by (fastforce simp: mem_box)+ have acdb: "a\i \ c\i \ d\i \ b\i" if i: "i \ Basis" and box: "box c d \ cbox a b" and cd: "\i. i \ Basis \ c\i < d\i" for i proof - have "box c d \ {}" using that unfolding box_eq_empty by force { let ?x = "(\j\Basis. (if j=i then ((min (a\j) (d\j))+c\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" assume *: "a\i > c\i" then have "c \ j < ?x \ j \ ?x \ j < d \ j" if "j \ Basis" for j using cd that by (fastforce simp add: i *) then have "?x \ box c d" unfolding mem_box by auto moreover have "?x \ cbox a b" using i cd * by (force simp: mem_box) ultimately have False using box by auto } then have "a\i \ c\i" by force moreover { let ?x = "(\j\Basis. (if j=i then ((max (b\j) (c\j))+d\j)/2 else (c\j+d\j)/2) *\<^sub>R j)::'a" assume *: "b\i < d\i" then have "d \ j > ?x \ j \ ?x \ j > c \ j" if "j \ Basis" for j using cd that by (fastforce simp add: i *) then have "?x \ box c d" unfolding mem_box by auto moreover have "?x \ cbox a b" using i cd * by (force simp: mem_box) ultimately have False using box by auto } then have "b\i \ d\i" by (rule ccontr) auto ultimately show ?thesis by auto qed show ?th3 using acdb by (fastforce simp add: mem_box) have acdb': "a\i \ c\i \ d\i \ b\i" if "i \ Basis" "box c d \ box a b" "\i. i \ Basis \ c\i < d\i" for i using box_subset_cbox[of a b] that acdb by auto show ?th4 using acdb' by (fastforce simp add: mem_box) qed lemma eq_cbox: "cbox a b = cbox c d \ cbox a b = {} \ cbox c d = {} \ a = c \ b = d" (is "?lhs = ?rhs") proof assume ?lhs then have "cbox a b \ cbox c d" "cbox c d \ cbox a b" by auto then show ?rhs by (force simp: subset_box box_eq_empty intro: antisym euclidean_eqI) next assume ?rhs then show ?lhs by force qed lemma eq_cbox_box [simp]: "cbox a b = box c d \ cbox a b = {} \ box c d = {}" (is "?lhs \ ?rhs") proof assume L: ?lhs then have "cbox a b \ box c d" "box c d \ cbox a b" by auto then show ?rhs apply (simp add: subset_box) using L box_ne_empty box_sing apply (fastforce simp add:) done qed force lemma eq_box_cbox [simp]: "box a b = cbox c d \ box a b = {} \ cbox c d = {}" by (metis eq_cbox_box) lemma eq_box: "box a b = box c d \ box a b = {} \ box c d = {} \ a = c \ b = d" (is "?lhs \ ?rhs") proof assume L: ?lhs then have "box a b \ box c d" "box c d \ box a b" by auto then show ?rhs apply (simp add: subset_box) using box_ne_empty(2) L apply auto apply (meson euclidean_eqI less_eq_real_def not_less)+ done qed force lemma subset_box_complex: "cbox a b \ cbox c d \ (Re a \ Re b \ Im a \ Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" "cbox a b \ box c d \ (Re a \ Re b \ Im a \ Im b) \ Re a > Re c \ Im a > Im c \ Re b < Re d \ Im b < Im d" "box a b \ cbox c d \ (Re a < Re b \ Im a < Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" "box a b \ box c d \ (Re a < Re b \ Im a < Im b) \ Re a \ Re c \ Im a \ Im c \ Re b \ Re d \ Im b \ Im d" by (subst subset_box; force simp: Basis_complex_def)+ +lemma in_cbox_complex_iff: + "x \ cbox a b \ Re x \ {Re a..Re b} \ Im x \ {Im a..Im b}" + by (cases x; cases a; cases b) (auto simp: cbox_Complex_eq) + +lemma box_Complex_eq: + "box (Complex a c) (Complex b d) = (\(x,y). Complex x y) ` (box a b \ box c d)" + by (auto simp: box_def Basis_complex_def image_iff complex_eq_iff) + +lemma in_box_complex_iff: + "x \ box a b \ Re x \ {Re a<.. Im x \ {Im a<.. cbox c d = cbox (\i\Basis. max (a\i) (c\i) *\<^sub>R i) (\i\Basis. min (b\i) (d\i) *\<^sub>R i)" unfolding set_eq_iff and Int_iff and mem_box by auto lemma disjoint_interval: fixes a::"'a::euclidean_space" shows "cbox a b \ cbox c d = {} \ (\i\Basis. (b\i < a\i \ d\i < c\i \ b\i < c\i \ d\i < a\i))" (is ?th1) and "cbox a b \ box c d = {} \ (\i\Basis. (b\i < a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th2) and "box a b \ cbox c d = {} \ (\i\Basis. (b\i \ a\i \ d\i < c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th3) and "box a b \ box c d = {} \ (\i\Basis. (b\i \ a\i \ d\i \ c\i \ b\i \ c\i \ d\i \ a\i))" (is ?th4) proof - let ?z = "(\i\Basis. (((max (a\i) (c\i)) + (min (b\i) (d\i))) / 2) *\<^sub>R i)::'a" have **: "\P Q. (\i :: 'a. i \ Basis \ Q ?z i \ P i) \ (\i x :: 'a. i \ Basis \ P i \ Q x i) \ (\x. \i\Basis. Q x i) \ (\i\Basis. P i)" by blast note * = set_eq_iff Int_iff empty_iff mem_box ball_conj_distrib[symmetric] eq_False ball_simps(10) show ?th1 unfolding * by (intro **) auto show ?th2 unfolding * by (intro **) auto show ?th3 unfolding * by (intro **) auto show ?th4 unfolding * by (intro **) auto qed lemma UN_box_eq_UNIV: "(\i::nat. box (- (real i *\<^sub>R One)) (real i *\<^sub>R One)) = UNIV" proof - have "\x \ b\ < real_of_int (\Max ((\b. \x \ b\)`Basis)\ + 1)" if [simp]: "b \ Basis" for x b :: 'a proof - have "\x \ b\ \ real_of_int \\x \ b\\" by (rule le_of_int_ceiling) also have "\ \ real_of_int \Max ((\b. \x \ b\)`Basis)\" by (auto intro!: ceiling_mono) also have "\ < real_of_int (\Max ((\b. \x \ b\)`Basis)\ + 1)" by simp finally show ?thesis . qed then have "\n::nat. \b\Basis. \x \ b\ < real n" for x :: 'a by (metis order.strict_trans reals_Archimedean2) moreover have "\x b::'a. \n::nat. \x \ b\ < real n \ - real n < x \ b \ x \ b < real n" by auto ultimately show ?thesis by (auto simp: box_def inner_sum_left inner_Basis sum.If_cases) qed lemma image_affinity_cbox: fixes m::real fixes a b c :: "'a::euclidean_space" shows "(\x. m *\<^sub>R x + c) ` cbox a b = (if cbox a b = {} then {} else (if 0 \ m then cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) else cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c)))" proof (cases "m = 0") case True { fix x assume "\i\Basis. x \ i \ c \ i" "\i\Basis. c \ i \ x \ i" then have "x = c" by (simp add: dual_order.antisym euclidean_eqI) } moreover have "c \ cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c)" unfolding True by (auto simp: cbox_sing) ultimately show ?thesis using True by (auto simp: cbox_def) next case False { fix y assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m > 0" then have "\i\Basis. (m *\<^sub>R a + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R b + c) \ i" by (auto simp: inner_distrib) } moreover { fix y assume "\i\Basis. a \ i \ y \ i" "\i\Basis. y \ i \ b \ i" "m < 0" then have "\i\Basis. (m *\<^sub>R b + c) \ i \ (m *\<^sub>R y + c) \ i" and "\i\Basis. (m *\<^sub>R y + c) \ i \ (m *\<^sub>R a + c) \ i" by (auto simp: mult_left_mono_neg inner_distrib) } moreover { fix y assume "m > 0" and "\i\Basis. (m *\<^sub>R a + c) \ i \ y \ i" and "\i\Basis. y \ i \ (m *\<^sub>R b + c) \ i" then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp: pos_le_divide_eq pos_divide_le_eq mult.commute inner_distrib inner_diff_left) done } moreover { fix y assume "\i\Basis. (m *\<^sub>R b + c) \ i \ y \ i" "\i\Basis. y \ i \ (m *\<^sub>R a + c) \ i" "m < 0" then have "y \ (\x. m *\<^sub>R x + c) ` cbox a b" unfolding image_iff Bex_def mem_box apply (intro exI[where x="(1 / m) *\<^sub>R (y - c)"]) apply (auto simp: neg_le_divide_eq neg_divide_le_eq mult.commute inner_distrib inner_diff_left) done } ultimately show ?thesis using False by (auto simp: cbox_def) qed lemma image_smult_cbox:"(\x. m *\<^sub>R (x::_::euclidean_space)) ` cbox a b = (if cbox a b = {} then {} else if 0 \ m then cbox (m *\<^sub>R a) (m *\<^sub>R b) else cbox (m *\<^sub>R b) (m *\<^sub>R a))" using image_affinity_cbox[of m 0 a b] by auto lemma swap_continuous: assumes "continuous_on (cbox (a,c) (b,d)) (\(x,y). f x y)" shows "continuous_on (cbox (c,a) (d,b)) (\(x, y). f y x)" proof - have "(\(x, y). f y x) = (\(x, y). f x y) \ prod.swap" by auto then show ?thesis apply (rule ssubst) apply (rule continuous_on_compose) apply (simp add: split_def) apply (rule continuous_intros | simp add: assms)+ done qed subsection \General Intervals\ definition\<^marker>\tag important\ "is_interval (s::('a::euclidean_space) set) \ (\a\s. \b\s. \x. (\i\Basis. ((a\i \ x\i \ x\i \ b\i) \ (b\i \ x\i \ x\i \ a\i))) \ x \ s)" lemma is_interval_1: "is_interval (s::real set) \ (\a\s. \b\s. \ x. a \ x \ x \ b \ x \ s)" unfolding is_interval_def by auto lemma is_interval_Int: "is_interval X \ is_interval Y \ is_interval (X \ Y)" unfolding is_interval_def by blast lemma is_interval_cbox [simp]: "is_interval (cbox a (b::'a::euclidean_space))" (is ?th1) and is_interval_box [simp]: "is_interval (box a b)" (is ?th2) unfolding is_interval_def mem_box Ball_def atLeastAtMost_iff by (meson order_trans le_less_trans less_le_trans less_trans)+ lemma is_interval_empty [iff]: "is_interval {}" unfolding is_interval_def by simp lemma is_interval_univ [iff]: "is_interval UNIV" unfolding is_interval_def by simp lemma mem_is_intervalI: assumes "is_interval s" and "a \ s" "b \ s" and "\i. i \ Basis \ a \ i \ x \ i \ x \ i \ b \ i \ b \ i \ x \ i \ x \ i \ a \ i" shows "x \ s" by (rule assms(1)[simplified is_interval_def, rule_format, OF assms(2,3,4)]) lemma interval_subst: fixes S::"'a::euclidean_space set" assumes "is_interval S" and "x \ S" "y j \ S" and "j \ Basis" shows "(\i\Basis. (if i = j then y i \ i else x \ i) *\<^sub>R i) \ S" by (rule mem_is_intervalI[OF assms(1,2)]) (auto simp: assms) lemma mem_box_componentwiseI: fixes S::"'a::euclidean_space set" assumes "is_interval S" assumes "\i. i \ Basis \ x \ i \ ((\x. x \ i) ` S)" shows "x \ S" proof - from assms have "\i \ Basis. \s \ S. x \ i = s \ i" by auto with finite_Basis obtain s and bs::"'a list" where s: "\i. i \ Basis \ x \ i = s i \ i" "\i. i \ Basis \ s i \ S" and bs: "set bs = Basis" "distinct bs" by (metis finite_distinct_list) from nonempty_Basis s obtain j where j: "j \ Basis" "s j \ S" by blast define y where "y = rec_list (s j) (\j _ Y. (\i\Basis. (if i = j then s i \ i else Y \ i) *\<^sub>R i))" have "x = (\i\Basis. (if i \ set bs then s i \ i else s j \ i) *\<^sub>R i)" using bs by (auto simp: s(1)[symmetric] euclidean_representation) also have [symmetric]: "y bs = \" using bs(2) bs(1)[THEN equalityD1] by (induct bs) (auto simp: y_def euclidean_representation intro!: euclidean_eqI[where 'a='a]) also have "y bs \ S" using bs(1)[THEN equalityD1] apply (induct bs) apply (auto simp: y_def j) apply (rule interval_subst[OF assms(1)]) apply (auto simp: s) done finally show ?thesis . qed lemma cbox01_nonempty [simp]: "cbox 0 One \ {}" by (simp add: box_ne_empty inner_Basis inner_sum_left sum_nonneg) lemma box01_nonempty [simp]: "box 0 One \ {}" by (simp add: box_ne_empty inner_Basis inner_sum_left) lemma empty_as_interval: "{} = cbox One (0::'a::euclidean_space)" using nonempty_Basis box01_nonempty box_eq_empty(1) box_ne_empty(1) by blast lemma interval_subset_is_interval: assumes "is_interval S" shows "cbox a b \ S \ cbox a b = {} \ a \ S \ b \ S" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs using box_ne_empty(1) mem_box(2) by fastforce next assume ?rhs have "cbox a b \ S" if "a \ S" "b \ S" using assms unfolding is_interval_def apply (clarsimp simp add: mem_box) using that by blast with \?rhs\ show ?lhs by blast qed lemma is_real_interval_union: "is_interval (X \ Y)" if X: "is_interval X" and Y: "is_interval Y" and I: "(X \ {} \ Y \ {} \ X \ Y \ {})" for X Y::"real set" proof - consider "X \ {}" "Y \ {}" | "X = {}" | "Y = {}" by blast then show ?thesis proof cases case 1 then obtain r where "r \ X \ X \ Y = {}" "r \ Y \ X \ Y = {}" by blast then show ?thesis using I 1 X Y unfolding is_interval_1 by (metis (full_types) Un_iff le_cases) qed (use that in auto) qed lemma is_interval_translationI: assumes "is_interval X" shows "is_interval ((+) x ` X)" unfolding is_interval_def proof safe fix b d e assume "b \ X" "d \ X" "\i\Basis. (x + b) \ i \ e \ i \ e \ i \ (x + d) \ i \ (x + d) \ i \ e \ i \ e \ i \ (x + b) \ i" hence "e - x \ X" by (intro mem_is_intervalI[OF assms \b \ X\ \d \ X\, of "e - x"]) (auto simp: algebra_simps) thus "e \ (+) x ` X" by force qed lemma is_interval_uminusI: assumes "is_interval X" shows "is_interval (uminus ` X)" unfolding is_interval_def proof safe fix b d e assume "b \ X" "d \ X" "\i\Basis. (- b) \ i \ e \ i \ e \ i \ (- d) \ i \ (- d) \ i \ e \ i \ e \ i \ (- b) \ i" hence "- e \ X" by (intro mem_is_intervalI[OF assms \b \ X\ \d \ X\, of "- e"]) (auto simp: algebra_simps) thus "e \ uminus ` X" by force qed lemma is_interval_uminus[simp]: "is_interval (uminus ` x) = is_interval x" using is_interval_uminusI[of x] is_interval_uminusI[of "uminus ` x"] by (auto simp: image_image) lemma is_interval_neg_translationI: assumes "is_interval X" shows "is_interval ((-) x ` X)" proof - have "(-) x ` X = (+) x ` uminus ` X" by (force simp: algebra_simps) also have "is_interval \" by (metis is_interval_uminusI is_interval_translationI assms) finally show ?thesis . qed lemma is_interval_translation[simp]: "is_interval ((+) x ` X) = is_interval X" using is_interval_neg_translationI[of "(+) x ` X" x] by (auto intro!: is_interval_translationI simp: image_image) lemma is_interval_minus_translation[simp]: shows "is_interval ((-) x ` X) = is_interval X" proof - have "(-) x ` X = (+) x ` uminus ` X" by (force simp: algebra_simps) also have "is_interval \ = is_interval X" by simp finally show ?thesis . qed lemma is_interval_minus_translation'[simp]: shows "is_interval ((\x. x - c) ` X) = is_interval X" using is_interval_translation[of "-c" X] by (metis image_cong uminus_add_conv_diff) lemma is_interval_cball_1[intro, simp]: "is_interval (cball a b)" for a b::real by (simp add: cball_eq_atLeastAtMost is_interval_def) lemma is_interval_ball_real: "is_interval (ball a b)" for a b::real by (simp add: ball_eq_greaterThanLessThan is_interval_def) subsection\<^marker>\tag unimportant\ \Bounded Projections\ lemma bounded_inner_imp_bdd_above: assumes "bounded s" shows "bdd_above ((\x. x \ a) ` s)" by (simp add: assms bounded_imp_bdd_above bounded_linear_image bounded_linear_inner_left) lemma bounded_inner_imp_bdd_below: assumes "bounded s" shows "bdd_below ((\x. x \ a) ` s)" by (simp add: assms bounded_imp_bdd_below bounded_linear_image bounded_linear_inner_left) subsection\<^marker>\tag unimportant\ \Structural rules for pointwise continuity\ lemma continuous_infnorm[continuous_intros]: "continuous F f \ continuous F (\x. infnorm (f x))" unfolding continuous_def by (rule tendsto_infnorm) lemma continuous_inner[continuous_intros]: assumes "continuous F f" and "continuous F g" shows "continuous F (\x. inner (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_inner) subsection\<^marker>\tag unimportant\ \Structural rules for setwise continuity\ lemma continuous_on_infnorm[continuous_intros]: "continuous_on s f \ continuous_on s (\x. infnorm (f x))" unfolding continuous_on by (fast intro: tendsto_infnorm) lemma continuous_on_inner[continuous_intros]: fixes g :: "'a::topological_space \ 'b::real_inner" assumes "continuous_on s f" and "continuous_on s g" shows "continuous_on s (\x. inner (f x) (g x))" using bounded_bilinear_inner assms by (rule bounded_bilinear.continuous_on) subsection\<^marker>\tag unimportant\ \Openness of halfspaces.\ lemma open_halfspace_lt: "open {x. inner a x < b}" by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) lemma open_halfspace_gt: "open {x. inner a x > b}" by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) lemma open_halfspace_component_lt: "open {x::'a::euclidean_space. x\i < a}" by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) lemma open_halfspace_component_gt: "open {x::'a::euclidean_space. x\i > a}" by (simp add: open_Collect_less continuous_on_inner continuous_on_const continuous_on_id) lemma eucl_less_eq_halfspaces: fixes a :: "'a::euclidean_space" shows "{x. x i\Basis. {x. x \ i < a \ i})" "{x. a i\Basis. {x. a \ i < x \ i})" by (auto simp: eucl_less_def) lemma open_Collect_eucl_less[simp, intro]: fixes a :: "'a::euclidean_space" shows "open {x. x \tag unimportant\ \Closure and Interior of halfspaces and hyperplanes\ lemma continuous_at_inner: "continuous (at x) (inner a)" unfolding continuous_at by (intro tendsto_intros) lemma closed_halfspace_le: "closed {x. inner a x \ b}" by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) lemma closed_halfspace_ge: "closed {x. inner a x \ b}" by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) lemma closed_hyperplane: "closed {x. inner a x = b}" by (simp add: closed_Collect_eq continuous_on_inner continuous_on_const continuous_on_id) lemma closed_halfspace_component_le: "closed {x::'a::euclidean_space. x\i \ a}" by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) lemma closed_halfspace_component_ge: "closed {x::'a::euclidean_space. x\i \ a}" by (simp add: closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) lemma closed_interval_left: fixes b :: "'a::euclidean_space" shows "closed {x::'a. \i\Basis. x\i \ b\i}" by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) lemma closed_interval_right: fixes a :: "'a::euclidean_space" shows "closed {x::'a. \i\Basis. a\i \ x\i}" by (simp add: Collect_ball_eq closed_INT closed_Collect_le continuous_on_inner continuous_on_const continuous_on_id) lemma interior_halfspace_le [simp]: assumes "a \ 0" shows "interior {x. a \ x \ b} = {x. a \ x < b}" proof - have *: "a \ x < b" if x: "x \ S" and S: "S \ {x. a \ x \ b}" and "open S" for S x proof - obtain e where "e>0" and e: "cball x e \ S" using \open S\ open_contains_cball x by blast then have "x + (e / norm a) *\<^sub>R a \ cball x e" by (simp add: dist_norm) then have "x + (e / norm a) *\<^sub>R a \ S" using e by blast then have "x + (e / norm a) *\<^sub>R a \ {x. a \ x \ b}" using S by blast moreover have "e * (a \ a) / norm a > 0" by (simp add: \0 < e\ assms) ultimately show ?thesis by (simp add: algebra_simps) qed show ?thesis by (rule interior_unique) (auto simp: open_halfspace_lt *) qed lemma interior_halfspace_ge [simp]: "a \ 0 \ interior {x. a \ x \ b} = {x. a \ x > b}" using interior_halfspace_le [of "-a" "-b"] by simp lemma closure_halfspace_lt [simp]: assumes "a \ 0" shows "closure {x. a \ x < b} = {x. a \ x \ b}" proof - have [simp]: "-{x. a \ x < b} = {x. a \ x \ b}" by (force simp:) then show ?thesis using interior_halfspace_ge [of a b] assms by (force simp: closure_interior) qed lemma closure_halfspace_gt [simp]: "a \ 0 \ closure {x. a \ x > b} = {x. a \ x \ b}" using closure_halfspace_lt [of "-a" "-b"] by simp lemma interior_hyperplane [simp]: assumes "a \ 0" shows "interior {x. a \ x = b} = {}" proof - have [simp]: "{x. a \ x = b} = {x. a \ x \ b} \ {x. a \ x \ b}" by (force simp:) then show ?thesis by (auto simp: assms) qed lemma frontier_halfspace_le: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x \ b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def closed_halfspace_le) qed lemma frontier_halfspace_ge: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x \ b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def closed_halfspace_ge) qed lemma frontier_halfspace_lt: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x < b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def interior_open open_halfspace_lt) qed lemma frontier_halfspace_gt: assumes "a \ 0 \ b \ 0" shows "frontier {x. a \ x > b} = {x. a \ x = b}" proof (cases "a = 0") case True with assms show ?thesis by simp next case False then show ?thesis by (force simp: frontier_def interior_open open_halfspace_gt) qed subsection\<^marker>\tag unimportant\\Some more convenient intermediate-value theorem formulations\ lemma connected_ivt_hyperplane: assumes "connected S" and xy: "x \ S" "y \ S" and b: "inner a x \ b" "b \ inner a y" shows "\z \ S. inner a z = b" proof (rule ccontr) assume as:"\ (\z\S. inner a z = b)" let ?A = "{x. inner a x < b}" let ?B = "{x. inner a x > b}" have "open ?A" "open ?B" using open_halfspace_lt and open_halfspace_gt by auto moreover have "?A \ ?B = {}" by auto moreover have "S \ ?A \ ?B" using as by auto ultimately show False using \connected S\[unfolded connected_def not_ex, THEN spec[where x="?A"], THEN spec[where x="?B"]] using xy b by auto qed lemma connected_ivt_component: fixes x::"'a::euclidean_space" shows "connected S \ x \ S \ y \ S \ x\k \ a \ a \ y\k \ (\z\S. z\k = a)" using connected_ivt_hyperplane[of S x y "k::'a" a] by (auto simp: inner_commute) subsection \Limit Component Bounds\ lemma Lim_component_le: fixes f :: "'a \ 'b::euclidean_space" assumes "(f \ l) net" and "\ (trivial_limit net)" and "eventually (\x. f(x)\i \ b) net" shows "l\i \ b" by (rule tendsto_le[OF assms(2) tendsto_const tendsto_inner[OF assms(1) tendsto_const] assms(3)]) lemma Lim_component_ge: fixes f :: "'a \ 'b::euclidean_space" assumes "(f \ l) net" and "\ (trivial_limit net)" and "eventually (\x. b \ (f x)\i) net" shows "b \ l\i" by (rule tendsto_le[OF assms(2) tendsto_inner[OF assms(1) tendsto_const] tendsto_const assms(3)]) lemma Lim_component_eq: fixes f :: "'a \ 'b::euclidean_space" assumes net: "(f \ l) net" "\ trivial_limit net" and ev:"eventually (\x. f(x)\i = b) net" shows "l\i = b" using ev[unfolded order_eq_iff eventually_conj_iff] using Lim_component_ge[OF net, of b i] using Lim_component_le[OF net, of i b] by auto lemma open_box[intro]: "open (box a b)" proof - have "open (\i\Basis. ((\) i) -` {a \ i <..< b \ i})" by (auto intro!: continuous_open_vimage continuous_inner continuous_ident continuous_const) also have "(\i\Basis. ((\) i) -` {a \ i <..< b \ i}) = box a b" by (auto simp: box_def inner_commute) finally show ?thesis . qed lemma closed_cbox[intro]: fixes a b :: "'a::euclidean_space" shows "closed (cbox a b)" proof - have "closed (\i\Basis. (\x. x\i) -` {a\i .. b\i})" by (intro closed_INT ballI continuous_closed_vimage allI linear_continuous_at closed_real_atLeastAtMost finite_Basis bounded_linear_inner_left) also have "(\i\Basis. (\x. x\i) -` {a\i .. b\i}) = cbox a b" by (auto simp: cbox_def) finally show "closed (cbox a b)" . qed lemma interior_cbox [simp]: fixes a b :: "'a::euclidean_space" shows "interior (cbox a b) = box a b" (is "?L = ?R") proof(rule subset_antisym) show "?R \ ?L" using box_subset_cbox open_box by (rule interior_maximal) { fix x assume "x \ interior (cbox a b)" then obtain s where s: "open s" "x \ s" "s \ cbox a b" .. then obtain e where "e>0" and e:"\x'. dist x' x < e \ x' \ cbox a b" unfolding open_dist and subset_eq by auto { fix i :: 'a assume i: "i \ Basis" have "dist (x - (e / 2) *\<^sub>R i) x < e" and "dist (x + (e / 2) *\<^sub>R i) x < e" unfolding dist_norm apply auto unfolding norm_minus_cancel using norm_Basis[OF i] \e>0\ apply auto done then have "a \ i \ (x - (e / 2) *\<^sub>R i) \ i" and "(x + (e / 2) *\<^sub>R i) \ i \ b \ i" using e[THEN spec[where x="x - (e/2) *\<^sub>R i"]] and e[THEN spec[where x="x + (e/2) *\<^sub>R i"]] unfolding mem_box using i by blast+ then have "a \ i < x \ i" and "x \ i < b \ i" using \e>0\ i by (auto simp: inner_diff_left inner_Basis inner_add_left) } then have "x \ box a b" unfolding mem_box by auto } then show "?L \ ?R" .. qed lemma bounded_cbox [simp]: fixes a :: "'a::euclidean_space" shows "bounded (cbox a b)" proof - let ?b = "\i\Basis. \a\i\ + \b\i\" { fix x :: "'a" assume "\i. i\Basis \ a \ i \ x \ i \ x \ i \ b \ i" then have "(\i\Basis. \x \ i\) \ ?b" by (force simp: intro!: sum_mono) then have "norm x \ ?b" using norm_le_l1[of x] by auto } then show ?thesis unfolding cbox_def bounded_iff by force qed lemma bounded_box [simp]: fixes a :: "'a::euclidean_space" shows "bounded (box a b)" using bounded_cbox[of a b] box_subset_cbox[of a b] bounded_subset[of "cbox a b" "box a b"] by simp lemma not_interval_UNIV [simp]: fixes a :: "'a::euclidean_space" shows "cbox a b \ UNIV" "box a b \ UNIV" using bounded_box[of a b] bounded_cbox[of a b] by force+ lemma not_interval_UNIV2 [simp]: fixes a :: "'a::euclidean_space" shows "UNIV \ cbox a b" "UNIV \ box a b" using bounded_box[of a b] bounded_cbox[of a b] by force+ lemma box_midpoint: fixes a :: "'a::euclidean_space" assumes "box a b \ {}" shows "((1/2) *\<^sub>R (a + b)) \ box a b" proof - have "a \ i < ((1 / 2) *\<^sub>R (a + b)) \ i \ ((1 / 2) *\<^sub>R (a + b)) \ i < b \ i" if "i \ Basis" for i using assms that by (auto simp: inner_add_left box_ne_empty) then show ?thesis unfolding mem_box by auto qed lemma open_cbox_convex: fixes x :: "'a::euclidean_space" assumes x: "x \ box a b" and y: "y \ cbox a b" and e: "0 < e" "e \ 1" shows "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ box a b" proof - { fix i :: 'a assume i: "i \ Basis" have "a \ i = e * (a \ i) + (1 - e) * (a \ i)" unfolding left_diff_distrib by simp also have "\ < e * (x \ i) + (1 - e) * (y \ i)" proof (rule add_less_le_mono) show "e * (a \ i) < e * (x \ i)" using \0 < e\ i mem_box(1) x by auto show "(1 - e) * (a \ i) \ (1 - e) * (y \ i)" by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) qed finally have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i" unfolding inner_simps by auto moreover { have "b \ i = e * (b\i) + (1 - e) * (b\i)" unfolding left_diff_distrib by simp also have "\ > e * (x \ i) + (1 - e) * (y \ i)" proof (rule add_less_le_mono) show "e * (x \ i) < e * (b \ i)" using \0 < e\ i mem_box(1) x by auto show "(1 - e) * (y \ i) \ (1 - e) * (b \ i)" by (meson diff_ge_0_iff_ge \e \ 1\ i mem_box(2) mult_left_mono y) qed finally have "(e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" unfolding inner_simps by auto } ultimately have "a \ i < (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i \ (e *\<^sub>R x + (1 - e) *\<^sub>R y) \ i < b \ i" by auto } then show ?thesis unfolding mem_box by auto qed lemma closure_cbox [simp]: "closure (cbox a b) = cbox a b" by (simp add: closed_cbox) lemma closure_box [simp]: fixes a :: "'a::euclidean_space" assumes "box a b \ {}" shows "closure (box a b) = cbox a b" proof - have ab: "a cbox a b" define f where [abs_def]: "f n = x + (inverse (real n + 1)) *\<^sub>R (?c - x)" for n { fix n assume fn: "f n a f n = x" and xc: "x \ ?c" have *: "0 < inverse (real n + 1)" "inverse (real n + 1) \ 1" unfolding inverse_le_1_iff by auto have "(inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b)) + (1 - inverse (real n + 1)) *\<^sub>R x = x + (inverse (real n + 1)) *\<^sub>R (((1 / 2) *\<^sub>R (a + b)) - x)" by (auto simp: algebra_simps) then have "f n (f \ x) sequentially" { fix e :: real assume "e > 0" then obtain N :: nat where N: "inverse (real (N + 1)) < e" using reals_Archimedean by auto have "inverse (real n + 1) < e" if "N \ n" for n by (auto intro!: that le_less_trans [OF _ N]) then have "\N::nat. \n\N. inverse (real n + 1) < e" by auto } then have "((\n. inverse (real n + 1)) \ 0) sequentially" unfolding lim_sequentially by(auto simp: dist_norm) then have "(f \ x) sequentially" unfolding f_def using tendsto_add[OF tendsto_const, of "\n::nat. (inverse (real n + 1)) *\<^sub>R ((1 / 2) *\<^sub>R (a + b) - x)" 0 sequentially x] using tendsto_scaleR [OF _ tendsto_const, of "\n::nat. inverse (real n + 1)" 0 sequentially "((1 / 2) *\<^sub>R (a + b) - x)"] by auto } ultimately have "x \ closure (box a b)" using as box_midpoint[OF assms] unfolding closure_def islimpt_sequential by (cases "x=?c") (auto simp: in_box_eucl_less) } then show ?thesis using closure_minimal[OF box_subset_cbox, of a b] by blast qed lemma bounded_subset_box_symmetric: fixes S :: "('a::euclidean_space) set" assumes "bounded S" obtains a where "S \ box (-a) a" proof - obtain b where "b>0" and b: "\x\S. norm x \ b" using assms[unfolded bounded_pos] by auto define a :: 'a where "a = (\i\Basis. (b + 1) *\<^sub>R i)" have "(-a)\i < x\i" and "x\i < a\i" if "x \ S" and i: "i \ Basis" for x i using b Basis_le_norm[OF i, of x] that by (auto simp: a_def) then have "S \ box (-a) a" by (auto simp: simp add: box_def) then show ?thesis .. qed lemma bounded_subset_cbox_symmetric: fixes S :: "('a::euclidean_space) set" assumes "bounded S" obtains a where "S \ cbox (-a) a" proof - obtain a where "S \ box (-a) a" using bounded_subset_box_symmetric[OF assms] by auto then show ?thesis by (meson box_subset_cbox dual_order.trans that) qed lemma frontier_cbox: fixes a b :: "'a::euclidean_space" shows "frontier (cbox a b) = cbox a b - box a b" unfolding frontier_def unfolding interior_cbox and closure_closed[OF closed_cbox] .. lemma frontier_box: fixes a b :: "'a::euclidean_space" shows "frontier (box a b) = (if box a b = {} then {} else cbox a b - box a b)" proof (cases "box a b = {}") case True then show ?thesis using frontier_empty by auto next case False then show ?thesis unfolding frontier_def and closure_box[OF False] and interior_open[OF open_box] by auto qed lemma Int_interval_mixed_eq_empty: fixes a :: "'a::euclidean_space" assumes "box c d \ {}" shows "box a b \ cbox c d = {} \ box a b \ box c d = {}" unfolding closure_box[OF assms, symmetric] unfolding open_Int_closure_eq_empty[OF open_box] .. subsection \Class Instances\ lemma compact_lemma: fixes f :: "nat \ 'a::euclidean_space" assumes "bounded (range f)" shows "\d\Basis. \l::'a. \ r. strict_mono r \ (\e>0. eventually (\n. \i\d. dist (f (r n) \ i) (l \ i) < e) sequentially)" by (rule compact_lemma_general[where unproj="\e. \i\Basis. e i *\<^sub>R i"]) (auto intro!: assms bounded_linear_inner_left bounded_linear_image simp: euclidean_representation) instance\<^marker>\tag important\ euclidean_space \ heine_borel proof fix f :: "nat \ 'a" assume f: "bounded (range f)" then obtain l::'a and r where r: "strict_mono r" and l: "\e>0. eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e) sequentially" using compact_lemma [OF f] by blast { fix e::real assume "e > 0" hence "e / real_of_nat DIM('a) > 0" by (simp add: DIM_positive) with l have "eventually (\n. \i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))) sequentially" by simp moreover { fix n assume n: "\i\Basis. dist (f (r n) \ i) (l \ i) < e / (real_of_nat DIM('a))" have "dist (f (r n)) l \ (\i\Basis. dist (f (r n) \ i) (l \ i))" apply (subst euclidean_dist_l2) using zero_le_dist apply (rule L2_set_le_sum) done also have "\ < (\i\(Basis::'a set). e / (real_of_nat DIM('a)))" apply (rule sum_strict_mono) using n apply auto done finally have "dist (f (r n)) l < e" by auto } ultimately have "eventually (\n. dist (f (r n)) l < e) sequentially" by (rule eventually_mono) } then have *: "((f \ r) \ l) sequentially" unfolding o_def tendsto_iff by simp with r show "\l r. strict_mono r \ ((f \ r) \ l) sequentially" by auto qed instance\<^marker>\tag important\ euclidean_space \ banach .. instance euclidean_space \ second_countable_topology proof define a where "a f = (\i\Basis. fst (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have a: "\f. (\i\Basis. fst (f i) *\<^sub>R i) = a f" by simp define b where "b f = (\i\Basis. snd (f i) *\<^sub>R i)" for f :: "'a \ real \ real" then have b: "\f. (\i\Basis. snd (f i) *\<^sub>R i) = b f" by simp define B where "B = (\f. box (a f) (b f)) ` (Basis \\<^sub>E (\ \ \))" have "Ball B open" by (simp add: B_def open_box) moreover have "(\A. open A \ (\B'\B. \B' = A))" proof safe fix A::"'a set" assume "open A" show "\B'\B. \B' = A" apply (rule exI[of _ "{b\B. b \ A}"]) apply (subst (3) open_UNION_box[OF \open A\]) apply (auto simp: a b B_def) done qed ultimately have "topological_basis B" unfolding topological_basis_def by blast moreover have "countable B" unfolding B_def by (intro countable_image countable_PiE finite_Basis countable_SIGMA countable_rat) ultimately show "\B::'a set set. countable B \ open = generate_topology B" by (blast intro: topological_basis_imp_subbasis) qed instance euclidean_space \ polish_space .. subsection \Compact Boxes\ lemma compact_cbox [simp]: fixes a :: "'a::euclidean_space" shows "compact (cbox a b)" using bounded_closed_imp_seq_compact[of "cbox a b"] using bounded_cbox[of a b] by (auto simp: compact_eq_seq_compact_metric) proposition is_interval_compact: "is_interval S \ compact S \ (\a b. S = cbox a b)" (is "?lhs = ?rhs") proof (cases "S = {}") case True with empty_as_interval show ?thesis by auto next case False show ?thesis proof assume L: ?lhs then have "is_interval S" "compact S" by auto define a where "a \ \i\Basis. (INF x\S. x \ i) *\<^sub>R i" define b where "b \ \i\Basis. (SUP x\S. x \ i) *\<^sub>R i" have 1: "\x i. \x \ S; i \ Basis\ \ (INF x\S. x \ i) \ x \ i" by (simp add: cInf_lower bounded_inner_imp_bdd_below compact_imp_bounded L) have 2: "\x i. \x \ S; i \ Basis\ \ x \ i \ (SUP x\S. x \ i)" by (simp add: cSup_upper bounded_inner_imp_bdd_above compact_imp_bounded L) have 3: "x \ S" if inf: "\i. i \ Basis \ (INF x\S. x \ i) \ x \ i" and sup: "\i. i \ Basis \ x \ i \ (SUP x\S. x \ i)" for x proof (rule mem_box_componentwiseI [OF \is_interval S\]) fix i::'a assume i: "i \ Basis" have cont: "continuous_on S (\x. x \ i)" by (intro continuous_intros) obtain a where "a \ S" and a: "\y. y\S \ a \ i \ y \ i" using continuous_attains_inf [OF \compact S\ False cont] by blast obtain b where "b \ S" and b: "\y. y\S \ y \ i \ b \ i" using continuous_attains_sup [OF \compact S\ False cont] by blast have "a \ i \ (INF x\S. x \ i)" by (simp add: False a cINF_greatest) also have "\ \ x \ i" by (simp add: i inf) finally have ai: "a \ i \ x \ i" . have "x \ i \ (SUP x\S. x \ i)" by (simp add: i sup) also have "(SUP x\S. x \ i) \ b \ i" by (simp add: False b cSUP_least) finally have bi: "x \ i \ b \ i" . show "x \ i \ (\x. x \ i) ` S" apply (rule_tac x="\j\Basis. (if j = i then x \ i else a \ j) *\<^sub>R j" in image_eqI) apply (simp add: i) apply (rule mem_is_intervalI [OF \is_interval S\ \a \ S\ \b \ S\]) using i ai bi apply force done qed have "S = cbox a b" by (auto simp: a_def b_def mem_box intro: 1 2 3) then show ?rhs by blast next assume R: ?rhs then show ?lhs using compact_cbox is_interval_cbox by blast qed qed subsection\<^marker>\tag unimportant\\Componentwise limits and continuity\ text\But is the premise really necessary? Need to generalise @{thm euclidean_dist_l2}\ lemma Euclidean_dist_upper: "i \ Basis \ dist (x \ i) (y \ i) \ dist x y" by (metis (no_types) member_le_L2_set euclidean_dist_l2 finite_Basis) text\But is the premise \<^term>\i \ Basis\ really necessary?\ lemma open_preimage_inner: assumes "open S" "i \ Basis" shows "open {x. x \ i \ S}" proof (rule openI, simp) fix x assume x: "x \ i \ S" with assms obtain e where "0 < e" and e: "ball (x \ i) e \ S" by (auto simp: open_contains_ball_eq) have "\e>0. ball (y \ i) e \ S" if dxy: "dist x y < e / 2" for y proof (intro exI conjI) have "dist (x \ i) (y \ i) < e / 2" by (meson \i \ Basis\ dual_order.trans Euclidean_dist_upper not_le that) then have "dist (x \ i) z < e" if "dist (y \ i) z < e / 2" for z by (metis dist_commute dist_triangle_half_l that) then have "ball (y \ i) (e / 2) \ ball (x \ i) e" using mem_ball by blast with e show "ball (y \ i) (e / 2) \ S" by (metis order_trans) qed (simp add: \0 < e\) then show "\e>0. ball x e \ {s. s \ i \ S}" by (metis (no_types, lifting) \0 < e\ \open S\ half_gt_zero_iff mem_Collect_eq mem_ball open_contains_ball_eq subsetI) qed proposition tendsto_componentwise_iff: fixes f :: "_ \ 'b::euclidean_space" shows "(f \ l) F \ (\i \ Basis. ((\x. (f x \ i)) \ (l \ i)) F)" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs unfolding tendsto_def apply clarify apply (drule_tac x="{s. s \ i \ S}" in spec) apply (auto simp: open_preimage_inner) done next assume R: ?rhs then have "\e. e > 0 \ \i\Basis. \\<^sub>F x in F. dist (f x \ i) (l \ i) < e" unfolding tendsto_iff by blast then have R': "\e. e > 0 \ \\<^sub>F x in F. \i\Basis. dist (f x \ i) (l \ i) < e" by (simp add: eventually_ball_finite_distrib [symmetric]) show ?lhs unfolding tendsto_iff proof clarify fix e::real assume "0 < e" have *: "L2_set (\i. dist (f x \ i) (l \ i)) Basis < e" if "\i\Basis. dist (f x \ i) (l \ i) < e / real DIM('b)" for x proof - have "L2_set (\i. dist (f x \ i) (l \ i)) Basis \ sum (\i. dist (f x \ i) (l \ i)) Basis" by (simp add: L2_set_le_sum) also have "... < DIM('b) * (e / real DIM('b))" apply (rule sum_bounded_above_strict) using that by auto also have "... = e" by (simp add: field_simps) finally show "L2_set (\i. dist (f x \ i) (l \ i)) Basis < e" . qed have "\\<^sub>F x in F. \i\Basis. dist (f x \ i) (l \ i) < e / DIM('b)" apply (rule R') using \0 < e\ by simp then show "\\<^sub>F x in F. dist (f x) l < e" apply (rule eventually_mono) apply (subst euclidean_dist_l2) using * by blast qed qed corollary continuous_componentwise: "continuous F f \ (\i \ Basis. continuous F (\x. (f x \ i)))" by (simp add: continuous_def tendsto_componentwise_iff [symmetric]) corollary continuous_on_componentwise: fixes S :: "'a :: t2_space set" shows "continuous_on S f \ (\i \ Basis. continuous_on S (\x. (f x \ i)))" apply (simp add: continuous_on_eq_continuous_within) using continuous_componentwise by blast lemma linear_componentwise_iff: "(linear f') \ (\i\Basis. linear (\x. f' x \ i))" apply (auto simp: linear_iff inner_left_distrib) apply (metis inner_left_distrib euclidean_eq_iff) by (metis euclidean_eqI inner_scaleR_left) lemma bounded_linear_componentwise_iff: "(bounded_linear f') \ (\i\Basis. bounded_linear (\x. f' x \ i))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (simp add: bounded_linear_inner_left_comp) next assume ?rhs then have "(\i\Basis. \K. \x. \f' x \ i\ \ norm x * K)" "linear f'" by (auto simp: bounded_linear_def bounded_linear_axioms_def linear_componentwise_iff [symmetric] ball_conj_distrib) then obtain F where F: "\i x. i \ Basis \ \f' x \ i\ \ norm x * F i" by metis have "norm (f' x) \ norm x * sum F Basis" for x proof - have "norm (f' x) \ (\i\Basis. \f' x \ i\)" by (rule norm_le_l1) also have "... \ (\i\Basis. norm x * F i)" by (metis F sum_mono) also have "... = norm x * sum F Basis" by (simp add: sum_distrib_left) finally show ?thesis . qed then show ?lhs by (force simp: bounded_linear_def bounded_linear_axioms_def \linear f'\) qed subsection\<^marker>\tag unimportant\ \Continuous Extension\ definition clamp :: "'a::euclidean_space \ 'a \ 'a \ 'a" where "clamp a b x = (if (\i\Basis. a \ i \ b \ i) then (\i\Basis. (if x\i < a\i then a\i else if x\i \ b\i then x\i else b\i) *\<^sub>R i) else a)" lemma clamp_in_interval[simp]: assumes "\i. i \ Basis \ a \ i \ b \ i" shows "clamp a b x \ cbox a b" unfolding clamp_def using box_ne_empty(1)[of a b] assms by (auto simp: cbox_def) lemma clamp_cancel_cbox[simp]: fixes x a b :: "'a::euclidean_space" assumes x: "x \ cbox a b" shows "clamp a b x = x" using assms by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a]) lemma clamp_empty_interval: assumes "i \ Basis" "a \ i > b \ i" shows "clamp a b = (\_. a)" using assms by (force simp: clamp_def[abs_def] split: if_splits intro!: ext) lemma dist_clamps_le_dist_args: fixes x :: "'a::euclidean_space" shows "dist (clamp a b y) (clamp a b x) \ dist y x" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" then have "(\i\Basis. (dist (clamp a b y \ i) (clamp a b x \ i))\<^sup>2) \ (\i\Basis. (dist (y \ i) (x \ i))\<^sup>2)" by (auto intro!: sum_mono simp: clamp_def dist_real_def abs_le_square_iff[symmetric]) then show ?thesis by (auto intro: real_sqrt_le_mono simp: euclidean_dist_l2[where y=x] euclidean_dist_l2[where y="clamp a b x"] L2_set_def) qed (auto simp: clamp_def) lemma clamp_continuous_at: fixes f :: "'a::euclidean_space \ 'b::metric_space" and x :: 'a assumes f_cont: "continuous_on (cbox a b) f" shows "continuous (at x) (\x. f (clamp a b x))" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" show ?thesis unfolding continuous_at_eps_delta proof safe fix x :: 'a fix e :: real assume "e > 0" moreover have "clamp a b x \ cbox a b" by (simp add: clamp_in_interval le) moreover note f_cont[simplified continuous_on_iff] ultimately obtain d where d: "0 < d" "\x'. x' \ cbox a b \ dist x' (clamp a b x) < d \ dist (f x') (f (clamp a b x)) < e" by force show "\d>0. \x'. dist x' x < d \ dist (f (clamp a b x')) (f (clamp a b x)) < e" using le by (auto intro!: d clamp_in_interval dist_clamps_le_dist_args[THEN le_less_trans]) qed qed (auto simp: clamp_empty_interval) lemma clamp_continuous_on: fixes f :: "'a::euclidean_space \ 'b::metric_space" assumes f_cont: "continuous_on (cbox a b) f" shows "continuous_on S (\x. f (clamp a b x))" using assms by (auto intro: continuous_at_imp_continuous_on clamp_continuous_at) lemma clamp_bounded: fixes f :: "'a::euclidean_space \ 'b::metric_space" assumes bounded: "bounded (f ` (cbox a b))" shows "bounded (range (\x. f (clamp a b x)))" proof cases assume le: "(\i\Basis. a \ i \ b \ i)" from bounded obtain c where f_bound: "\x\f ` cbox a b. dist undefined x \ c" by (auto simp: bounded_any_center[where a=undefined]) then show ?thesis by (auto intro!: exI[where x=c] clamp_in_interval[OF le[rule_format]] simp: bounded_any_center[where a=undefined]) qed (auto simp: clamp_empty_interval image_def) definition ext_cont :: "('a::euclidean_space \ 'b::metric_space) \ 'a \ 'a \ 'a \ 'b" where "ext_cont f a b = (\x. f (clamp a b x))" lemma ext_cont_cancel_cbox[simp]: fixes x a b :: "'a::euclidean_space" assumes x: "x \ cbox a b" shows "ext_cont f a b x = f x" using assms unfolding ext_cont_def by (auto simp: clamp_def mem_box intro!: euclidean_eqI[where 'a='a] arg_cong[where f=f]) lemma continuous_on_ext_cont[continuous_intros]: "continuous_on (cbox a b) f \ continuous_on S (ext_cont f a b)" by (auto intro!: clamp_continuous_on simp: ext_cont_def) subsection \Separability\ lemma univ_second_countable_sequence: obtains B :: "nat \ 'a::euclidean_space set" where "inj B" "\n. open(B n)" "\S. open S \ \k. S = \{B n |n. n \ k}" proof - obtain \ :: "'a set set" where "countable \" and opn: "\C. C \ \ \ open C" and Un: "\S. open S \ \U. U \ \ \ S = \U" using univ_second_countable by blast have *: "infinite (range (\n. ball (0::'a) (inverse(Suc n))))" apply (rule Infinite_Set.range_inj_infinite) apply (simp add: inj_on_def ball_eq_ball_iff) done have "infinite \" proof assume "finite \" then have "finite (Union ` (Pow \))" by simp then have "finite (range (\n. ball (0::'a) (inverse(Suc n))))" apply (rule rev_finite_subset) by (metis (no_types, lifting) PowI image_eqI image_subset_iff Un [OF open_ball]) with * show False by simp qed obtain f :: "nat \ 'a set" where "\ = range f" "inj f" by (blast intro: countable_as_injective_image [OF \countable \\ \infinite \\]) have *: "\k. S = \{f n |n. n \ k}" if "open S" for S using Un [OF that] apply clarify apply (rule_tac x="f-`U" in exI) using \inj f\ \\ = range f\ apply force done show ?thesis apply (rule that [OF \inj f\ _ *]) apply (auto simp: \\ = range f\ opn) done qed proposition separable: fixes S :: "'a::{metric_space, second_countable_topology} set" obtains T where "countable T" "T \ S" "S \ closure T" proof - obtain \ :: "'a set set" where "countable \" and "{} \ \" and ope: "\C. C \ \ \ openin(top_of_set S) C" and if_ope: "\T. openin(top_of_set S) T \ \\. \ \ \ \ T = \\" by (meson subset_second_countable) then obtain f where f: "\C. C \ \ \ f C \ C" by (metis equals0I) show ?thesis proof show "countable (f ` \)" by (simp add: \countable \\) show "f ` \ \ S" using ope f openin_imp_subset by blast show "S \ closure (f ` \)" proof (clarsimp simp: closure_approachable) fix x and e::real assume "x \ S" "0 < e" have "openin (top_of_set S) (S \ ball x e)" by (simp add: openin_Int_open) with if_ope obtain \ where \: "\ \ \" "S \ ball x e = \\" by meson show "\C \ \. dist (f C) x < e" proof (cases "\ = {}") case True then show ?thesis using \0 < e\ \ \x \ S\ by auto next case False then obtain C where "C \ \" by blast show ?thesis proof show "dist (f C) x < e" by (metis Int_iff Union_iff \ \C \ \\ dist_commute f mem_ball subsetCE) show "C \ \" using \\ \ \\ \C \ \\ by blast qed qed qed qed qed subsection\<^marker>\tag unimportant\ \Diameter\ lemma diameter_cball [simp]: fixes a :: "'a::euclidean_space" shows "diameter(cball a r) = (if r < 0 then 0 else 2*r)" proof - have "diameter(cball a r) = 2*r" if "r \ 0" proof (rule order_antisym) show "diameter (cball a r) \ 2*r" proof (rule diameter_le) fix x y assume "x \ cball a r" "y \ cball a r" then have "norm (x - a) \ r" "norm (a - y) \ r" by (auto simp: dist_norm norm_minus_commute) then have "norm (x - y) \ r+r" using norm_diff_triangle_le by blast then show "norm (x - y) \ 2*r" by simp qed (simp add: that) have "2*r = dist (a + r *\<^sub>R (SOME i. i \ Basis)) (a - r *\<^sub>R (SOME i. i \ Basis))" apply (simp add: dist_norm) by (metis abs_of_nonneg mult.right_neutral norm_numeral norm_scaleR norm_some_Basis real_norm_def scaleR_2 that) also have "... \ diameter (cball a r)" apply (rule diameter_bounded_bound) using that by (auto simp: dist_norm) finally show "2*r \ diameter (cball a r)" . qed then show ?thesis by simp qed lemma diameter_ball [simp]: fixes a :: "'a::euclidean_space" shows "diameter(ball a r) = (if r < 0 then 0 else 2*r)" proof - have "diameter(ball a r) = 2*r" if "r > 0" by (metis bounded_ball diameter_closure closure_ball diameter_cball less_eq_real_def linorder_not_less that) then show ?thesis by (simp add: diameter_def) qed lemma diameter_closed_interval [simp]: "diameter {a..b} = (if b < a then 0 else b-a)" proof - have "{a .. b} = cball ((a+b)/2) ((b-a)/2)" by (auto simp: dist_norm abs_if field_split_simps split: if_split_asm) then show ?thesis by simp qed lemma diameter_open_interval [simp]: "diameter {a<..i \ Basis. a \ i \ b \ i) \ diameter (cbox a b) = dist a b" by (force simp: diameter_def intro!: cSup_eq_maximum L2_set_mono simp: euclidean_dist_l2[where 'a='a] cbox_def dist_norm) subsection\<^marker>\tag unimportant\\Relating linear images to open/closed/interior/closure/connected\ proposition open_surjective_linear_image: fixes f :: "'a::real_normed_vector \ 'b::euclidean_space" assumes "open A" "linear f" "surj f" shows "open(f ` A)" unfolding open_dist proof clarify fix x assume "x \ A" have "bounded (inv f ` Basis)" by (simp add: finite_imp_bounded) with bounded_pos obtain B where "B > 0" and B: "\x. x \ inv f ` Basis \ norm x \ B" by metis obtain e where "e > 0" and e: "\z. dist z x < e \ z \ A" by (metis open_dist \x \ A\ \open A\) define \ where "\ \ e / B / DIM('b)" show "\e>0. \y. dist y (f x) < e \ y \ f ` A" proof (intro exI conjI) show "\ > 0" using \e > 0\ \B > 0\ by (simp add: \_def field_split_simps) have "y \ f ` A" if "dist y (f x) * (B * real DIM('b)) < e" for y proof - define u where "u \ y - f x" show ?thesis proof (rule image_eqI) show "y = f (x + (\i\Basis. (u \ i) *\<^sub>R inv f i))" apply (simp add: linear_add linear_sum linear.scaleR \linear f\ surj_f_inv_f \surj f\) apply (simp add: euclidean_representation u_def) done have "dist (x + (\i\Basis. (u \ i) *\<^sub>R inv f i)) x \ (\i\Basis. norm ((u \ i) *\<^sub>R inv f i))" by (simp add: dist_norm sum_norm_le) also have "... = (\i\Basis. \u \ i\ * norm (inv f i))" by simp also have "... \ (\i\Basis. \u \ i\) * B" by (simp add: B sum_distrib_right sum_mono mult_left_mono) also have "... \ DIM('b) * dist y (f x) * B" apply (rule mult_right_mono [OF sum_bounded_above]) using \0 < B\ by (auto simp: Basis_le_norm dist_norm u_def) also have "... < e" by (metis mult.commute mult.left_commute that) finally show "x + (\i\Basis. (u \ i) *\<^sub>R inv f i) \ A" by (rule e) qed qed then show "\y. dist y (f x) < \ \ y \ f ` A" using \e > 0\ \B > 0\ by (auto simp: \_def field_split_simps mult_less_0_iff) qed qed corollary open_bijective_linear_image_eq: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "bij f" shows "open(f ` A) \ open A" proof assume "open(f ` A)" then have "open(f -` (f ` A))" using assms by (force simp: linear_continuous_at linear_conv_bounded_linear continuous_open_vimage) then show "open A" by (simp add: assms bij_is_inj inj_vimage_image_eq) next assume "open A" then show "open(f ` A)" by (simp add: assms bij_is_surj open_surjective_linear_image) qed corollary interior_bijective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "bij f" shows "interior (f ` S) = f ` interior S" (is "?lhs = ?rhs") proof safe fix x assume x: "x \ ?lhs" then obtain T where "open T" and "x \ T" and "T \ f ` S" by (metis interiorE) then show "x \ ?rhs" by (metis (no_types, hide_lams) assms subsetD interior_maximal open_bijective_linear_image_eq subset_image_iff) next fix x assume x: "x \ interior S" then show "f x \ interior (f ` S)" by (meson assms imageI image_mono interiorI interior_subset open_bijective_linear_image_eq open_interior) qed lemma interior_injective_linear_image: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "linear f" "inj f" shows "interior(f ` S) = f ` (interior S)" by (simp add: linear_injective_imp_surjective assms bijI interior_bijective_linear_image) lemma interior_surjective_linear_image: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "linear f" "surj f" shows "interior(f ` S) = f ` (interior S)" by (simp add: assms interior_injective_linear_image linear_surjective_imp_injective) lemma interior_negations: fixes S :: "'a::euclidean_space set" shows "interior(uminus ` S) = image uminus (interior S)" by (simp add: bij_uminus interior_bijective_linear_image linear_uminus) lemma connected_linear_image: fixes f :: "'a::euclidean_space \ 'b::real_normed_vector" assumes "linear f" and "connected s" shows "connected (f ` s)" using connected_continuous_image assms linear_continuous_on linear_conv_bounded_linear by blast subsection\<^marker>\tag unimportant\ \"Isometry" (up to constant bounds) of Injective Linear Map\ proposition injective_imp_isometric: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes s: "closed s" "subspace s" and f: "bounded_linear f" "\x\s. f x = 0 \ x = 0" shows "\e>0. \x\s. norm (f x) \ e * norm x" proof (cases "s \ {0::'a}") case True have "norm x \ norm (f x)" if "x \ s" for x proof - from True that have "x = 0" by auto then show ?thesis by simp qed then show ?thesis by (auto intro!: exI[where x=1]) next case False interpret f: bounded_linear f by fact from False obtain a where a: "a \ 0" "a \ s" by auto from False have "s \ {}" by auto let ?S = "{f x| x. x \ s \ norm x = norm a}" let ?S' = "{x::'a. x\s \ norm x = norm a}" let ?S'' = "{x::'a. norm x = norm a}" have "?S'' = frontier (cball 0 (norm a))" by (simp add: sphere_def dist_norm) then have "compact ?S''" by (metis compact_cball compact_frontier) moreover have "?S' = s \ ?S''" by auto ultimately have "compact ?S'" using closed_Int_compact[of s ?S''] using s(1) by auto moreover have *:"f ` ?S' = ?S" by auto ultimately have "compact ?S" using compact_continuous_image[OF linear_continuous_on[OF f(1)], of ?S'] by auto then have "closed ?S" using compact_imp_closed by auto moreover from a have "?S \ {}" by auto ultimately obtain b' where "b'\?S" "\y\?S. norm b' \ norm y" using distance_attains_inf[of ?S 0] unfolding dist_0_norm by auto then obtain b where "b\s" and ba: "norm b = norm a" and b: "\x\{x \ s. norm x = norm a}. norm (f b) \ norm (f x)" unfolding *[symmetric] unfolding image_iff by auto let ?e = "norm (f b) / norm b" have "norm b > 0" using ba and a and norm_ge_zero by auto moreover have "norm (f b) > 0" using f(2)[THEN bspec[where x=b], OF \b\s\] using \norm b >0\ by simp ultimately have "0 < norm (f b) / norm b" by simp moreover have "norm (f b) / norm b * norm x \ norm (f x)" if "x\s" for x proof (cases "x = 0") case True then show "norm (f b) / norm b * norm x \ norm (f x)" by auto next case False with \a \ 0\ have *: "0 < norm a / norm x" unfolding zero_less_norm_iff[symmetric] by simp have "\x\s. c *\<^sub>R x \ s" for c using s[unfolded subspace_def] by simp with \x \ s\ \x \ 0\ have "(norm a / norm x) *\<^sub>R x \ {x \ s. norm x = norm a}" by simp with \x \ 0\ \a \ 0\ show "norm (f b) / norm b * norm x \ norm (f x)" using b[THEN bspec[where x="(norm a / norm x) *\<^sub>R x"]] unfolding f.scaleR and ba by (auto simp: mult.commute pos_le_divide_eq pos_divide_le_eq) qed ultimately show ?thesis by auto qed proposition closed_injective_image_subspace: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace s" "bounded_linear f" "\x\s. f x = 0 \ x = 0" "closed s" shows "closed(f ` s)" proof - obtain e where "e > 0" and e: "\x\s. e * norm x \ norm (f x)" using injective_imp_isometric[OF assms(4,1,2,3)] by auto show ?thesis using complete_isometric_image[OF \e>0\ assms(1,2) e] and assms(4) unfolding complete_eq_closed[symmetric] by auto qed subsection\<^marker>\tag unimportant\ \Some properties of a canonical subspace\ lemma closed_substandard: "closed {x::'a::euclidean_space. \i\Basis. P i \ x\i = 0}" (is "closed ?A") proof - let ?D = "{i\Basis. P i}" have "closed (\i\?D. {x::'a. x\i = 0})" by (simp add: closed_INT closed_Collect_eq continuous_on_inner continuous_on_const continuous_on_id) also have "(\i\?D. {x::'a. x\i = 0}) = ?A" by auto finally show "closed ?A" . qed lemma closed_subspace: fixes s :: "'a::euclidean_space set" assumes "subspace s" shows "closed s" proof - have "dim s \ card (Basis :: 'a set)" using dim_subset_UNIV by auto with ex_card[OF this] obtain d :: "'a set" where t: "card d = dim s" and d: "d \ Basis" by auto let ?t = "{x::'a. \i\Basis. i \ d \ x\i = 0}" have "\f. linear f \ f ` {x::'a. \i\Basis. i \ d \ x \ i = 0} = s \ inj_on f {x::'a. \i\Basis. i \ d \ x \ i = 0}" using dim_substandard[of d] t d assms by (intro subspace_isomorphism[OF subspace_substandard[of "\i. i \ d"]]) (auto simp: inner_Basis) then obtain f where f: "linear f" "f ` {x. \i\Basis. i \ d \ x \ i = 0} = s" "inj_on f {x. \i\Basis. i \ d \ x \ i = 0}" by blast interpret f: bounded_linear f using f by (simp add: linear_conv_bounded_linear) have "x \ ?t \ f x = 0 \ x = 0" for x using f.zero d f(3)[THEN inj_onD, of x 0] by auto moreover have "closed ?t" by (rule closed_substandard) moreover have "subspace ?t" by (rule subspace_substandard) ultimately show ?thesis using closed_injective_image_subspace[of ?t f] unfolding f(2) using f(1) unfolding linear_conv_bounded_linear by auto qed lemma complete_subspace: "subspace s \ complete s" for s :: "'a::euclidean_space set" using complete_eq_closed closed_subspace by auto lemma closed_span [iff]: "closed (span s)" for s :: "'a::euclidean_space set" by (simp add: closed_subspace subspace_span) lemma dim_closure [simp]: "dim (closure s) = dim s" (is "?dc = ?d") for s :: "'a::euclidean_space set" proof - have "?dc \ ?d" using closure_minimal[OF span_superset, of s] using closed_subspace[OF subspace_span, of s] using dim_subset[of "closure s" "span s"] by simp then show ?thesis using dim_subset[OF closure_subset, of s] by simp qed subsection \Set Distance\ lemma setdist_compact_closed: fixes A :: "'a::heine_borel set" assumes A: "compact A" and B: "closed B" and "A \ {}" "B \ {}" shows "\x \ A. \y \ B. dist x y = setdist A B" proof - obtain x where "x \ A" "setdist A B = infdist x B" by (metis A assms(3) setdist_attains_inf setdist_sym) moreover obtain y where"y \ B" "infdist x B = dist x y" using B \B \ {}\ infdist_attains_inf by blast ultimately show ?thesis using \x \ A\ \y \ B\ by auto qed lemma setdist_closed_compact: fixes S :: "'a::heine_borel set" assumes S: "closed S" and T: "compact T" and "S \ {}" "T \ {}" shows "\x \ S. \y \ T. dist x y = setdist S T" using setdist_compact_closed [OF T S \T \ {}\ \S \ {}\] by (metis dist_commute setdist_sym) lemma setdist_eq_0_compact_closed: assumes S: "compact S" and T: "closed T" shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" proof (cases "S = {} \ T = {}") case True then show ?thesis by force next case False then show ?thesis by (metis S T disjoint_iff_not_equal in_closed_iff_infdist_zero setdist_attains_inf setdist_eq_0I setdist_sym) qed corollary setdist_gt_0_compact_closed: assumes S: "compact S" and T: "closed T" shows "setdist S T > 0 \ (S \ {} \ T \ {} \ S \ T = {})" using setdist_pos_le [of S T] setdist_eq_0_compact_closed [OF assms] by linarith lemma setdist_eq_0_closed_compact: assumes S: "closed S" and T: "compact T" shows "setdist S T = 0 \ S = {} \ T = {} \ S \ T \ {}" using setdist_eq_0_compact_closed [OF T S] by (metis Int_commute setdist_sym) lemma setdist_eq_0_bounded: fixes S :: "'a::heine_borel set" assumes "bounded S \ bounded T" shows "setdist S T = 0 \ S = {} \ T = {} \ closure S \ closure T \ {}" proof (cases "S = {} \ T = {}") case False then show ?thesis using setdist_eq_0_compact_closed [of "closure S" "closure T"] setdist_eq_0_closed_compact [of "closure S" "closure T"] assms by (force simp: bounded_closure compact_eq_bounded_closed) qed force lemma setdist_eq_0_sing_1: "setdist {x} S = 0 \ S = {} \ x \ closure S" by (metis in_closure_iff_infdist_zero infdist_def infdist_eq_setdist) lemma setdist_eq_0_sing_2: "setdist S {x} = 0 \ S = {} \ x \ closure S" by (metis setdist_eq_0_sing_1 setdist_sym) lemma setdist_neq_0_sing_1: "\setdist {x} S = a; a \ 0\ \ S \ {} \ x \ closure S" by (metis setdist_closure_2 setdist_empty2 setdist_eq_0I singletonI) lemma setdist_neq_0_sing_2: "\setdist S {x} = a; a \ 0\ \ S \ {} \ x \ closure S" by (simp add: setdist_neq_0_sing_1 setdist_sym) lemma setdist_sing_in_set: "x \ S \ setdist {x} S = 0" by (simp add: setdist_eq_0I) lemma setdist_eq_0_closed: "closed S \ (setdist {x} S = 0 \ S = {} \ x \ S)" by (simp add: setdist_eq_0_sing_1) lemma setdist_eq_0_closedin: shows "\closedin (top_of_set U) S; x \ U\ \ (setdist {x} S = 0 \ S = {} \ x \ S)" by (auto simp: closedin_limpt setdist_eq_0_sing_1 closure_def) lemma setdist_gt_0_closedin: shows "\closedin (top_of_set U) S; x \ U; S \ {}; x \ S\ \ setdist {x} S > 0" using less_eq_real_def setdist_eq_0_closedin by fastforce no_notation eucl_less (infix "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\ -(* TODO: Move *) -lemma closed_segmentI: - "u \ {0..1} \ z = (1 - u) *\<^sub>R a + u *\<^sub>R b \ z \ closed_segment a b" - by (auto simp: closed_segment_def) - -lemma in_cbox_complex_iff: - "x \ cbox a b \ Re x \ {Re a..Re b} \ Im x \ {Im a..Im b}" - by (cases x; cases a; cases b) (auto simp: cbox_Complex_eq) - -lemma box_Complex_eq: - "box (Complex a c) (Complex b d) = (\(x,y). Complex x y) ` (box a b \ box c d)" - by (auto simp: box_def Basis_complex_def image_iff complex_eq_iff) - -lemma in_box_complex_iff: - "x \ box a b \ Re x \ {Re a<.. Im x \ {Im 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 (intro closed_segmentI[of u]) (auto simp: u 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 (intro closed_segmentI[of u]) (auto simp: u 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 (intro closed_segmentI[of u]) (auto simp: u 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 (intro closed_segmentI[of u]) (auto simp: u algebra_simps scaleR_conv_of_real complex_eq_iff) -qed - 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/Deriv.thy b/src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy +++ b/src/HOL/Deriv.thy @@ -1,2291 +1,2308 @@ (* Title: HOL/Deriv.thy Author: Jacques D. Fleuriot, University of Cambridge, 1998 Author: Brian Huffman Author: Lawrence C Paulson, 2004 Author: Benjamin Porter, 2005 *) section \Differentiation\ theory Deriv imports Limits begin subsection \Frechet derivative\ definition has_derivative :: "('a::real_normed_vector \ 'b::real_normed_vector) \ ('a \ 'b) \ 'a filter \ bool" (infix "(has'_derivative)" 50) where "(f has_derivative f') F \ bounded_linear f' \ ((\y. ((f y - f (Lim F (\x. x))) - f' (y - Lim F (\x. x))) /\<^sub>R norm (y - Lim F (\x. x))) \ 0) F" text \ Usually the filter \<^term>\F\ is \<^term>\at x within s\. \<^term>\(f has_derivative D) (at x within s)\ means: \<^term>\D\ is the derivative of function \<^term>\f\ at point \<^term>\x\ within the set \<^term>\s\. Where \<^term>\s\ is used to express left or right sided derivatives. In most cases \<^term>\s\ is either a variable or \<^term>\UNIV\. \ text \These are the only cases we'll care about, probably.\ lemma has_derivative_within: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. (1 / norm(y - x)) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x within s)" unfolding has_derivative_def tendsto_iff by (subst eventually_Lim_ident_at) (auto simp add: field_simps) lemma has_derivative_eq_rhs: "(f has_derivative f') F \ f' = g' \ (f has_derivative g') F" by simp definition has_field_derivative :: "('a::real_normed_field \ 'a) \ 'a \ 'a filter \ bool" (infix "(has'_field'_derivative)" 50) where "(f has_field_derivative D) F \ (f has_derivative (*) D) F" lemma DERIV_cong: "(f has_field_derivative X) F \ X = Y \ (f has_field_derivative Y) F" by simp definition has_vector_derivative :: "(real \ 'b::real_normed_vector) \ 'b \ real filter \ bool" (infix "has'_vector'_derivative" 50) where "(f has_vector_derivative f') net \ (f has_derivative (\x. x *\<^sub>R f')) net" lemma has_vector_derivative_eq_rhs: "(f has_vector_derivative X) F \ X = Y \ (f has_vector_derivative Y) F" by simp named_theorems derivative_intros "structural introduction rules for derivatives" setup \ let val eq_thms = @{thms has_derivative_eq_rhs DERIV_cong has_vector_derivative_eq_rhs} fun eq_rule thm = get_first (try (fn eq_thm => eq_thm OF [thm])) eq_thms in Global_Theory.add_thms_dynamic (\<^binding>\derivative_eq_intros\, fn context => Named_Theorems.get (Context.proof_of context) \<^named_theorems>\derivative_intros\ |> map_filter eq_rule) end \ text \ The following syntax is only used as a legacy syntax. \ abbreviation (input) FDERIV :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a \ ('a \ 'b) \ bool" ("(FDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "FDERIV f x :> f' \ (f has_derivative f') (at x)" lemma has_derivative_bounded_linear: "(f has_derivative f') F \ bounded_linear f'" by (simp add: has_derivative_def) lemma has_derivative_linear: "(f has_derivative f') F \ linear f'" using bounded_linear.linear[OF has_derivative_bounded_linear] . lemma has_derivative_ident[derivative_intros, simp]: "((\x. x) has_derivative (\x. x)) F" by (simp add: has_derivative_def) lemma has_derivative_id [derivative_intros, simp]: "(id has_derivative id) (at a)" by (metis eq_id_iff has_derivative_ident) lemma has_derivative_const[derivative_intros, simp]: "((\x. c) has_derivative (\x. 0)) F" by (simp add: has_derivative_def) lemma (in bounded_linear) bounded_linear: "bounded_linear f" .. lemma (in bounded_linear) has_derivative: "(g has_derivative g') F \ ((\x. f (g x)) has_derivative (\x. f (g' x))) F" unfolding has_derivative_def by (auto simp add: bounded_linear_compose [OF bounded_linear] scaleR diff dest: tendsto) lemmas has_derivative_scaleR_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_right] lemmas has_derivative_scaleR_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_scaleR_left] lemmas has_derivative_mult_right [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_right] lemmas has_derivative_mult_left [derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_mult_left] lemmas has_derivative_of_real[derivative_intros, simp] = bounded_linear.has_derivative[OF bounded_linear_of_real] lemma has_derivative_add[simp, derivative_intros]: assumes f: "(f has_derivative f') F" and g: "(g has_derivative g') F" shows "((\x. f x + g x) has_derivative (\x. f' x + g' x)) F" unfolding has_derivative_def proof safe let ?x = "Lim F (\x. x)" let ?D = "\f f' y. ((f y - f ?x) - f' (y - ?x)) /\<^sub>R norm (y - ?x)" have "((\x. ?D f f' x + ?D g g' x) \ (0 + 0)) F" using f g by (intro tendsto_add) (auto simp: has_derivative_def) then show "(?D (\x. f x + g x) (\x. f' x + g' x) \ 0) F" by (simp add: field_simps scaleR_add_right scaleR_diff_right) qed (blast intro: bounded_linear_add f g has_derivative_bounded_linear) lemma has_derivative_sum[simp, derivative_intros]: "(\i. i \ I \ (f i has_derivative f' i) F) \ ((\x. \i\I. f i x) has_derivative (\x. \i\I. f' i x)) F" by (induct I rule: infinite_finite_induct) simp_all lemma has_derivative_minus[simp, derivative_intros]: "(f has_derivative f') F \ ((\x. - f x) has_derivative (\x. - f' x)) F" using has_derivative_scaleR_right[of f f' F "-1"] by simp lemma has_derivative_diff[simp, derivative_intros]: "(f has_derivative f') F \ (g has_derivative g') F \ ((\x. f x - g x) has_derivative (\x. f' x - g' x)) F" by (simp only: diff_conv_add_uminus has_derivative_add has_derivative_minus) lemma has_derivative_at_within: "(f has_derivative f') (at x within s) \ (bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s))" by (cases "at x within s = bot") (simp_all add: has_derivative_def Lim_ident_at) lemma has_derivative_iff_norm: "(f has_derivative f') (at x within s) \ bounded_linear f' \ ((\y. norm ((f y - f x) - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" using tendsto_norm_zero_iff[of _ "at x within s", where 'b="'b", symmetric] by (simp add: has_derivative_at_within divide_inverse ac_simps) lemma has_derivative_at: "(f has_derivative D) (at x) \ (bounded_linear D \ (\h. norm (f (x + h) - f x - D h) / norm h) \0\ 0)" unfolding has_derivative_iff_norm LIM_offset_zero_iff[of _ _ x] by simp lemma field_has_derivative_at: fixes x :: "'a::real_normed_field" shows "(f has_derivative (*) D) (at x) \ (\h. (f (x + h) - f x) / h) \0\ D" (is "?lhs = ?rhs") proof - have "?lhs = (\h. norm (f (x + h) - f x - D * h) / norm h) \0 \ 0" by (simp add: bounded_linear_mult_right has_derivative_at) also have "... = (\y. norm ((f (x + y) - f x - D * y) / y)) \0\ 0" by (simp cong: LIM_cong flip: nonzero_norm_divide) also have "... = (\y. norm ((f (x + y) - f x) / y - D / y * y)) \0\ 0" by (simp only: diff_divide_distrib times_divide_eq_left [symmetric]) also have "... = ?rhs" by (simp add: tendsto_norm_zero_iff LIM_zero_iff cong: LIM_cong) finally show ?thesis . qed lemma has_derivative_iff_Ex: "(f has_derivative f') (at x) \ bounded_linear f' \ (\e. (\h. f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" unfolding has_derivative_at by force lemma has_derivative_at_within_iff_Ex: assumes "x \ S" "open S" shows "(f has_derivative f') (at x within S) \ bounded_linear f' \ (\e. (\h. x+h \ S \ f (x+h) = f x + f' h + e h) \ ((\h. norm (e h) / norm h) \ 0) (at 0))" (is "?lhs = ?rhs") proof safe show "bounded_linear f'" if "(f has_derivative f') (at x within S)" using has_derivative_bounded_linear that by blast show "\e. (\h. x + h \ S \ f (x + h) = f x + f' h + e h) \ (\h. norm (e h) / norm h) \0\ 0" if "(f has_derivative f') (at x within S)" by (metis (full_types) assms that has_derivative_iff_Ex at_within_open) show "(f has_derivative f') (at x within S)" if "bounded_linear f'" and eq [rule_format]: "\h. x + h \ S \ f (x + h) = f x + f' h + e h" and 0: "(\h. norm (e (h::'a)::'b) / norm h) \0\ 0" for e proof - have 1: "f y - f x = f' (y-x) + e (y-x)" if "y \ S" for y using eq [of "y-x"] that by simp have 2: "((\y. norm (e (y-x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: "0" assms tendsto_offset_zero_iff) have "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within S)" by (simp add: Lim_cong_within 1 2) then show ?thesis by (simp add: has_derivative_iff_norm \bounded_linear f'\) qed qed lemma has_derivativeI: "bounded_linear f' \ ((\y. ((f y - f x) - f' (y - x)) /\<^sub>R norm (y - x)) \ 0) (at x within s) \ (f has_derivative f') (at x within s)" by (simp add: has_derivative_at_within) lemma has_derivativeI_sandwich: assumes e: "0 < e" and bounded: "bounded_linear f'" and sandwich: "(\y. y \ s \ y \ x \ dist y x < e \ norm ((f y - f x) - f' (y - x)) / norm (y - x) \ H y)" and "(H \ 0) (at x within s)" shows "(f has_derivative f') (at x within s)" unfolding has_derivative_iff_norm proof safe show "((\y. norm (f y - f x - f' (y - x)) / norm (y - x)) \ 0) (at x within s)" proof (rule tendsto_sandwich[where f="\x. 0"]) show "(H \ 0) (at x within s)" by fact show "eventually (\n. norm (f n - f x - f' (n - x)) / norm (n - x) \ H n) (at x within s)" unfolding eventually_at using e sandwich by auto qed (auto simp: le_divide_eq) qed fact lemma has_derivative_subset: "(f has_derivative f') (at x within s) \ t \ s \ (f has_derivative f') (at x within t)" by (auto simp add: has_derivative_iff_norm intro: tendsto_within_subset) lemmas has_derivative_within_subset = has_derivative_subset lemma has_derivative_within_singleton_iff: "(f has_derivative g) (at x within {x}) \ bounded_linear g" by (auto intro!: has_derivativeI_sandwich[where e=1] has_derivative_bounded_linear) subsubsection \Limit transformation for derivatives\ lemma has_derivative_transform_within: assumes "(f has_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_derivative f') (at x within s)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within) lemma has_derivative_transform_within_open: assumes "(f has_derivative f') (at x within t)" and "open s" and "x \ s" and "\x. x\s \ f x = g x" shows "(g has_derivative f') (at x within t)" using assms unfolding has_derivative_within by (force simp add: intro: Lim_transform_within_open) lemma has_derivative_transform: assumes "x \ s" "\x. x \ s \ g x = f x" assumes "(f has_derivative f') (at x within s)" shows "(g has_derivative f') (at x within s)" using assms by (intro has_derivative_transform_within[OF _ zero_less_one, where g=g]) auto lemma has_derivative_transform_eventually: assumes "(f has_derivative f') (at x within s)" "(\\<^sub>F x' in at x within s. f x' = g x')" assumes "f x = g x" "x \ s" shows "(g has_derivative f') (at x within s)" using assms proof - from assms(2,3) obtain d where "d > 0" "\x'. x' \ s \ dist x' x < d \ f x' = g x'" by (force simp: eventually_at) from has_derivative_transform_within[OF assms(1) this(1) assms(4) this(2)] show ?thesis . qed +lemma has_field_derivative_transform_within: + assumes "(f has_field_derivative f') (at a within S)" + and "0 < d" + and "a \ S" + and "\x. \x \ S; dist x a < d\ \ f x = g x" + shows "(g has_field_derivative f') (at a within S)" + using assms unfolding has_field_derivative_def + by (metis has_derivative_transform_within) + +lemma has_field_derivative_transform_within_open: + assumes "(f has_field_derivative f') (at a)" + and "open S" "a \ S" + and "\x. x \ S \ f x = g x" + shows "(g has_field_derivative f') (at a)" + using assms unfolding has_field_derivative_def + by (metis has_derivative_transform_within_open) + subsection \Continuity\ lemma has_derivative_continuous: assumes f: "(f has_derivative f') (at x within s)" shows "continuous (at x within s) f" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) note F.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x))" using f unfolding has_derivative_iff_norm by blast then have "?L (\y. norm ((f y - f x) - f' (y - x)) / norm (y - x) * norm (y - x))" (is ?m) by (rule tendsto_mult_zero) (auto intro!: tendsto_eq_intros) also have "?m \ ?L (\y. norm ((f y - f x) - f' (y - x)))" by (intro filterlim_cong) (simp_all add: eventually_at_filter) finally have "?L (\y. (f y - f x) - f' (y - x))" by (rule tendsto_norm_zero_cancel) then have "?L (\y. ((f y - f x) - f' (y - x)) + f' (y - x))" by (rule tendsto_eq_intros) (auto intro!: tendsto_eq_intros simp: F.zero) then have "?L (\y. f y - f x)" by simp from tendsto_add[OF this tendsto_const, of "f x"] show ?thesis by (simp add: continuous_within) qed subsection \Composition\ lemma tendsto_at_iff_tendsto_nhds_within: "f x = y \ (f \ y) (at x within s) \ (f \ y) (inf (nhds x) (principal s))" unfolding tendsto_def eventually_inf_principal eventually_at_filter by (intro ext all_cong imp_cong) (auto elim!: eventually_mono) lemma has_derivative_in_compose: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at (f x) within (f`s))" shows "((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" proof - from f interpret F: bounded_linear f' by (rule has_derivative_bounded_linear) from g interpret G: bounded_linear g' by (rule has_derivative_bounded_linear) from F.bounded obtain kF where kF: "\x. norm (f' x) \ norm x * kF" by fast from G.bounded obtain kG where kG: "\x. norm (g' x) \ norm x * kG" by fast note G.tendsto[tendsto_intros] let ?L = "\f. (f \ 0) (at x within s)" let ?D = "\f f' x y. (f y - f x) - f' (y - x)" let ?N = "\f f' x y. norm (?D f f' x y) / norm (y - x)" let ?gf = "\x. g (f x)" and ?gf' = "\x. g' (f' x)" define Nf where "Nf = ?N f f' x" define Ng where [abs_def]: "Ng y = ?N g g' (f x) (f y)" for y show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\x. g' (f' x))" using f g by (blast intro: bounded_linear_compose has_derivative_bounded_linear) next fix y :: 'a assume neq: "y \ x" have "?N ?gf ?gf' x y = norm (g' (?D f f' x y) + ?D g g' (f x) (f y)) / norm (y - x)" by (simp add: G.diff G.add field_simps) also have "\ \ norm (g' (?D f f' x y)) / norm (y - x) + Ng y * (norm (f y - f x) / norm (y - x))" by (simp add: add_divide_distrib[symmetric] divide_right_mono norm_triangle_ineq G.zero Ng_def) also have "\ \ Nf y * kG + Ng y * (Nf y + kF)" proof (intro add_mono mult_left_mono) have "norm (f y - f x) = norm (?D f f' x y + f' (y - x))" by simp also have "\ \ norm (?D f f' x y) + norm (f' (y - x))" by (rule norm_triangle_ineq) also have "\ \ norm (?D f f' x y) + norm (y - x) * kF" using kF by (intro add_mono) simp finally show "norm (f y - f x) / norm (y - x) \ Nf y + kF" by (simp add: neq Nf_def field_simps) qed (use kG in \simp_all add: Ng_def Nf_def neq zero_le_divide_iff field_simps\) finally show "?N ?gf ?gf' x y \ Nf y * kG + Ng y * (Nf y + kF)" . next have [tendsto_intros]: "?L Nf" using f unfolding has_derivative_iff_norm Nf_def .. from f have "(f \ f x) (at x within s)" by (blast intro: has_derivative_continuous continuous_within[THEN iffD1]) then have f': "LIM x at x within s. f x :> inf (nhds (f x)) (principal (f`s))" unfolding filterlim_def by (simp add: eventually_filtermap eventually_at_filter le_principal) have "((?N g g' (f x)) \ 0) (at (f x) within f`s)" using g unfolding has_derivative_iff_norm .. then have g': "((?N g g' (f x)) \ 0) (inf (nhds (f x)) (principal (f`s)))" by (rule tendsto_at_iff_tendsto_nhds_within[THEN iffD1, rotated]) simp have [tendsto_intros]: "?L Ng" unfolding Ng_def by (rule filterlim_compose[OF g' f']) show "((\y. Nf y * kG + Ng y * (Nf y + kF)) \ 0) (at x within s)" by (intro tendsto_eq_intros) auto qed simp qed lemma has_derivative_compose: "(f has_derivative f') (at x within s) \ (g has_derivative g') (at (f x)) \ ((\x. g (f x)) has_derivative (\x. g' (f' x))) (at x within s)" by (blast intro: has_derivative_in_compose has_derivative_subset) lemma has_derivative_in_compose2: assumes "\x. x \ t \ (g has_derivative g' x) (at x within t)" assumes "f ` s \ t" "x \ s" assumes "(f has_derivative f') (at x within s)" shows "((\x. g (f x)) has_derivative (\y. g' (f x) (f' y))) (at x within s)" using assms by (auto intro: has_derivative_within_subset intro!: has_derivative_in_compose[of f f' x s g]) lemma (in bounded_bilinear) FDERIV: assumes f: "(f has_derivative f') (at x within s)" and g: "(g has_derivative g') (at x within s)" shows "((\x. f x ** g x) has_derivative (\h. f x ** g' h + f' h ** g x)) (at x within s)" proof - from bounded_linear.bounded [OF has_derivative_bounded_linear [OF f]] obtain KF where norm_F: "\x. norm (f' x) \ norm x * KF" by fast from pos_bounded obtain K where K: "0 < K" and norm_prod: "\a b. norm (a ** b) \ norm a * norm b * K" by fast let ?D = "\f f' y. f y - f x - f' (y - x)" let ?N = "\f f' y. norm (?D f f' y) / norm (y - x)" define Ng where "Ng = ?N g g'" define Nf where "Nf = ?N f f'" let ?fun1 = "\y. norm (f y ** g y - f x ** g x - (f x ** g' (y - x) + f' (y - x) ** g x)) / norm (y - x)" let ?fun2 = "\y. norm (f x) * Ng y * K + Nf y * norm (g y) * K + KF * norm (g y - g x) * K" let ?F = "at x within s" show ?thesis proof (rule has_derivativeI_sandwich[of 1]) show "bounded_linear (\h. f x ** g' h + f' h ** g x)" by (intro bounded_linear_add bounded_linear_compose [OF bounded_linear_right] bounded_linear_compose [OF bounded_linear_left] has_derivative_bounded_linear [OF g] has_derivative_bounded_linear [OF f]) next from g have "(g \ g x) ?F" by (intro continuous_within[THEN iffD1] has_derivative_continuous) moreover from f g have "(Nf \ 0) ?F" "(Ng \ 0) ?F" by (simp_all add: has_derivative_iff_norm Ng_def Nf_def) ultimately have "(?fun2 \ norm (f x) * 0 * K + 0 * norm (g x) * K + KF * norm (0::'b) * K) ?F" by (intro tendsto_intros) (simp_all add: LIM_zero_iff) then show "(?fun2 \ 0) ?F" by simp next fix y :: 'd assume "y \ x" have "?fun1 y = norm (f x ** ?D g g' y + ?D f f' y ** g y + f' (y - x) ** (g y - g x)) / norm (y - x)" by (simp add: diff_left diff_right add_left add_right field_simps) also have "\ \ (norm (f x) * norm (?D g g' y) * K + norm (?D f f' y) * norm (g y) * K + norm (y - x) * KF * norm (g y - g x) * K) / norm (y - x)" by (intro divide_right_mono mult_mono' order_trans [OF norm_triangle_ineq add_mono] order_trans [OF norm_prod mult_right_mono] mult_nonneg_nonneg order_refl norm_ge_zero norm_F K [THEN order_less_imp_le]) also have "\ = ?fun2 y" by (simp add: add_divide_distrib Ng_def Nf_def) finally show "?fun1 y \ ?fun2 y" . qed simp qed lemmas has_derivative_mult[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_mult] lemmas has_derivative_scaleR[simp, derivative_intros] = bounded_bilinear.FDERIV[OF bounded_bilinear_scaleR] lemma has_derivative_prod[simp, derivative_intros]: fixes f :: "'i \ 'a::real_normed_vector \ 'b::real_normed_field" shows "(\i. i \ I \ (f i has_derivative f' i) (at x within S)) \ ((\x. \i\I. f i x) has_derivative (\y. \i\I. f' i y * (\j\I - {i}. f j x))) (at x within S)" proof (induct I rule: infinite_finite_induct) case infinite then show ?case by simp next case empty then show ?case by simp next case (insert i I) let ?P = "\y. f i x * (\i\I. f' i y * (\j\I - {i}. f j x)) + (f' i y) * (\i\I. f i x)" have "((\x. f i x * (\i\I. f i x)) has_derivative ?P) (at x within S)" using insert by (intro has_derivative_mult) auto also have "?P = (\y. \i'\insert i I. f' i' y * (\j\insert i I - {i'}. f j x))" using insert(1,2) by (auto simp add: sum_distrib_left insert_Diff_if intro!: ext sum.cong) finally show ?case using insert by simp qed lemma has_derivative_power[simp, derivative_intros]: fixes f :: "'a :: real_normed_vector \ 'b :: real_normed_field" assumes f: "(f has_derivative f') (at x within S)" shows "((\x. f x^n) has_derivative (\y. of_nat n * f' y * f x^(n - 1))) (at x within S)" using has_derivative_prod[OF f, of "{..< n}"] by (simp add: prod_constant ac_simps) lemma has_derivative_inverse': fixes x :: "'a::real_normed_div_algebra" assumes x: "x \ 0" shows "(inverse has_derivative (\h. - (inverse x * h * inverse x))) (at x within S)" (is "(_ has_derivative ?f) _") proof (rule has_derivativeI_sandwich) show "bounded_linear (\h. - (inverse x * h * inverse x))" by (simp add: bounded_linear_minus bounded_linear_mult_const bounded_linear_mult_right) show "0 < norm x" using x by simp have "(inverse \ inverse x) (at x within S)" using tendsto_inverse tendsto_ident_at x by auto then show "((\y. norm (inverse y - inverse x) * norm (inverse x)) \ 0) (at x within S)" by (simp add: LIM_zero_iff tendsto_mult_left_zero tendsto_norm_zero) next fix y :: 'a assume h: "y \ x" "dist y x < norm x" then have "y \ 0" by auto have "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) = norm (- (inverse y * (y - x) * inverse x - inverse x * (y - x) * inverse x)) / norm (y - x)" by (simp add: \y \ 0\ inverse_diff_inverse x) also have "... = norm ((inverse y - inverse x) * (y - x) * inverse x) / norm (y - x)" by (simp add: left_diff_distrib norm_minus_commute) also have "\ \ norm (inverse y - inverse x) * norm (y - x) * norm (inverse x) / norm (y - x)" by (simp add: norm_mult) also have "\ = norm (inverse y - inverse x) * norm (inverse x)" by simp finally show "norm (inverse y - inverse x - ?f (y -x)) / norm (y - x) \ norm (inverse y - inverse x) * norm (inverse x)" . qed lemma has_derivative_inverse[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes x: "f x \ 0" and f: "(f has_derivative f') (at x within S)" shows "((\x. inverse (f x)) has_derivative (\h. - (inverse (f x) * f' h * inverse (f x)))) (at x within S)" using has_derivative_compose[OF f has_derivative_inverse', OF x] . lemma has_derivative_divide[simp, derivative_intros]: fixes f :: "_ \ 'a::real_normed_div_algebra" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" assumes x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. - f x * (inverse (g x) * g' h * inverse (g x)) + f' h / g x)) (at x within S)" using has_derivative_mult[OF f has_derivative_inverse[OF x g]] by (simp add: field_simps) text \Conventional form requires mult-AC laws. Types real and complex only.\ lemma has_derivative_divide'[derivative_intros]: fixes f :: "_ \ 'a::real_normed_field" assumes f: "(f has_derivative f') (at x within S)" and g: "(g has_derivative g') (at x within S)" and x: "g x \ 0" shows "((\x. f x / g x) has_derivative (\h. (f' h * g x - f x * g' h) / (g x * g x))) (at x within S)" proof - have "f' h / g x - f x * (inverse (g x) * g' h * inverse (g x)) = (f' h * g x - f x * g' h) / (g x * g x)" for h by (simp add: field_simps x) then show ?thesis using has_derivative_divide [OF f g] x by simp qed subsection \Uniqueness\ text \ This can not generally shown for \<^const>\has_derivative\, as we need to approach the point from all directions. There is a proof in \Analysis\ for \euclidean_space\. \ lemma has_derivative_at2: "(f has_derivative f') (at x) \ bounded_linear f' \ ((\y. (1 / (norm(y - x))) *\<^sub>R (f y - (f x + f' (y - x)))) \ 0) (at x)" using has_derivative_within [of f f' x UNIV] by simp lemma has_derivative_zero_unique: assumes "((\x. 0) has_derivative F) (at x)" shows "F = (\h. 0)" proof - interpret F: bounded_linear F using assms by (rule has_derivative_bounded_linear) let ?r = "\h. norm (F h) / norm h" have *: "?r \0\ 0" using assms unfolding has_derivative_at by simp show "F = (\h. 0)" proof show "F h = 0" for h proof (rule ccontr) assume **: "\ ?thesis" then have h: "h \ 0" by (auto simp add: F.zero) with ** have "0 < ?r h" by simp from LIM_D [OF * this] obtain S where S: "0 < S" and r: "\x. x \ 0 \ norm x < S \ ?r x < ?r h" by auto from dense [OF S] obtain t where t: "0 < t \ t < S" .. let ?x = "scaleR (t / norm h) h" have "?x \ 0" and "norm ?x < S" using t h by simp_all then have "?r ?x < ?r h" by (rule r) then show False using t h by (simp add: F.scaleR) qed qed qed lemma has_derivative_unique: assumes "(f has_derivative F) (at x)" and "(f has_derivative F') (at x)" shows "F = F'" proof - have "((\x. 0) has_derivative (\h. F h - F' h)) (at x)" using has_derivative_diff [OF assms] by simp then have "(\h. F h - F' h) = (\h. 0)" by (rule has_derivative_zero_unique) then show "F = F'" unfolding fun_eq_iff right_minus_eq . qed subsection \Differentiability predicate\ definition differentiable :: "('a::real_normed_vector \ 'b::real_normed_vector) \ 'a filter \ bool" (infix "differentiable" 50) where "f differentiable F \ (\D. (f has_derivative D) F)" lemma differentiable_subset: "f differentiable (at x within s) \ t \ s \ f differentiable (at x within t)" unfolding differentiable_def by (blast intro: has_derivative_subset) lemmas differentiable_within_subset = differentiable_subset lemma differentiable_ident [simp, derivative_intros]: "(\x. x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_ident) lemma differentiable_const [simp, derivative_intros]: "(\z. a) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_const) lemma differentiable_in_compose: "f differentiable (at (g x) within (g`s)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_in_compose) lemma differentiable_compose: "f differentiable (at (g x)) \ g differentiable (at x within s) \ (\x. f (g x)) differentiable (at x within s)" by (blast intro: differentiable_in_compose differentiable_subset) lemma differentiable_add [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x + g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_add) lemma differentiable_sum[simp, derivative_intros]: assumes "finite s" "\a\s. (f a) differentiable net" shows "(\x. sum (\a. f a x) s) differentiable net" proof - from bchoice[OF assms(2)[unfolded differentiable_def]] show ?thesis by (auto intro!: has_derivative_sum simp: differentiable_def) qed lemma differentiable_minus [simp, derivative_intros]: "f differentiable F \ (\x. - f x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_minus) lemma differentiable_diff [simp, derivative_intros]: "f differentiable F \ g differentiable F \ (\x. f x - g x) differentiable F" unfolding differentiable_def by (blast intro: has_derivative_diff) lemma differentiable_mult [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_algebra" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x * g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_mult) lemma differentiable_inverse [simp, derivative_intros]: fixes f :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ f x \ 0 \ (\x. inverse (f x)) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_inverse) lemma differentiable_divide [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ g differentiable (at x within s) \ g x \ 0 \ (\x. f x / g x) differentiable (at x within s)" unfolding divide_inverse by simp lemma differentiable_power [simp, derivative_intros]: fixes f g :: "'a::real_normed_vector \ 'b::real_normed_field" shows "f differentiable (at x within s) \ (\x. f x ^ n) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_power) lemma differentiable_scaleR [simp, derivative_intros]: "f differentiable (at x within s) \ g differentiable (at x within s) \ (\x. f x *\<^sub>R g x) differentiable (at x within s)" unfolding differentiable_def by (blast intro: has_derivative_scaleR) lemma has_derivative_imp_has_field_derivative: "(f has_derivative D) F \ (\x. x * D' = D x) \ (f has_field_derivative D') F" unfolding has_field_derivative_def by (rule has_derivative_eq_rhs[of f D]) (simp_all add: fun_eq_iff mult.commute) lemma has_field_derivative_imp_has_derivative: "(f has_field_derivative D) F \ (f has_derivative (*) D) F" by (simp add: has_field_derivative_def) lemma DERIV_subset: "(f has_field_derivative f') (at x within s) \ t \ s \ (f has_field_derivative f') (at x within t)" by (simp add: has_field_derivative_def has_derivative_within_subset) lemma has_field_derivative_at_within: "(f has_field_derivative f') (at x) \ (f has_field_derivative f') (at x within s)" using DERIV_subset by blast abbreviation (input) DERIV :: "('a::real_normed_field \ 'a) \ 'a \ 'a \ bool" ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where "DERIV f x :> D \ (f has_field_derivative D) (at x)" abbreviation has_real_derivative :: "(real \ real) \ real \ real filter \ bool" (infix "(has'_real'_derivative)" 50) where "(f has_real_derivative D) F \ (f has_field_derivative D) F" lemma real_differentiable_def: "f differentiable at x within s \ (\D. (f has_real_derivative D) (at x within s))" proof safe assume "f differentiable at x within s" then obtain f' where *: "(f has_derivative f') (at x within s)" unfolding differentiable_def by auto then obtain c where "f' = ((*) c)" by (metis real_bounded_linear has_derivative_bounded_linear mult.commute fun_eq_iff) with * show "\D. (f has_real_derivative D) (at x within s)" unfolding has_field_derivative_def by auto qed (auto simp: differentiable_def has_field_derivative_def) lemma real_differentiableE [elim?]: assumes f: "f differentiable (at x within s)" obtains df where "(f has_real_derivative df) (at x within s)" using assms by (auto simp: real_differentiable_def) lemma has_field_derivative_iff: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" proof - have "((\y. norm (f y - f x - D * (y - x)) / norm (y - x)) \ 0) (at x within S) = ((\y. (f y - f x) / (y - x) - D) \ 0) (at x within S)" apply (subst tendsto_norm_zero_iff[symmetric], rule filterlim_cong) apply (simp_all add: eventually_at_filter field_simps nonzero_norm_divide) done then show ?thesis by (simp add: has_field_derivative_def has_derivative_iff_norm bounded_linear_mult_right LIM_zero_iff) qed lemma DERIV_def: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" unfolding field_has_derivative_at has_field_derivative_def has_field_derivative_iff .. lemma mult_commute_abs: "(\x. x * c) = (*) c" for c :: "'a::ab_semigroup_mult" by (simp add: fun_eq_iff mult.commute) lemma DERIV_compose_FDERIV: fixes f::"real\real" assumes "DERIV f (g x) :> f'" assumes "(g has_derivative g') (at x within s)" shows "((\x. f (g x)) has_derivative (\x. g' x * f')) (at x within s)" using assms has_derivative_compose[of g g' x s f "(*) f'"] by (auto simp: has_field_derivative_def ac_simps) subsection \Vector derivative\ lemma has_field_derivative_iff_has_vector_derivative: "(f has_field_derivative y) F \ (f has_vector_derivative y) F" unfolding has_vector_derivative_def has_field_derivative_def real_scaleR_def mult_commute_abs .. lemma has_field_derivative_subset: "(f has_field_derivative y) (at x within s) \ t \ s \ (f has_field_derivative y) (at x within t)" unfolding has_field_derivative_def by (rule has_derivative_subset) lemma has_vector_derivative_const[simp, derivative_intros]: "((\x. c) has_vector_derivative 0) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_id[simp, derivative_intros]: "((\x. x) has_vector_derivative 1) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_minus[derivative_intros]: "(f has_vector_derivative f') net \ ((\x. - f x) has_vector_derivative (- f')) net" by (auto simp: has_vector_derivative_def) lemma has_vector_derivative_add[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x + g x) has_vector_derivative (f' + g')) net" by (auto simp: has_vector_derivative_def scaleR_right_distrib) lemma has_vector_derivative_sum[derivative_intros]: "(\i. i \ I \ (f i has_vector_derivative f' i) net) \ ((\x. \i\I. f i x) has_vector_derivative (\i\I. f' i)) net" by (auto simp: has_vector_derivative_def fun_eq_iff scaleR_sum_right intro!: derivative_eq_intros) lemma has_vector_derivative_diff[derivative_intros]: "(f has_vector_derivative f') net \ (g has_vector_derivative g') net \ ((\x. f x - g x) has_vector_derivative (f' - g')) net" by (auto simp: has_vector_derivative_def scaleR_diff_right) lemma has_vector_derivative_add_const: "((\t. g t + z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" apply (intro iffI) apply (force dest: has_vector_derivative_diff [where g = "\t. z", OF _ has_vector_derivative_const]) apply (force dest: has_vector_derivative_add [OF _ has_vector_derivative_const]) done lemma has_vector_derivative_diff_const: "((\t. g t - z) has_vector_derivative f') net = ((\t. g t) has_vector_derivative f') net" using has_vector_derivative_add_const [where z = "-z"] by simp lemma (in bounded_linear) has_vector_derivative: assumes "(g has_vector_derivative g') F" shows "((\x. f (g x)) has_vector_derivative f g') F" using has_derivative[OF assms[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR) lemma (in bounded_bilinear) has_vector_derivative: assumes "(f has_vector_derivative f') (at x within s)" and "(g has_vector_derivative g') (at x within s)" shows "((\x. f x ** g x) has_vector_derivative (f x ** g' + f' ** g x)) (at x within s)" using FDERIV[OF assms(1-2)[unfolded has_vector_derivative_def]] by (simp add: has_vector_derivative_def scaleR_right scaleR_left scaleR_right_distrib) lemma has_vector_derivative_scaleR[derivative_intros]: "(f has_field_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x *\<^sub>R g x) has_vector_derivative (f x *\<^sub>R g' + f' *\<^sub>R g x)) (at x within s)" unfolding has_field_derivative_iff_has_vector_derivative by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_scaleR]) lemma has_vector_derivative_mult[derivative_intros]: "(f has_vector_derivative f') (at x within s) \ (g has_vector_derivative g') (at x within s) \ ((\x. f x * g x) has_vector_derivative (f x * g' + f' * g x)) (at x within s)" for f g :: "real \ 'a::real_normed_algebra" by (rule bounded_bilinear.has_vector_derivative[OF bounded_bilinear_mult]) lemma has_vector_derivative_of_real[derivative_intros]: "(f has_field_derivative D) F \ ((\x. of_real (f x)) has_vector_derivative (of_real D)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_of_real]) (simp add: has_field_derivative_iff_has_vector_derivative) lemma has_vector_derivative_real_field: "(f has_field_derivative f') (at (of_real a)) \ ((\x. f (of_real x)) has_vector_derivative f') (at a within s)" using has_derivative_compose[of of_real of_real a _ f "(*) f'"] by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) lemma has_vector_derivative_continuous: "(f has_vector_derivative D) (at x within s) \ continuous (at x within s) f" by (auto intro: has_derivative_continuous simp: has_vector_derivative_def) lemma continuous_on_vector_derivative: "(\x. x \ S \ (f has_vector_derivative f' x) (at x within S)) \ continuous_on S f" by (auto simp: continuous_on_eq_continuous_within intro!: has_vector_derivative_continuous) lemma has_vector_derivative_mult_right[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. a * f x) has_vector_derivative (a * x)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_right]) lemma has_vector_derivative_mult_left[derivative_intros]: fixes a :: "'a::real_normed_algebra" shows "(f has_vector_derivative x) F \ ((\x. f x * a) has_vector_derivative (x * a)) F" by (rule bounded_linear.has_vector_derivative[OF bounded_linear_mult_left]) subsection \Derivatives\ lemma DERIV_D: "DERIV f x :> D \ (\h. (f (x + h) - f x) / h) \0\ D" by (simp add: DERIV_def) lemma has_field_derivativeD: "(f has_field_derivative D) (at x within S) \ ((\y. (f y - f x) / (y - x)) \ D) (at x within S)" by (simp add: has_field_derivative_iff) lemma DERIV_const [simp, derivative_intros]: "((\x. k) has_field_derivative 0) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_const]) auto lemma DERIV_ident [simp, derivative_intros]: "((\x. x) has_field_derivative 1) F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_ident]) auto lemma field_differentiable_add[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z + g z) has_field_derivative f' + g') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_add]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_add: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x + g x) has_field_derivative D + E) (at x within s)" by (rule field_differentiable_add) lemma field_differentiable_minus[derivative_intros]: "(f has_field_derivative f') F \ ((\z. - (f z)) has_field_derivative -f') F" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_minus]) (auto simp: has_field_derivative_def field_simps mult_commute_abs) corollary DERIV_minus: "(f has_field_derivative D) (at x within s) \ ((\x. - f x) has_field_derivative -D) (at x within s)" by (rule field_differentiable_minus) lemma field_differentiable_diff[derivative_intros]: "(f has_field_derivative f') F \ (g has_field_derivative g') F \ ((\z. f z - g z) has_field_derivative f' - g') F" by (simp only: diff_conv_add_uminus field_differentiable_add field_differentiable_minus) corollary DERIV_diff: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x - g x) has_field_derivative D - E) (at x within s)" by (rule field_differentiable_diff) lemma DERIV_continuous: "(f has_field_derivative D) (at x within s) \ continuous (at x within s) f" by (drule has_derivative_continuous[OF has_field_derivative_imp_has_derivative]) simp corollary DERIV_isCont: "DERIV f x :> D \ isCont f x" by (rule DERIV_continuous) lemma DERIV_atLeastAtMost_imp_continuous_on: assumes "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y" shows "continuous_on {a..b} f" by (meson DERIV_isCont assms atLeastAtMost_iff continuous_at_imp_continuous_at_within continuous_on_eq_continuous_within) lemma DERIV_continuous_on: "(\x. x \ s \ (f has_field_derivative (D x)) (at x within s)) \ continuous_on s f" unfolding continuous_on_eq_continuous_within by (intro continuous_at_imp_continuous_on ballI DERIV_continuous) lemma DERIV_mult': "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ ((\x. f x * g x) has_field_derivative f x * E + D * g x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_mult[derivative_intros]: "(f has_field_derivative Da) (at x within s) \ (g has_field_derivative Db) (at x within s) \ ((\x. f x * g x) has_field_derivative Da * g x + Db * f x) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_mult]) (auto simp: field_simps dest: has_field_derivative_imp_has_derivative) text \Derivative of linear multiplication\ lemma DERIV_cmult: "(f has_field_derivative D) (at x within s) \ ((\x. c * f x) has_field_derivative c * D) (at x within s)" by (drule DERIV_mult' [OF DERIV_const]) simp lemma DERIV_cmult_right: "(f has_field_derivative D) (at x within s) \ ((\x. f x * c) has_field_derivative D * c) (at x within s)" using DERIV_cmult by (auto simp add: ac_simps) lemma DERIV_cmult_Id [simp]: "((*) c has_field_derivative c) (at x within s)" using DERIV_ident [THEN DERIV_cmult, where c = c and x = x] by simp lemma DERIV_cdivide: "(f has_field_derivative D) (at x within s) \ ((\x. f x / c) has_field_derivative D / c) (at x within s)" using DERIV_cmult_right[of f D x s "1 / c"] by simp lemma DERIV_unique: "DERIV f x :> D \ DERIV f x :> E \ D = E" unfolding DERIV_def by (rule LIM_unique) lemma DERIV_sum[derivative_intros]: "(\ n. n \ S \ ((\x. f x n) has_field_derivative (f' x n)) F) \ ((\x. sum (f x) S) has_field_derivative sum (f' x) S) F" by (rule has_derivative_imp_has_field_derivative [OF has_derivative_sum]) (auto simp: sum_distrib_left mult_commute_abs dest: has_field_derivative_imp_has_derivative) lemma DERIV_inverse'[derivative_intros]: assumes "(f has_field_derivative D) (at x within s)" and "f x \ 0" shows "((\x. inverse (f x)) has_field_derivative - (inverse (f x) * D * inverse (f x))) (at x within s)" proof - have "(f has_derivative (\x. x * D)) = (f has_derivative (*) D)" by (rule arg_cong [of "\x. x * D"]) (simp add: fun_eq_iff) with assms have "(f has_derivative (\x. x * D)) (at x within s)" by (auto dest!: has_field_derivative_imp_has_derivative) then show ?thesis using \f x \ 0\ by (auto intro: has_derivative_imp_has_field_derivative has_derivative_inverse) qed text \Power of \-1\\ lemma DERIV_inverse: "x \ 0 \ ((\x. inverse(x)) has_field_derivative - (inverse x ^ Suc (Suc 0))) (at x within s)" by (drule DERIV_inverse' [OF DERIV_ident]) simp text \Derivative of inverse\ lemma DERIV_inverse_fun: "(f has_field_derivative d) (at x within s) \ f x \ 0 \ ((\x. inverse (f x)) has_field_derivative (- (d * inverse(f x ^ Suc (Suc 0))))) (at x within s)" by (drule (1) DERIV_inverse') (simp add: ac_simps nonzero_inverse_mult_distrib) text \Derivative of quotient\ lemma DERIV_divide[derivative_intros]: "(f has_field_derivative D) (at x within s) \ (g has_field_derivative E) (at x within s) \ g x \ 0 \ ((\x. f x / g x) has_field_derivative (D * g x - f x * E) / (g x * g x)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_divide]) (auto dest: has_field_derivative_imp_has_derivative simp: field_simps) lemma DERIV_quotient: "(f has_field_derivative d) (at x within s) \ (g has_field_derivative e) (at x within s)\ g x \ 0 \ ((\y. f y / g y) has_field_derivative (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))) (at x within s)" by (drule (2) DERIV_divide) (simp add: mult.commute) lemma DERIV_power_Suc: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ Suc n) has_field_derivative (1 + of_nat n) * (D * f x ^ n)) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_power[derivative_intros]: "(f has_field_derivative D) (at x within s) \ ((\x. f x ^ n) has_field_derivative of_nat n * (D * f x ^ (n - Suc 0))) (at x within s)" by (rule has_derivative_imp_has_field_derivative[OF has_derivative_power]) (auto simp: has_field_derivative_def) lemma DERIV_pow: "((\x. x ^ n) has_field_derivative real n * (x ^ (n - Suc 0))) (at x within s)" using DERIV_power [OF DERIV_ident] by simp lemma DERIV_chain': "(f has_field_derivative D) (at x within s) \ DERIV g (f x) :> E \ ((\x. g (f x)) has_field_derivative E * D) (at x within s)" using has_derivative_compose[of f "(*) D" x s g "(*) E"] by (simp only: has_field_derivative_def mult_commute_abs ac_simps) corollary DERIV_chain2: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ ((\x. f (g x)) has_field_derivative Da * Db) (at x within s)" by (rule DERIV_chain') text \Standard version\ lemma DERIV_chain: "DERIV f (g x) :> Da \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" by (drule (1) DERIV_chain', simp add: o_def mult.commute) lemma DERIV_image_chain: "(f has_field_derivative Da) (at (g x) within (g ` s)) \ (g has_field_derivative Db) (at x within s) \ (f \ g has_field_derivative Da * Db) (at x within s)" using has_derivative_in_compose [of g "(*) Db" x s f "(*) Da "] by (simp add: has_field_derivative_def o_def mult_commute_abs ac_simps) (*These two are from HOL Light: HAS_COMPLEX_DERIVATIVE_CHAIN*) lemma DERIV_chain_s: assumes "(\x. x \ s \ DERIV g x :> g'(x))" and "DERIV f x :> f'" and "f x \ s" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis (full_types) DERIV_chain' mult.commute assms) lemma DERIV_chain3: (*HAS_COMPLEX_DERIVATIVE_CHAIN_UNIV*) assumes "(\x. DERIV g x :> g'(x))" and "DERIV f x :> f'" shows "DERIV (\x. g(f x)) x :> f' * g'(f x)" by (metis UNIV_I DERIV_chain_s [of UNIV] assms) text \Alternative definition for differentiability\ lemma DERIV_LIM_iff: fixes f :: "'a::{real_normed_vector,inverse} \ 'a" shows "((\h. (f (a + h) - f a) / h) \0\ D) = ((\x. (f x - f a) / (x - a)) \a\ D)" (is "?lhs = ?rhs") proof assume ?lhs then have "(\x. (f (a + (x + - a)) - f a) / (x + - a)) \0 - - a\ D" by (rule LIM_offset) then show ?rhs by simp next assume ?rhs then have "(\x. (f (x+a) - f a) / ((x+a) - a)) \a-a\ D" by (rule LIM_offset) then show ?lhs by (simp add: add.commute) qed lemma has_field_derivative_cong_ev: assumes "x = y" and *: "eventually (\x. x \ S \ f x = g x) (nhds x)" and "u = v" "S = t" "x \ S" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative v) (at y within t)" unfolding has_field_derivative_iff proof (rule filterlim_cong) from assms have "f y = g y" by (auto simp: eventually_nhds) with * show "\\<^sub>F z in at x within S. (f z - f x) / (z - x) = (g z - g y) / (z - y)" unfolding eventually_at_filter by eventually_elim (auto simp: assms \f y = g y\) qed (simp_all add: assms) lemma has_field_derivative_cong_eventually: assumes "eventually (\x. f x = g x) (at x within S)" "f x = g x" shows "(f has_field_derivative u) (at x within S) = (g has_field_derivative u) (at x within S)" unfolding has_field_derivative_iff proof (rule tendsto_cong) show "\\<^sub>F y in at x within S. (f y - f x) / (y - x) = (g y - g x) / (y - x)" using assms by (auto elim: eventually_mono) qed lemma DERIV_cong_ev: "x = y \ eventually (\x. f x = g x) (nhds x) \ u = v \ DERIV f x :> u \ DERIV g y :> v" by (rule has_field_derivative_cong_ev) simp_all lemma DERIV_shift: "(f has_field_derivative y) (at (x + z)) = ((\x. f (x + z)) has_field_derivative y) (at x)" by (simp add: DERIV_def field_simps) lemma DERIV_mirror: "(DERIV f (- x) :> y) \ (DERIV (\x. f (- x)) x :> - y)" for f :: "real \ real" and x y :: real by (simp add: DERIV_def filterlim_at_split filterlim_at_left_to_right tendsto_minus_cancel_left field_simps conj_commute) lemma floor_has_real_derivative: fixes f :: "real \ 'a::{floor_ceiling,order_topology}" assumes "isCont f x" and "f x \ \" shows "((\x. floor (f x)) has_real_derivative 0) (at x)" proof (subst DERIV_cong_ev[OF refl _ refl]) show "((\_. floor (f x)) has_real_derivative 0) (at x)" by simp have "\\<^sub>F y in at x. \f y\ = \f x\" by (rule eventually_floor_eq[OF assms[unfolded continuous_at]]) then show "\\<^sub>F y in nhds x. real_of_int \f y\ = real_of_int \f x\" unfolding eventually_at_filter by eventually_elim auto qed lemmas has_derivative_floor[derivative_intros] = floor_has_real_derivative[THEN DERIV_compose_FDERIV] lemma continuous_floor: fixes x::real shows "x \ \ \ continuous (at x) (real_of_int \ floor)" using floor_has_real_derivative [where f=id] by (auto simp: o_def has_field_derivative_def intro: has_derivative_continuous) lemma continuous_frac: fixes x::real assumes "x \ \" shows "continuous (at x) frac" proof - have "isCont (\x. real_of_int \x\) x" using continuous_floor [OF assms] by (simp add: o_def) then have *: "continuous (at x) (\x. x - real_of_int \x\)" by (intro continuous_intros) moreover have "\\<^sub>F x in nhds x. frac x = x - real_of_int \x\" by (simp add: frac_def) ultimately show ?thesis by (simp add: LIM_imp_LIM frac_def isCont_def) qed text \Caratheodory formulation of derivative at a point\ lemma CARAT_DERIV: "(DERIV f x :> l) \ (\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l)" (is "?lhs = ?rhs") proof assume ?lhs show "\g. (\z. f z - f x = g z * (z - x)) \ isCont g x \ g x = l" 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 "isCont ?g x" using \?lhs\ by (simp add: isCont_iff DERIV_def cong: LIM_equal [rule_format]) show "?g x = l" by simp qed next assume ?rhs then show ?lhs by (auto simp add: isCont_iff DERIV_def cong: LIM_cong) qed subsection \Local extrema\ text \If \<^term>\0 < f' x\ then \<^term>\x\ is Locally Strictly Increasing At The Right.\ lemma has_real_derivative_pos_inc_right: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and l: "0 < l" shows "\d > 0. \h > 0. x + h \ S \ h < d \ f x < f (x + h)" using assms proof - from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x + h \ S" with all [of "x + h"] show "f x < f (x+h)" proof (simp add: abs_if dist_real_def pos_less_divide_eq split: if_split_asm) assume "\ (f (x + h) - f x) / h < l" and h: "0 < h" with l have "0 < (f (x + h) - f x) / h" by arith then show "f x < f (x + h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_pos_inc_right: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "0 < l" shows "\d > 0. \h > 0. h < d \ f x < f (x + h)" using has_real_derivative_pos_inc_right[OF assms] by auto lemma has_real_derivative_neg_dec_left: fixes f :: "real \ real" assumes der: "(f has_real_derivative l) (at x within S)" and "l < 0" shows "\d > 0. \h > 0. x - h \ S \ h < d \ f x < f (x - h)" proof - from \l < 0\ have l: "- l > 0" by simp from der [THEN has_field_derivativeD, THEN tendstoD, OF l, unfolded eventually_at] obtain s where s: "0 < s" and all: "\xa. xa\S \ xa \ x \ dist xa x < s \ \(f xa - f x) / (xa - x) - l\ < - l" by (auto simp: dist_real_def) then show ?thesis proof (intro exI conjI strip) show "0 < s" by (rule s) next fix h :: real assume "0 < h" "h < s" "x - h \ S" with all [of "x - h"] show "f x < f (x-h)" proof (simp add: abs_if pos_less_divide_eq dist_real_def split: if_split_asm) assume "- ((f (x-h) - f x) / h) < l" and h: "0 < h" with l have "0 < (f (x-h) - f x) / h" by arith then show "f x < f (x - h)" by (simp add: pos_less_divide_eq h) qed qed qed lemma DERIV_neg_dec_left: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and l: "l < 0" shows "\d > 0. \h > 0. h < d \ f x < f (x - h)" using has_real_derivative_neg_dec_left[OF assms] by auto lemma has_real_derivative_pos_inc_left: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ 0 < l \ \d>0. \h>0. x - h \ S \ h < d \ f (x - h) < f x" by (rule has_real_derivative_neg_dec_left [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_pos_inc_left: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < l \ \d > 0. \h > 0. h < d \ f (x - h) < f x" using has_real_derivative_pos_inc_left by blast lemma has_real_derivative_neg_dec_right: fixes f :: "real \ real" shows "(f has_real_derivative l) (at x within S) \ l < 0 \ \d > 0. \h > 0. x + h \ S \ h < d \ f x > f (x + h)" by (rule has_real_derivative_pos_inc_right [of "\x. - f x" "-l" x S, simplified]) (auto simp add: DERIV_minus) lemma DERIV_neg_dec_right: fixes f :: "real \ real" shows "DERIV f x :> l \ l < 0 \ \d > 0. \h > 0. h < d \ f x > f (x + h)" using has_real_derivative_neg_dec_right by blast lemma DERIV_local_max: fixes f :: "real \ real" assumes der: "DERIV f x :> l" and d: "0 < d" and le: "\y. \x - y\ < d \ f y \ f x" shows "l = 0" proof (cases rule: linorder_cases [of l 0]) case equal then show ?thesis . next case less from DERIV_neg_dec_left [OF der less] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x - h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x - e"]] show ?thesis by (auto simp add: abs_if) next case greater from DERIV_pos_inc_right [OF der greater] obtain d' where d': "0 < d'" and lt: "\h > 0. h < d' \ f x < f (x + h)" by blast obtain e where "0 < e \ e < d \ e < d'" using field_lbound_gt_zero [OF d d'] .. with lt le [THEN spec [where x="x + e"]] show ?thesis by (auto simp add: abs_if) qed text \Similar theorem for a local minimum\ lemma DERIV_local_min: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x \ f y \ l = 0" by (drule DERIV_minus [THEN DERIV_local_max]) auto text\In particular, if a function is locally flat\ lemma DERIV_local_const: fixes f :: "real \ real" shows "DERIV f x :> l \ 0 < d \ \y. \x - y\ < d \ f x = f y \ l = 0" by (auto dest!: DERIV_local_max) subsection \Rolle's Theorem\ text \Lemma about introducing open ball in open interval\ lemma lemma_interval_lt: fixes a b x :: real assumes "a < x" "x < b" shows "\d. 0 < d \ (\y. \x - y\ < d \ a < y \ y < b)" using linorder_linear [of "x - a" "b - x"] proof assume "x - a \ b - x" with assms show ?thesis by (rule_tac x = "x - a" in exI) auto next assume "b - x \ x - a" with assms show ?thesis by (rule_tac x = "b - x" in exI) auto qed lemma lemma_interval: "a < x \ x < b \ \d. 0 < d \ (\y. \x - y\ < d \ a \ y \ y \ b)" for a b x :: real by (force dest: lemma_interval_lt) text \Rolle's Theorem. If \<^term>\f\ is defined and continuous on the closed interval \[a,b]\ and differentiable on the open interval \(a,b)\, and \<^term>\f a = f b\, then there exists \x0 \ (a,b)\ such that \<^term>\f' x0 = 0\\ theorem Rolle_deriv: fixes f :: "real \ real" assumes "a < b" and fab: "f a = f b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" shows "\z. a < z \ z < b \ f' z = (\v. 0)" proof - have le: "a \ b" using \a < b\ by simp have "(a + b) / 2 \ {a..b}" using assms(1) by auto then have *: "{a..b} \ {}" by auto obtain x where x_max: "\z. a \ z \ z \ b \ f z \ f x" and "a \ x" "x \ b" using continuous_attains_sup[OF compact_Icc * contf] by (meson atLeastAtMost_iff) obtain x' where x'_min: "\z. a \ z \ z \ b \ f x' \ f z" and "a \ x'" "x' \ b" using continuous_attains_inf[OF compact_Icc * contf] by (meson atLeastAtMost_iff) consider "a < x" "x < b" | "x = a \ x = b" using \a \ x\ \x \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its maximum within the interval\ then obtain l where der: "DERIV f x :> l" using derf differentiable_def real_differentiable_def by blast obtain d where d: "0 < d" and bound: "\y. \x - y\ < d \ a \ y \ y \ b" using lemma_interval [OF 1] by blast then have bound': "\y. \x - y\ < d \ f y \ f x" using x_max by blast \ \the derivative at a local maximum is zero\ have "l = 0" by (rule DERIV_local_max [OF der d bound']) with 1 der derf [of x] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 then have fx: "f b = f x" by (auto simp add: fab) consider "a < x'" "x' < b" | "x' = a \ x' = b" using \a \ x'\ \x' \ b\ by arith then show ?thesis proof cases case 1 \ \\<^term>\f\ attains its minimum within the interval\ then obtain l where der: "DERIV f x' :> l" using derf differentiable_def real_differentiable_def by blast from lemma_interval [OF 1] obtain d where d: "0y. \x'-y\ < d \ a \ y \ y \ b" by blast then have bound': "\y. \x' - y\ < d \ f x' \ f y" using x'_min by blast have "l = 0" by (rule DERIV_local_min [OF der d bound']) \ \the derivative at a local minimum is zero\ then show ?thesis using 1 der derf [of x'] by (metis has_derivative_unique has_field_derivative_def mult_zero_left) next case 2 \ \\<^term>\f\ is constant throughout the interval\ then have fx': "f b = f x'" by (auto simp: fab) from dense [OF \a < b\] obtain r where r: "a < r" "r < b" by blast obtain d where d: "0 < d" and bound: "\y. \r - y\ < d \ a \ y \ y \ b" using lemma_interval [OF r] by blast have eq_fb: "f z = f b" if "a \ z" and "z \ b" for z proof (rule order_antisym) show "f z \ f b" by (simp add: fx x_max that) show "f b \ f z" by (simp add: fx' x'_min that) qed have bound': "\y. \r - y\ < d \ f r = f y" proof (intro strip) fix y :: real assume lt: "\r - y\ < d" then have "f y = f b" by (simp add: eq_fb bound) then show "f r = f y" by (simp add: eq_fb r order_less_imp_le) qed obtain l where der: "DERIV f r :> l" using derf differentiable_def r(1) r(2) real_differentiable_def by blast have "l = 0" by (rule DERIV_local_const [OF der d bound']) \ \the derivative of a constant function is zero\ with r der derf [of r] show ?thesis by (metis has_derivative_unique has_field_derivative_def mult_zero_left) qed qed qed corollary Rolle: fixes a b :: real assumes ab: "a < b" "f a = f b" "continuous_on {a..b} f" and dif [rule_format]: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\z. a < z \ z < b \ DERIV f z :> 0" proof - obtain f' where f': "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then have "\z. a < z \ z < b \ f' z = (\v. 0)" by (metis Rolle_deriv [OF ab]) then show ?thesis using f' has_derivative_imp_has_field_derivative by fastforce qed subsection \Mean Value Theorem\ theorem mvt: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ (f has_derivative f' x) (at x)" obtains \ where "a < \" "\ < b" "f b - f a = (f' \) (b - a)" proof - have "\x. a < x \ x < b \ (\y. f' x y - (f b - f a) / (b - a) * y) = (\v. 0)" proof (intro Rolle_deriv[OF \a < b\]) fix x assume x: "a < x" "x < b" show "((\x. f x - (f b - f a) / (b - a) * x) has_derivative (\y. f' x y - (f b - f a) / (b - a) * y)) (at x)" by (intro derivative_intros derf[OF x]) qed (use assms in \auto intro!: continuous_intros simp: field_simps\) then obtain \ where "a < \" "\ < b" "(\y. f' \ y - (f b - f a) / (b - a) * y) = (\v. 0)" by metis then show ?thesis by (metis (no_types, hide_lams) that add.right_neutral add_diff_cancel_left' add_diff_eq \a < b\ less_irrefl nonzero_eq_divide_eq) qed theorem MVT: fixes a b :: real assumes lt: "a < b" and contf: "continuous_on {a..b} f" and dif: "\x. \a < x; x < b\ \ f differentiable (at x)" shows "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" proof - obtain f' :: "real \ real \ real" where derf: "\x. a < x \ x < b \ (f has_derivative f' x) (at x)" using dif unfolding differentiable_def by metis then obtain z where "a < z" "z < b" "f b - f a = (f' z) (b - a)" using mvt [OF lt contf] by blast then show ?thesis by (simp add: ac_simps) (metis derf dif has_derivative_unique has_field_derivative_imp_has_derivative real_differentiable_def) qed corollary MVT2: assumes "a < b" and der: "\x. \a \ x; x \ b\ \ DERIV f x :> f' x" shows "\z::real. a < z \ z < b \ (f b - f a = (b - a) * f' z)" proof - have "\l z. a < z \ z < b \ (f has_real_derivative l) (at z) \ f b - f a = (b - a) * l" proof (rule MVT [OF \a < b\]) show "continuous_on {a..b} f" by (meson DERIV_continuous atLeastAtMost_iff continuous_at_imp_continuous_on der) show "\x. \a < x; x < b\ \ f differentiable (at x)" using assms by (force dest: order_less_imp_le simp add: real_differentiable_def) qed with assms show ?thesis by (blast dest: DERIV_unique order_less_imp_le) qed lemma pos_deriv_imp_strict_mono: assumes "\x. (f has_real_derivative f' x) (at x)" assumes "\x. f' x > 0" shows "strict_mono f" proof (rule strict_monoI) fix x y :: real assume xy: "x < y" from assms and xy have "\z>x. z < y \ f y - f x = (y - x) * f' z" by (intro MVT2) (auto dest: connectedD_interval) then obtain z where z: "z > x" "z < y" "f y - f x = (y - x) * f' z" by blast note \f y - f x = (y - x) * f' z\ also have "(y - x) * f' z > 0" using xy assms by (intro mult_pos_pos) auto finally show "f x < f y" by simp qed proposition deriv_nonneg_imp_mono: assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" assumes ab: "a \ b" shows "g a \ g b" proof (cases "a < b") assume "a < b" from deriv have "\x. \x \ a; x \ b\ \ (g has_real_derivative g' x) (at x)" by simp with MVT2[OF \a < b\] and deriv obtain \ where \_ab: "\ > a" "\ < b" and g_ab: "g b - g a = (b - a) * g' \" by blast from \_ab ab nonneg have "(b - a) * g' \ \ 0" by simp with g_ab show ?thesis by simp qed (insert ab, simp) subsubsection \A function is constant if its derivative is 0 over an interval.\ lemma DERIV_isconst_end: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and 0: "\x. \a < x; x < b\ \ DERIV f x :> 0" shows "f b = f a" using MVT [OF \a < b\] "0" DERIV_unique contf real_differentiable_def by (fastforce simp: algebra_simps) lemma DERIV_isconst2: fixes f :: "real \ real" assumes "a < b" and contf: "continuous_on {a..b} f" and derf: "\x. \a < x; x < b\ \ DERIV f x :> 0" and "a \ x" "x \ b" shows "f x = f a" proof (cases "a < x") case True have *: "continuous_on {a..x} f" using \x \ b\ contf continuous_on_subset by fastforce show ?thesis by (rule DERIV_isconst_end [OF True *]) (use \x \ b\ derf in auto) qed (use \a \ x\ in auto) lemma DERIV_isconst3: fixes a b x y :: real assumes "a < b" and "x \ {a <..< b}" and "y \ {a <..< b}" and derivable: "\x. x \ {a <..< b} \ DERIV f x :> 0" shows "f x = f y" proof (cases "x = y") case False let ?a = "min x y" let ?b = "max x y" have *: "DERIV f z :> 0" if "?a \ z" "z \ ?b" for z proof - have "a < z" and "z < b" using that \x \ {a <..< b}\ and \y \ {a <..< b}\ by auto then have "z \ {a<.. 0" by (rule derivable) qed have isCont: "continuous_on {?a..?b} f" by (meson * DERIV_continuous_on atLeastAtMost_iff has_field_derivative_at_within) have DERIV: "\z. \?a < z; z < ?b\ \ DERIV f z :> 0" using * by auto have "?a < ?b" using \x \ y\ by auto from DERIV_isconst2[OF this isCont DERIV, of x] and DERIV_isconst2[OF this isCont DERIV, of y] show ?thesis by auto qed auto lemma DERIV_isconst_all: fixes f :: "real \ real" shows "\x. DERIV f x :> 0 \ f x = f y" apply (rule linorder_cases [of x y]) apply (metis DERIV_continuous DERIV_isconst_end continuous_at_imp_continuous_on)+ done lemma DERIV_const_ratio_const: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "f b - f a = (b - a) * k" proof (cases a b rule: linorder_cases) case less show ?thesis using MVT [OF less] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) next case greater have "f a - f b = (a - b) * k" using MVT [OF greater] df by (metis DERIV_continuous DERIV_unique continuous_at_imp_continuous_on real_differentiable_def) then show ?thesis by (simp add: algebra_simps) qed auto lemma DERIV_const_ratio_const2: fixes f :: "real \ real" assumes "a \ b" and df: "\x. DERIV f x :> k" shows "(f b - f a) / (b - a) = k" using DERIV_const_ratio_const [OF assms] \a \ b\ by auto lemma real_average_minus_first [simp]: "(a + b) / 2 - a = (b - a) / 2" for a b :: real by simp lemma real_average_minus_second [simp]: "(b + a) / 2 - a = (b - a) / 2" for a b :: real by simp text \Gallileo's "trick": average velocity = av. of end velocities.\ lemma DERIV_const_average: fixes v :: "real \ real" and a b :: real assumes neq: "a \ b" and der: "\x. DERIV v x :> k" shows "v ((a + b) / 2) = (v a + v b) / 2" proof (cases rule: linorder_cases [of a b]) case equal with neq show ?thesis by simp next case less have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have "(v ((a + b) / 2) - v a) / ((a + b) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by force next case greater have "(v b - v a) / (b - a) = k" by (rule DERIV_const_ratio_const2 [OF neq der]) then have "(b - a) * ((v b - v a) / (b - a)) = (b - a) * k" by simp moreover have " (v ((b + a) / 2) - v a) / ((b + a) / 2 - a) = k" by (rule DERIV_const_ratio_const2 [OF _ der]) (simp add: neq) ultimately show ?thesis using neq by (force simp add: add.commute) qed subsubsection\A function with positive derivative is increasing\ text \A simple proof using the MVT, by Jeremy Avigad. And variants.\ lemma DERIV_pos_imp_increasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ (\y. DERIV f x :> y \ y > 0)" and con: "continuous_on {a..b} f" shows "f a < f b" proof (rule ccontr) assume f: "\ ?thesis" have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" by (rule MVT) (use assms real_differentiable_def in \force+\) then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" by auto with assms f have "\ l > 0" by (metis linorder_not_le mult_le_0_iff diff_le_0_iff_le) with assms z show False by (metis DERIV_unique) qed lemma DERIV_pos_imp_increasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y > 0" shows "f a < f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_pos_imp_increasing_open [OF \a < b\] der) lemma DERIV_nonneg_imp_nondecreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof (rule ccontr, cases "a = b") assume "\ ?thesis" and "a = b" then show False by auto next assume *: "\ ?thesis" assume "a \ b" with \a \ b\ have "a < b" by linarith moreover have "continuous_on {a..b} f" by (meson DERIV_isCont assms(2) atLeastAtMost_iff continuous_at_imp_continuous_on) ultimately have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" using assms MVT [OF \a < b\, of f] real_differentiable_def less_eq_real_def by blast then obtain l z where lz: "a < z" "z < b" "DERIV f z :> l" and **: "f b - f a = (b - a) * l" by auto with * have "a < b" "f b < f a" by auto with ** have "\ l \ 0" by (auto simp add: not_le algebra_simps) (metis * add_le_cancel_right assms(1) less_eq_real_def mult_right_mono add_left_mono linear order_refl) with assms lz show False by (metis DERIV_unique order_less_imp_le) qed lemma DERIV_neg_imp_decreasing_open: fixes a b :: real and f :: "real \ real" assumes "a < b" and "\x. a < x \ x < b \ \y. DERIV f x :> y \ y < 0" and con: "continuous_on {a..b} f" shows "f a > f b" proof - have "(\x. -f x) a < (\x. -f x) b" proof (rule DERIV_pos_imp_increasing_open [of a b]) show "\x. \a < x; x < b\ \ \y. ((\x. - f x) has_real_derivative y) (at x) \ 0 < y" using assms by simp (metis field_differentiable_minus neg_0_less_iff_less) show "continuous_on {a..b} (\x. - f x)" using con continuous_on_minus by blast qed (use assms in auto) then show ?thesis by simp qed lemma DERIV_neg_imp_decreasing: fixes a b :: real and f :: "real \ real" assumes "a < b" and der: "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y < 0" shows "f a > f b" by (metis less_le_not_le DERIV_atLeastAtMost_imp_continuous_on DERIV_neg_imp_decreasing_open [OF \a < b\] der) lemma DERIV_nonpos_imp_nonincreasing: fixes a b :: real and f :: "real \ real" assumes "a \ b" and "\x. \a \ x; x \ b\ \ \y. DERIV f x :> y \ y \ 0" shows "f a \ f b" proof - have "(\x. -f x) a \ (\x. -f x) b" using DERIV_nonneg_imp_nondecreasing [of a b "\x. -f x"] assms DERIV_minus by fastforce then show ?thesis by simp qed lemma DERIV_pos_imp_increasing_at_bot: fixes f :: "real \ real" assumes "\x. x \ b \ (\y. DERIV f x :> y \ y > 0)" and lim: "(f \ flim) at_bot" shows "flim < f b" proof - have "\N. \n\N. f n \ f (b - 1)" by (rule_tac x="b - 2" in exI) (force intro: order.strict_implies_order DERIV_pos_imp_increasing assms) then have "flim \ f (b - 1)" by (auto simp: eventually_at_bot_linorder tendsto_upperbound [OF lim]) also have "\ < f b" by (force intro: DERIV_pos_imp_increasing [where f=f] assms) finally show ?thesis . qed lemma DERIV_neg_imp_decreasing_at_top: fixes f :: "real \ real" assumes der: "\x. x \ b \ \y. DERIV f x :> y \ y < 0" and lim: "(f \ flim) at_top" shows "flim < f b" apply (rule DERIV_pos_imp_increasing_at_bot [where f = "\i. f (-i)" and b = "-b", simplified]) apply (metis DERIV_mirror der le_minus_iff neg_0_less_iff_less) apply (metis filterlim_at_top_mirror lim) done text \Derivative of inverse function\ lemma DERIV_inverse_function: fixes f g :: "real \ real" assumes der: "DERIV f (g x) :> D" and neq: "D \ 0" and x: "a < x" "x < b" and inj: "\y. \a < y; y < b\ \ f (g y) = y" and cont: "isCont g x" shows "DERIV g x :> inverse D" unfolding has_field_derivative_iff proof (rule LIM_equal2) show "0 < min (x - a) (b - x)" using x by arith next fix y assume "norm (y - x) < min (x - a) (b - x)" then have "a < y" and "y < b" by (simp_all add: abs_less_iff) then show "(g y - g x) / (y - x) = inverse ((f (g y) - x) / (g y - g x))" by (simp add: inj) next have "(\z. (f z - f (g x)) / (z - g x)) \g x\ D" by (rule der [unfolded has_field_derivative_iff]) then have 1: "(\z. (f z - x) / (z - g x)) \g x\ D" using inj x by simp have 2: "\d>0. \y. y \ x \ norm (y - x) < d \ g y \ g x" proof (rule exI, safe) show "0 < min (x - a) (b - x)" using x by simp next fix y assume "norm (y - x) < min (x - a) (b - x)" then have y: "a < y" "y < b" by (simp_all add: abs_less_iff) assume "g y = g x" then have "f (g y) = f (g x)" by simp then have "y = x" using inj y x by simp also assume "y \ x" finally show False by simp qed have "(\y. (f (g y) - x) / (g y - g x)) \x\ D" using cont 1 2 by (rule isCont_LIM_compose2) then show "(\y. inverse ((f (g y) - x) / (g y - g x))) \x\ inverse D" using neq by (rule tendsto_inverse) qed subsection \Generalized Mean Value Theorem\ theorem GMVT: fixes a b :: real assumes alb: "a < b" and fc: "\x. a \ x \ x \ b \ isCont f x" and fd: "\x. a < x \ x < b \ f differentiable (at x)" and gc: "\x. a \ x \ x \ b \ isCont g x" and gd: "\x. a < x \ x < b \ g differentiable (at x)" shows "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" proof - let ?h = "\x. (f b - f a) * g x - (g b - g a) * f x" have "\l z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" proof (rule MVT) from assms show "a < b" by simp show "continuous_on {a..b} ?h" by (simp add: continuous_at_imp_continuous_on fc gc) show "\x. \a < x; x < b\ \ ?h differentiable (at x)" using fd gd by simp qed then obtain l where l: "\z. a < z \ z < b \ DERIV ?h z :> l \ ?h b - ?h a = (b - a) * l" .. then obtain c where c: "a < c \ c < b \ DERIV ?h c :> l \ ?h b - ?h a = (b - a) * l" .. from c have cint: "a < c \ c < b" by auto then obtain g'c where g'c: "DERIV g c :> g'c" using gd real_differentiable_def by blast from c have "a < c \ c < b" by auto then obtain f'c where f'c: "DERIV f c :> f'c" using fd real_differentiable_def by blast from c have "DERIV ?h c :> l" by auto moreover have "DERIV ?h c :> g'c * (f b - f a) - f'c * (g b - g a)" using g'c f'c by (auto intro!: derivative_eq_intros) ultimately have leq: "l = g'c * (f b - f a) - f'c * (g b - g a)" by (rule DERIV_unique) have "?h b - ?h a = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" proof - from c have "?h b - ?h a = (b - a) * l" by auto also from leq have "\ = (b - a) * (g'c * (f b - f a) - f'c * (g b - g a))" by simp finally show ?thesis by simp qed moreover have "?h b - ?h a = 0" proof - have "?h b - ?h a = ((f b)*(g b) - (f a)*(g b) - (g b)*(f b) + (g a)*(f b)) - ((f b)*(g a) - (f a)*(g a) - (g b)*(f a) + (g a)*(f a))" by (simp add: algebra_simps) then show ?thesis by auto qed ultimately have "(b - a) * (g'c * (f b - f a) - f'c * (g b - g a)) = 0" by auto with alb have "g'c * (f b - f a) - f'c * (g b - g a) = 0" by simp then have "g'c * (f b - f a) = f'c * (g b - g a)" by simp then have "(f b - f a) * g'c = (g b - g a) * f'c" by (simp add: ac_simps) with g'c f'c cint show ?thesis by auto qed lemma GMVT': fixes f g :: "real \ real" assumes "a < b" and isCont_f: "\z. a \ z \ z \ b \ isCont f z" and isCont_g: "\z. a \ z \ z \ b \ isCont g z" and DERIV_g: "\z. a < z \ z < b \ DERIV g z :> (g' z)" and DERIV_f: "\z. a < z \ z < b \ DERIV f z :> (f' z)" shows "\c. a < c \ c < b \ (f b - f a) * g' c = (g b - g a) * f' c" proof - have "\g'c f'c c. DERIV g c :> g'c \ DERIV f c :> f'c \ a < c \ c < b \ (f b - f a) * g'c = (g b - g a) * f'c" using assms by (intro GMVT) (force simp: real_differentiable_def)+ then obtain c where "a < c" "c < b" "(f b - f a) * g' c = (g b - g a) * f' c" using DERIV_f DERIV_g by (force dest: DERIV_unique) then show ?thesis by auto qed subsection \L'Hopitals rule\ lemma isCont_If_ge: fixes a :: "'a :: linorder_topology" assumes "continuous (at_left a) g" and f: "(f \ g a) (at_right a)" shows "isCont (\x. if x \ a then g x else f x) a" (is "isCont ?gf a") proof - have g: "(g \ g a) (at_left a)" using assms continuous_within by blast show ?thesis unfolding isCont_def continuous_within proof (intro filterlim_split_at; simp) show "(?gf \ g a) (at_left a)" by (subst filterlim_cong[OF refl refl, where g=g]) (simp_all add: eventually_at_filter less_le g) show "(?gf \ g a) (at_right a)" by (subst filterlim_cong[OF refl refl, where g=f]) (simp_all add: eventually_at_filter less_le f) qed qed lemma lhopital_right_0: fixes f0 g0 :: "real \ real" assumes f_0: "(f0 \ 0) (at_right 0)" and g_0: "(g0 \ 0) (at_right 0)" and ev: "eventually (\x. g0 x \ 0) (at_right 0)" "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f0 x :> f' x) (at_right 0)" "eventually (\x. DERIV g0 x :> g' x) (at_right 0)" and lim: "filterlim (\ x. (f' x / g' x)) F (at_right 0)" shows "filterlim (\ x. f0 x / g0 x) F (at_right 0)" proof - define f where [abs_def]: "f x = (if x \ 0 then 0 else f0 x)" for x then have "f 0 = 0" by simp define g where [abs_def]: "g x = (if x \ 0 then 0 else g0 x)" for x then have "g 0 = 0" by simp have "eventually (\x. g0 x \ 0 \ g' x \ 0 \ DERIV f0 x :> (f' x) \ DERIV g0 x :> (g' x)) (at_right 0)" using ev by eventually_elim auto then obtain a where [arith]: "0 < a" and g0_neq_0: "\x. 0 < x \ x < a \ g0 x \ 0" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x < a \ DERIV f0 x :> (f' x)" and g0: "\x. 0 < x \ x < a \ DERIV g0 x :> (g' x)" unfolding eventually_at by (auto simp: dist_real_def) have g_neq_0: "\x. 0 < x \ x < a \ g x \ 0" using g0_neq_0 by (simp add: g_def) have f: "DERIV f x :> (f' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ f0[OF x]]) (auto simp: f_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have g: "DERIV g x :> (g' x)" if x: "0 < x" "x < a" for x using that by (intro DERIV_cong_ev[THEN iffD1, OF _ _ _ g0[OF x]]) (auto simp: g_def eventually_nhds_metric dist_real_def intro!: exI[of _ x]) have "isCont f 0" unfolding f_def by (intro isCont_If_ge f_0 continuous_const) have "isCont g 0" unfolding g_def by (intro isCont_If_ge g_0 continuous_const) have "\\. \x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" proof (rule bchoice, rule ballI) fix x assume "x \ {0 <..< a}" then have x[arith]: "0 < x" "x < a" by auto with g'_neq_0 g_neq_0 \g 0 = 0\ have g': "\x. 0 < x \ x < a \ 0 \ g' x" "g 0 \ g x" by auto have "\x. 0 \ x \ x < a \ isCont f x" using \isCont f 0\ f by (auto intro: DERIV_isCont simp: le_less) moreover have "\x. 0 \ x \ x < a \ isCont g x" using \isCont g 0\ g by (auto intro: DERIV_isCont simp: le_less) ultimately have "\c. 0 < c \ c < x \ (f x - f 0) * g' c = (g x - g 0) * f' c" using f g \x < a\ by (intro GMVT') auto then obtain c where *: "0 < c" "c < x" "(f x - f 0) * g' c = (g x - g 0) * f' c" by blast moreover from * g'(1)[of c] g'(2) have "(f x - f 0) / (g x - g 0) = f' c / g' c" by (simp add: field_simps) ultimately show "\y. 0 < y \ y < x \ f x / g x = f' y / g' y" using \f 0 = 0\ \g 0 = 0\ by (auto intro!: exI[of _ c]) qed then obtain \ where "\x\{0 <..< a}. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)" .. then have \: "eventually (\x. 0 < \ x \ \ x < x \ f x / g x = f' (\ x) / g' (\ x)) (at_right 0)" unfolding eventually_at by (intro exI[of _ a]) (auto simp: dist_real_def) moreover from \ have "eventually (\x. norm (\ x) \ x) (at_right 0)" by eventually_elim auto then have "((\x. norm (\ x)) \ 0) (at_right 0)" by (rule_tac real_tendsto_sandwich[where f="\x. 0" and h="\x. x"]) auto then have "(\ \ 0) (at_right 0)" by (rule tendsto_norm_zero_cancel) with \ have "filterlim \ (at_right 0) (at_right 0)" by (auto elim!: eventually_mono simp: filterlim_at) from this lim have "filterlim (\t. f' (\ t) / g' (\ t)) F (at_right 0)" by (rule_tac filterlim_compose[of _ _ _ \]) ultimately have "filterlim (\t. f t / g t) F (at_right 0)" (is ?P) by (rule_tac filterlim_cong[THEN iffD1, OF refl refl]) (auto elim: eventually_mono) also have "?P \ ?thesis" by (rule filterlim_cong) (auto simp: f_def g_def eventually_at_filter) finally show ?thesis . qed lemma lhopital_right: "(f \ 0) (at_right x) \ (g \ 0) (at_right x) \ eventually (\x. g x \ 0) (at_right x) \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ filterlim (\ x. (f' x / g' x)) F (at_right x) \ filterlim (\ x. f x / g x) F (at_right x)" for x :: real unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0) lemma lhopital_left: "(f \ 0) (at_left x) \ (g \ 0) (at_left x) \ eventually (\x. g x \ 0) (at_left x) \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ filterlim (\ x. (f' x / g' x)) F (at_left x) \ filterlim (\ x. f x / g x) F (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital: "(f \ 0) (at x) \ (g \ 0) (at x) \ eventually (\x. g x \ 0) (at x) \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ filterlim (\ x. (f' x / g' x)) F (at x) \ filterlim (\ x. f x / g x) F (at x)" for x :: real unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right[of f x g g' f'] lhopital_left[of f x g g' f']) lemma lhopital_right_0_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_right 0. g x :> at_top" and ev: "eventually (\x. g' x \ 0) (at_right 0)" "eventually (\x. DERIV f x :> f' x) (at_right 0)" "eventually (\x. DERIV g x :> g' x) (at_right 0)" and lim: "((\ x. (f' x / g' x)) \ x) (at_right 0)" shows "((\ x. f x / g x) \ x) (at_right 0)" unfolding tendsto_iff proof safe fix e :: real assume "0 < e" with lim[unfolded tendsto_iff, rule_format, of "e / 4"] have "eventually (\t. dist (f' t / g' t) x < e / 4) (at_right 0)" by simp from eventually_conj[OF eventually_conj[OF ev(1) ev(2)] eventually_conj[OF ev(3) this]] obtain a where [arith]: "0 < a" and g'_neq_0: "\x. 0 < x \ x < a \ g' x \ 0" and f0: "\x. 0 < x \ x \ a \ DERIV f x :> (f' x)" and g0: "\x. 0 < x \ x \ a \ DERIV g x :> (g' x)" and Df: "\t. 0 < t \ t < a \ dist (f' t / g' t) x < e / 4" unfolding eventually_at_le by (auto simp: dist_real_def) from Df have "eventually (\t. t < a) (at_right 0)" "eventually (\t::real. 0 < t) (at_right 0)" unfolding eventually_at by (auto intro!: exI[of _ a] simp: dist_real_def) moreover have "eventually (\t. 0 < g t) (at_right 0)" "eventually (\t. g a < g t) (at_right 0)" using g_0 by (auto elim: eventually_mono simp: filterlim_at_top_dense) moreover have inv_g: "((\x. inverse (g x)) \ 0) (at_right 0)" using tendsto_inverse_0 filterlim_mono[OF g_0 at_top_le_at_infinity order_refl] by (rule filterlim_compose) then have "((\x. norm (1 - g a * inverse (g x))) \ norm (1 - g a * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\x. norm (1 - g a / g x)) \ 1) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of 1] have "eventually (\x. norm (1 - g a / g x) < 2) (at_right 0)" by (auto elim!: eventually_mono simp: dist_real_def) moreover from inv_g have "((\t. norm ((f a - x * g a) * inverse (g t))) \ norm ((f a - x * g a) * 0)) (at_right 0)" by (intro tendsto_intros) then have "((\t. norm (f a - x * g a) / norm (g t)) \ 0) (at_right 0)" by (simp add: inverse_eq_divide) from this[unfolded tendsto_iff, rule_format, of "e / 2"] \0 < e\ have "eventually (\t. norm (f a - x * g a) / norm (g t) < e / 2) (at_right 0)" by (auto simp: dist_real_def) ultimately show "eventually (\t. dist (f t / g t) x < e) (at_right 0)" proof eventually_elim fix t assume t[arith]: "0 < t" "t < a" "g a < g t" "0 < g t" assume ineq: "norm (1 - g a / g t) < 2" "norm (f a - x * g a) / norm (g t) < e / 2" have "\y. t < y \ y < a \ (g a - g t) * f' y = (f a - f t) * g' y" using f0 g0 t(1,2) by (intro GMVT') (force intro!: DERIV_isCont)+ then obtain y where [arith]: "t < y" "y < a" and D_eq0: "(g a - g t) * f' y = (f a - f t) * g' y" by blast from D_eq0 have D_eq: "(f t - f a) / (g t - g a) = f' y / g' y" using \g a < g t\ g'_neq_0[of y] by (auto simp add: field_simps) have *: "f t / g t - x = ((f t - f a) / (g t - g a) - x) * (1 - g a / g t) + (f a - x * g a) / g t" by (simp add: field_simps) have "norm (f t / g t - x) \ norm (((f t - f a) / (g t - g a) - x) * (1 - g a / g t)) + norm ((f a - x * g a) / g t)" unfolding * by (rule norm_triangle_ineq) also have "\ = dist (f' y / g' y) x * norm (1 - g a / g t) + norm (f a - x * g a) / norm (g t)" by (simp add: abs_mult D_eq dist_real_def) also have "\ < (e / 4) * 2 + e / 2" using ineq Df[of y] \0 < e\ by (intro add_le_less_mono mult_mono) auto finally show "dist (f t / g t) x < e" by (simp add: dist_real_def) qed qed lemma lhopital_right_at_top: "LIM x at_right x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at_right x) \ eventually (\x. DERIV f x :> f' x) (at_right x) \ eventually (\x. DERIV g x :> g' x) (at_right x) \ ((\ x. (f' x / g' x)) \ y) (at_right x) \ ((\ x. f x / g x) \ y) (at_right x)" unfolding eventually_at_right_to_0[of _ x] filterlim_at_right_to_0[of _ _ x] DERIV_shift by (rule lhopital_right_0_at_top) lemma lhopital_left_at_top: "LIM x at_left x. g x :> at_top \ eventually (\x. g' x \ 0) (at_left x) \ eventually (\x. DERIV f x :> f' x) (at_left x) \ eventually (\x. DERIV g x :> g' x) (at_left x) \ ((\ x. (f' x / g' x)) \ y) (at_left x) \ ((\ x. f x / g x) \ y) (at_left x)" for x :: real unfolding eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror by (rule lhopital_right_at_top[where f'="\x. - f' (- x)"]) (auto simp: DERIV_mirror) lemma lhopital_at_top: "LIM x at x. (g::real \ real) x :> at_top \ eventually (\x. g' x \ 0) (at x) \ eventually (\x. DERIV f x :> f' x) (at x) \ eventually (\x. DERIV g x :> g' x) (at x) \ ((\ x. (f' x / g' x)) \ y) (at x) \ ((\ x. f x / g x) \ y) (at x)" unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top[of g x g' f f'] lhopital_left_at_top[of g x g' f f']) lemma lhospital_at_top_at_top: fixes f g :: "real \ real" assumes g_0: "LIM x at_top. g x :> at_top" and g': "eventually (\x. g' x \ 0) at_top" and Df: "eventually (\x. DERIV f x :> f' x) at_top" and Dg: "eventually (\x. DERIV g x :> g' x) at_top" and lim: "((\ x. (f' x / g' x)) \ x) at_top" shows "((\ x. f x / g x) \ x) at_top" unfolding filterlim_at_top_to_right proof (rule lhopital_right_0_at_top) let ?F = "\x. f (inverse x)" let ?G = "\x. g (inverse x)" let ?R = "at_right (0::real)" let ?D = "\f' x. f' (inverse x) * - (inverse x ^ Suc (Suc 0))" show "LIM x ?R. ?G x :> at_top" using g_0 unfolding filterlim_at_top_to_right . show "eventually (\x. DERIV ?G x :> ?D g' x) ?R" unfolding eventually_at_right_to_top using Dg eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. DERIV ?F x :> ?D f' x) ?R" unfolding eventually_at_right_to_top using Df eventually_ge_at_top[where c=1] by eventually_elim (rule derivative_eq_intros DERIV_chain'[where f=inverse] | simp)+ show "eventually (\x. ?D g' x \ 0) ?R" unfolding eventually_at_right_to_top using g' eventually_ge_at_top[where c=1] by eventually_elim auto show "((\x. ?D f' x / ?D g' x) \ x) ?R" unfolding filterlim_at_right_to_top apply (intro filterlim_cong[THEN iffD2, OF refl refl _ lim]) using eventually_ge_at_top[where c=1] by eventually_elim simp qed lemma lhopital_right_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_right a)" shows "filterlim (\ x. f x / g x) at_top (at_right a)" proof - from lim have pos: "eventually (\x. f' x / g' x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast have "((\x. g x / f x) \ 0) (at_right a)" proof (rule lhopital_right_at_top) from pos show "eventually (\x. f' x \ 0) (at_right a)" by eventually_elim auto from tendsto_inverse_0_at_top[OF lim] show "((\x. g' x / f' x) \ 0) (at_right a)" by simp qed fact+ moreover from f_0 g_0 have "eventually (\x. f x > 0) (at_right a)" "eventually (\x. g x > 0) (at_right a)" unfolding filterlim_at_top_dense by blast+ hence "eventually (\x. g x / f x > 0) (at_right a)" by eventually_elim simp ultimately have "filterlim (\x. inverse (g x / f x)) at_top (at_right a)" by (rule filterlim_inverse_at_top) thus ?thesis by simp qed lemma lhopital_right_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_right a. f x :> at_top" assumes g_0: "LIM x at_right a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_right a)" "eventually (\x. DERIV g x :> g' x) (at_right a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_right a)" shows "filterlim (\ x. f x / g x) at_bot (at_right a)" proof - from ev(2) have ev': "eventually (\x. DERIV (\x. -g x) x :> -g' x) (at_right a)" by eventually_elim (auto intro: derivative_intros) have "filterlim (\x. f x / (-g x)) at_top (at_right a)" by (rule lhopital_right_at_top_at_top[where f' = f' and g' = "\x. -g' x"]) (insert assms ev', auto simp: filterlim_uminus_at_bot) hence "filterlim (\x. -(f x / g x)) at_top (at_right a)" by simp thus ?thesis by (simp add: filterlim_uminus_at_bot) qed lemma lhopital_left_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at_left a)" shows "filterlim (\ x. f x / g x) at_top (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_top[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_left_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at_left a. f x :> at_top" assumes g_0: "LIM x at_left a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at_left a)" "eventually (\x. DERIV g x :> g' x) (at_left a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at_left a)" shows "filterlim (\ x. f x / g x) at_bot (at_left a)" by (insert assms, unfold eventually_at_left_to_right filterlim_at_left_to_right DERIV_mirror, rule lhopital_right_at_top_at_bot[where f'="\x. - f' (- x)"]) (insert assms, auto simp: DERIV_mirror) lemma lhopital_at_top_at_top: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_top" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_top (at a)" shows "filterlim (\ x. f x / g x) at_top (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_top[of f a g f' g'] lhopital_left_at_top_at_top[of f a g f' g']) lemma lhopital_at_top_at_bot: fixes f g :: "real \ real" assumes f_0: "LIM x at a. f x :> at_top" assumes g_0: "LIM x at a. g x :> at_bot" and ev: "eventually (\x. DERIV f x :> f' x) (at a)" "eventually (\x. DERIV g x :> g' x) (at a)" and lim: "filterlim (\ x. (f' x / g' x)) at_bot (at a)" shows "filterlim (\ x. f x / g x) at_bot (at a)" using assms unfolding eventually_at_split filterlim_at_split by (auto intro!: lhopital_right_at_top_at_bot[of f a g f' g'] lhopital_left_at_top_at_bot[of f a g f' g']) end