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,1183 +1,1157 @@ (* 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 Brouwer_Fixpoint "HOL-Library.Nonpos_Ints" + 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 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) 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) 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 -lemma has_field_derivative_inverse_strong: - fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" - shows "DERIV f x :> f' \ - f' \ 0 \ - open S \ - x \ S \ - continuous_on S f \ - (\z. z \ S \ g (f z) = z) - \ DERIV g (f x) :> inverse (f')" - unfolding has_field_derivative_def - apply (rule has_derivative_inverse_strong [of S x f g ]) - by auto - -lemma has_field_derivative_inverse_strong_x: - fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" - shows "DERIV f (g y) :> f' \ - f' \ 0 \ - open S \ - continuous_on S f \ - g y \ S \ f(g y) = y \ - (\z. z \ S \ g (f z) = z) - \ DERIV g y :> inverse (f')" - unfolding has_field_derivative_def - apply (rule has_derivative_inverse_strong_x [of S g y f]) - by auto - 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,4306 +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) 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) 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 -subsection\Formulation of loop homotopy in terms of maps out of type complex\ - -lemma homotopic_circlemaps_imp_homotopic_loops: - assumes "homotopic_with_canon (\h. True) (sphere 0 1) S f g" - shows "homotopic_loops S (f \ exp \ (\t. 2 * of_real pi * of_real t * \)) - (g \ exp \ (\t. 2 * of_real pi * of_real t * \))" -proof - - have "homotopic_with_canon (\f. True) {z. cmod z = 1} S f g" - using assms by (auto simp: sphere_def) - moreover have "continuous_on {0..1} (exp \ (\t. 2 * of_real pi * of_real t * \))" - by (intro continuous_intros) - moreover have "(exp \ (\t. 2 * of_real pi * of_real t * \)) ` {0..1} \ {z. cmod z = 1}" - by (auto simp: norm_mult) - ultimately - show ?thesis - apply (simp add: homotopic_loops_def comp_assoc) - apply (rule homotopic_with_compose_continuous_right) - apply (auto simp: pathstart_def pathfinish_def) - done -qed - -lemma homotopic_loops_imp_homotopic_circlemaps: - assumes "homotopic_loops S p q" - shows "homotopic_with_canon (\h. True) (sphere 0 1) S - (p \ (\z. (Arg2pi z / (2 * pi)))) - (q \ (\z. (Arg2pi z / (2 * pi))))" -proof - - obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" - and him: "h ` ({0..1} \ {0..1}) \ S" - and h0: "(\x. h (0, x) = p x)" - and h1: "(\x. h (1, x) = q x)" - and h01: "(\t\{0..1}. h (t, 1) = h (t, 0)) " - using assms - by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def) - define j where "j \ \z. if 0 \ Im (snd z) - then h (fst z, Arg2pi (snd z) / (2 * pi)) - else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))" - have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \ Arg2pi y = 0 \ Arg2pi (cnj y) = 0" if "cmod y = 1" for y - using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj field_split_simps) - show ?thesis - proof (simp add: homotopic_with; intro conjI ballI exI) - show "continuous_on ({0..1} \ sphere 0 1) (\w. h (fst w, Arg2pi (snd w) / (2 * pi)))" - proof (rule continuous_on_eq) - show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \ {0..1} \ sphere 0 1" for x - using Arg2pi_eq that h01 by (force simp: j_def) - have eq: "S = S \ (UNIV \ {z. 0 \ Im z}) \ S \ (UNIV \ {z. Im z \ 0})" for S :: "(real*complex)set" - by auto - have c1: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. 0 \ Im z}) (\x. h (fst x, Arg2pi (snd x) / (2 * pi)))" - apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) - apply (auto simp: Arg2pi) - apply (meson Arg2pi_lt_2pi linear not_le) - done - have c2: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. Im z \ 0}) (\x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))" - apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) - apply (auto simp: Arg2pi) - apply (meson Arg2pi_lt_2pi linear not_le) - done - show "continuous_on ({0..1} \ sphere 0 1) j" - apply (simp add: j_def) - apply (subst eq) - apply (rule continuous_on_cases_local) - apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2) - using Arg2pi_eq h01 - by force - qed - have "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ h ` ({0..1} \ {0..1})" - by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le) - also have "... \ S" - using him by blast - finally show "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ S" . - qed (auto simp: h0 h1) -qed - -lemma simply_connected_homotopic_loops: - "simply_connected S \ - (\p q. homotopic_loops S p p \ homotopic_loops S q q \ homotopic_loops S p q)" -unfolding simply_connected_def using homotopic_loops_refl by metis - - -lemma simply_connected_eq_homotopic_circlemaps1: - fixes f :: "complex \ 'a::topological_space" and g :: "complex \ 'a" - assumes S: "simply_connected S" - and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \ S" - and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \ S" - shows "homotopic_with_canon (\h. True) (sphere 0 1) S f g" -proof - - have "homotopic_loops S (f \ exp \ (\t. of_real(2 * pi * t) * \)) (g \ exp \ (\t. of_real(2 * pi * t) * \))" - apply (rule S [unfolded simply_connected_homotopic_loops, rule_format]) - apply (simp add: homotopic_circlemaps_imp_homotopic_loops homotopic_with_refl contf fim contg gim) - done - then show ?thesis - apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps]) - apply (auto simp: o_def complex_norm_eq_1_exp mult.commute) - done -qed - -lemma simply_connected_eq_homotopic_circlemaps2a: - fixes h :: "complex \ 'a::topological_space" - assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \ S" - and hom: "\f g::complex \ 'a. - \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; - continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ - \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" - shows "\a. homotopic_with_canon (\h. True) (sphere 0 1) S h (\x. a)" - apply (rule_tac x="h 1" in exI) - apply (rule hom) - using assms - by (auto simp: continuous_on_const) - -lemma simply_connected_eq_homotopic_circlemaps2b: - fixes S :: "'a::real_normed_vector set" - assumes "\f g::complex \ 'a. - \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; - continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ - \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" - shows "path_connected S" -proof (clarsimp simp add: path_connected_eq_homotopic_points) - fix a b - assume "a \ S" "b \ S" - then show "homotopic_loops S (linepath a a) (linepath b b)" - using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\x. a" "\x. b"]] - by (auto simp: o_def continuous_on_const linepath_def) -qed - -lemma simply_connected_eq_homotopic_circlemaps3: - fixes h :: "complex \ 'a::real_normed_vector" - assumes "path_connected S" - and hom: "\f::complex \ 'a. - \continuous_on (sphere 0 1) f; f `(sphere 0 1) \ S\ - \ \a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)" - shows "simply_connected S" -proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms) - fix p - assume p: "path p" "path_image p \ S" "pathfinish p = pathstart p" - then have "homotopic_loops S p p" - by (simp add: homotopic_loops_refl) - then obtain a where homp: "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. Arg2pi z / (2 * pi))) (\x. a)" - by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom) - show "\a. a \ S \ homotopic_loops S p (linepath a a)" - proof (intro exI conjI) - show "a \ S" - using homotopic_with_imp_subset2 [OF homp] - by (metis dist_0_norm image_subset_iff mem_sphere norm_one) - have teq: "\t. \0 \ t; t \ 1\ - \ t = Arg2pi (exp (2 * of_real pi * of_real t * \)) / (2 * pi) \ t=1 \ Arg2pi (exp (2 * of_real pi * of_real t * \)) = 0" - apply (rule disjCI) - using Arg2pi_of_real [of 1] apply (auto simp: Arg2pi_exp) - done - have "homotopic_loops S p (p \ (\z. Arg2pi z / (2 * pi)) \ exp \ (\t. 2 * complex_of_real pi * complex_of_real t * \))" - apply (rule homotopic_loops_eq [OF p]) - using p teq apply (fastforce simp: pathfinish_def pathstart_def) - done - then - show "homotopic_loops S p (linepath a a)" - by (simp add: linepath_refl homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]]) - qed -qed - - -proposition simply_connected_eq_homotopic_circlemaps: - fixes S :: "'a::real_normed_vector set" - shows "simply_connected S \ - (\f g::complex \ 'a. - continuous_on (sphere 0 1) f \ f ` (sphere 0 1) \ S \ - continuous_on (sphere 0 1) g \ g ` (sphere 0 1) \ S - \ homotopic_with_canon (\h. True) (sphere 0 1) S f g)" - apply (rule iffI) - apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1) - by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) - -proposition simply_connected_eq_contractible_circlemap: - fixes S :: "'a::real_normed_vector set" - shows "simply_connected S \ - path_connected S \ - (\f::complex \ 'a. - continuous_on (sphere 0 1) f \ f `(sphere 0 1) \ S - \ (\a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)))" - apply (rule iffI) - apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b) - using simply_connected_eq_homotopic_circlemaps3 by blast - -corollary homotopy_eqv_simple_connectedness: - fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" - shows "S homotopy_eqv T \ simply_connected S \ simply_connected T" - by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality) - - -subsection\Homeomorphism of simple closed curves to circles\ - -proposition homeomorphic_simple_path_image_circle: - fixes a :: complex and \ :: "real \ 'a::t2_space" - assumes "simple_path \" and loop: "pathfinish \ = pathstart \" and "0 < r" - shows "(path_image \) homeomorphic sphere a r" -proof - - have "homotopic_loops (path_image \) \ \" - by (simp add: assms homotopic_loops_refl simple_path_imp_path) - then have hom: "homotopic_with_canon (\h. True) (sphere 0 1) (path_image \) - (\ \ (\z. Arg2pi z / (2*pi))) (\ \ (\z. Arg2pi z / (2*pi)))" - by (rule homotopic_loops_imp_homotopic_circlemaps) - have "\g. homeomorphism (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) g" - proof (rule homeomorphism_compact) - show "continuous_on (sphere 0 1) (\ \ (\z. Arg2pi z / (2*pi)))" - using hom homotopic_with_imp_continuous by blast - show "inj_on (\ \ (\z. Arg2pi z / (2*pi))) (sphere 0 1)" - proof - fix x y - assume xy: "x \ sphere 0 1" "y \ sphere 0 1" - and eq: "(\ \ (\z. Arg2pi z / (2*pi))) x = (\ \ (\z. Arg2pi z / (2*pi))) y" - then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))" - proof - - have "(Arg2pi x / (2*pi)) \ {0..1}" "(Arg2pi y / (2*pi)) \ {0..1}" - using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+ - with eq show ?thesis - using \simple_path \\ Arg2pi_lt_2pi unfolding simple_path_def o_def - by (metis eq_divide_eq_1 not_less_iff_gr_or_eq) - qed - with xy show "x = y" - by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere) - qed - have "\z. cmod z = 1 \ \x\{0..1}. \ (Arg2pi z / (2*pi)) = \ x" - by (metis Arg2pi_ge_0 Arg2pi_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral) - moreover have "\z\sphere 0 1. \ x = \ (Arg2pi z / (2*pi))" if "0 \ x" "x \ 1" for x - proof (cases "x=1") - case True - with Arg2pi_of_real [of 1] loop show ?thesis - by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \0 \ x\) - next - case False - then have *: "(Arg2pi (exp (\*(2* of_real pi* of_real x))) / (2*pi)) = x" - using that by (auto simp: Arg2pi_exp field_split_simps) - show ?thesis - by (rule_tac x="exp(\ * of_real(2*pi*x))" in bexI) (auto simp: *) - qed - ultimately show "(\ \ (\z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \" - by (auto simp: path_image_def image_iff) - qed auto - then have "path_image \ homeomorphic sphere (0::complex) 1" - using homeomorphic_def homeomorphic_sym by blast - also have "... homeomorphic sphere a r" - by (simp add: assms homeomorphic_spheres) - finally show ?thesis . -qed - -lemma homeomorphic_simple_path_images: - fixes \1 :: "real \ 'a::t2_space" and \2 :: "real \ 'b::t2_space" - assumes "simple_path \1" and loop: "pathfinish \1 = pathstart \1" - assumes "simple_path \2" and loop: "pathfinish \2 = pathstart \2" - shows "(path_image \1) homeomorphic (path_image \2)" - by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero) - end diff --git a/src/HOL/Analysis/Conformal_Mappings.thy b/src/HOL/Analysis/Conformal_Mappings.thy --- a/src/HOL/Analysis/Conformal_Mappings.thy +++ b/src/HOL/Analysis/Conformal_Mappings.thy @@ -1,5070 +1,5096 @@ section \Conformal Mappings and Consequences of Cauchy's Integral Theorem\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\ text\Also Cauchy's residue theorem by Wenda Li (2016)\ theory Conformal_Mappings imports Cauchy_Integral_Theorem begin (* FIXME mv to Cauchy_Integral_Theorem.thy *) subsection\Cauchy's inequality and more versions of Liouville\ lemma Cauchy_higher_deriv_bound: assumes holf: "f holomorphic_on (ball z r)" and contf: "continuous_on (cball z r) f" and fin : "\w. w \ ball z r \ f w \ ball y B0" and "0 < r" and "0 < n" shows "norm ((deriv ^^ n) f z) \ (fact n) * B0 / r^n" proof - have "0 < B0" using \0 < r\ fin [of z] by (metis ball_eq_empty ex_in_conv fin not_less) have le_B0: "\w. cmod (w - z) \ r \ cmod (f w - y) \ B0" apply (rule continuous_on_closure_norm_le [of "ball z r" "\w. f w - y"]) apply (auto simp: \0 < r\ dist_norm norm_minus_commute) apply (rule continuous_intros contf)+ using fin apply (simp add: dist_commute dist_norm less_eq_real_def) done have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w) z - (deriv ^^ n) (\w. y) z" using \0 < n\ by simp also have "... = (deriv ^^ n) (\w. f w - y) z" by (rule higher_deriv_diff [OF holf, symmetric]) (auto simp: \0 < r\) finally have "(deriv ^^ n) f z = (deriv ^^ n) (\w. f w - y) z" . have contf': "continuous_on (cball z r) (\u. f u - y)" by (rule contf continuous_intros)+ have holf': "(\u. (f u - y)) holomorphic_on (ball z r)" by (simp add: holf holomorphic_on_diff) define a where "a = (2 * pi)/(fact n)" have "0 < a" by (simp add: a_def) have "B0/r^(Suc n)*2 * pi * r = a*((fact n)*B0/r^n)" using \0 < r\ by (simp add: a_def field_split_simps) have der_dif: "(deriv ^^ n) (\w. f w - y) z = (deriv ^^ n) f z" using \0 < r\ \0 < n\ by (auto simp: higher_deriv_diff [OF holf holomorphic_on_const]) have "norm ((2 * of_real pi * \)/(fact n) * (deriv ^^ n) (\w. f w - y) z) \ (B0/r^(Suc n)) * (2 * pi * r)" apply (rule has_contour_integral_bound_circlepath [of "(\u. (f u - y)/(u - z)^(Suc n))" _ z]) using Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf' holf'] using \0 < B0\ \0 < r\ apply (auto simp: norm_divide norm_mult norm_power divide_simps le_B0) done then show ?thesis using \0 < r\ by (auto simp: norm_divide norm_mult norm_power field_simps der_dif le_B0) qed lemma Cauchy_inequality: assumes holf: "f holomorphic_on (ball \ r)" and contf: "continuous_on (cball \ r) f" and "0 < r" and nof: "\x. norm(\-x) = r \ norm(f x) \ B" shows "norm ((deriv ^^ n) f \) \ (fact n) * B / r^n" proof - obtain x where "norm (\-x) = r" by (metis abs_of_nonneg add_diff_cancel_left' \0 < r\ diff_add_cancel dual_order.strict_implies_order norm_of_real) then have "0 \ B" by (metis nof norm_not_less_zero not_le order_trans) have "((\u. f u / (u - \) ^ Suc n) has_contour_integral (2 * pi) * \ / fact n * (deriv ^^ n) f \) (circlepath \ r)" apply (rule Cauchy_has_contour_integral_higher_derivative_circlepath [OF contf holf]) using \0 < r\ by simp then have "norm ((2 * pi * \)/(fact n) * (deriv ^^ n) f \) \ (B / r^(Suc n)) * (2 * pi * r)" apply (rule has_contour_integral_bound_circlepath) using \0 \ B\ \0 < r\ apply (simp_all add: norm_divide norm_power nof frac_le norm_minus_commute del: power_Suc) done then show ?thesis using \0 < r\ by (simp add: norm_divide norm_mult field_simps) qed lemma Liouville_polynomial: assumes holf: "f holomorphic_on UNIV" and nof: "\z. A \ norm z \ norm(f z) \ B * norm z ^ n" shows "f \ = (\k\n. (deriv^^k) f 0 / fact k * \ ^ k)" proof (cases rule: le_less_linear [THEN disjE]) assume "B \ 0" then have "\z. A \ norm z \ norm(f z) = 0" by (metis nof less_le_trans zero_less_mult_iff neqE norm_not_less_zero norm_power not_le) then have f0: "(f \ 0) at_infinity" using Lim_at_infinity by force then have [simp]: "f = (\w. 0)" using Liouville_weak [OF holf, of 0] by (simp add: eventually_at_infinity f0) meson show ?thesis by simp next assume "0 < B" have "((\k. (deriv ^^ k) f 0 / (fact k) * (\ - 0)^k) sums f \)" apply (rule holomorphic_power_series [where r = "norm \ + 1"]) using holf holomorphic_on_subset apply auto done then have sumsf: "((\k. (deriv ^^ k) f 0 / (fact k) * \^k) sums f \)" by simp have "(deriv ^^ k) f 0 / fact k * \ ^ k = 0" if "k>n" for k proof (cases "(deriv ^^ k) f 0 = 0") case True then show ?thesis by simp next case False define w where "w = complex_of_real (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" have "1 \ abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge1: "1 \ norm w" by (metis norm_of_real w_def) then have "w \ 0" by auto have kB: "0 < fact k * B" using \0 < B\ by simp then have "0 \ fact k * B / cmod ((deriv ^^ k) f 0)" by simp then have wgeA: "A \ cmod w" by (simp only: w_def norm_of_real) have "fact k * B / cmod ((deriv ^^ k) f 0) < abs (fact k * B / cmod ((deriv ^^ k) f 0) + (\A\ + 1))" using \0 < B\ by simp then have wge: "fact k * B / cmod ((deriv ^^ k) f 0) < norm w" by (metis norm_of_real w_def) then have "fact k * B / norm w < cmod ((deriv ^^ k) f 0)" using False by (simp add: field_split_simps mult.commute split: if_split_asm) also have "... \ fact k * (B * norm w ^ n) / norm w ^ k" apply (rule Cauchy_inequality) using holf holomorphic_on_subset apply force using holf holomorphic_on_imp_continuous_on holomorphic_on_subset apply blast using \w \ 0\ apply simp by (metis nof wgeA dist_0_norm dist_norm) also have "... = fact k * (B * 1 / cmod w ^ (k-n))" apply (simp only: mult_cancel_left times_divide_eq_right [symmetric]) using \k>n\ \w \ 0\ \0 < B\ apply (simp add: field_split_simps semiring_normalization_rules) done also have "... = fact k * B / cmod w ^ (k-n)" by simp finally have "fact k * B / cmod w < fact k * B / cmod w ^ (k - n)" . then have "1 / cmod w < 1 / cmod w ^ (k - n)" by (metis kB divide_inverse inverse_eq_divide mult_less_cancel_left_pos) then have "cmod w ^ (k - n) < cmod w" by (metis frac_le le_less_trans norm_ge_zero norm_one not_less order_refl wge1 zero_less_one) with self_le_power [OF wge1] have False by (meson diff_is_0_eq not_gr0 not_le that) then show ?thesis by blast qed then have "(deriv ^^ (k + Suc n)) f 0 / fact (k + Suc n) * \ ^ (k + Suc n) = 0" for k using not_less_eq by blast then have "(\i. (deriv ^^ (i + Suc n)) f 0 / fact (i + Suc n) * \ ^ (i + Suc n)) sums 0" by (rule sums_0) with sums_split_initial_segment [OF sumsf, where n = "Suc n"] show ?thesis using atLeast0AtMost lessThan_Suc_atMost sums_unique2 by fastforce qed text\Every bounded entire function is a constant function.\ theorem Liouville_theorem: assumes holf: "f holomorphic_on UNIV" and bf: "bounded (range f)" obtains c where "\z. f z = c" proof - obtain B where "\z. cmod (f z) \ B" by (meson bf bounded_pos rangeI) then show ?thesis using Liouville_polynomial [OF holf, of 0 B 0, simplified] that by blast qed text\A holomorphic function f has only isolated zeros unless f is 0.\ lemma powser_0_nonzero: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes r: "0 < r" and sm: "\x. norm (x - \) < r \ (\n. a n * (x - \) ^ n) sums (f x)" and [simp]: "f \ = 0" and m0: "a m \ 0" and "m>0" obtains s where "0 < s" and "\z. z \ cball \ s - {\} \ f z \ 0" proof - have "r \ conv_radius a" using sm sums_summable by (auto simp: le_conv_radius_iff [where \=\]) obtain m where am: "a m \ 0" and az [simp]: "(\n. n a n = 0)" apply (rule_tac m = "LEAST n. a n \ 0" in that) using m0 apply (rule LeastI2) apply (fastforce intro: dest!: not_less_Least)+ done define b where "b i = a (i+m) / a m" for i define g where "g x = suminf (\i. b i * (x - \) ^ i)" for x have [simp]: "b 0 = 1" by (simp add: am b_def) { fix x::'a assume "norm (x - \) < r" then have "(\n. (a m * (x - \)^m) * (b n * (x - \)^n)) sums (f x)" using am az sm sums_zero_iff_shift [of m "(\n. a n * (x - \) ^ n)" "f x"] by (simp add: b_def monoid_mult_class.power_add algebra_simps) then have "x \ \ \ (\n. b n * (x - \)^n) sums (f x / (a m * (x - \)^m))" using am by (simp add: sums_mult_D) } note bsums = this then have "norm (x - \) < r \ summable (\n. b n * (x - \)^n)" for x using sums_summable by (cases "x=\") auto then have "r \ conv_radius b" by (simp add: le_conv_radius_iff [where \=\]) then have "r/2 < conv_radius b" using not_le order_trans r by fastforce then have "continuous_on (cball \ (r/2)) g" using powser_continuous_suminf [of "r/2" b \] by (simp add: g_def) then obtain s where "s>0" "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ dist (g x) (g \) < 1/2" apply (rule continuous_onE [where x=\ and e = "1/2"]) using r apply (auto simp: norm_minus_commute dist_norm) done moreover have "g \ = 1" by (simp add: g_def) ultimately have gnz: "\x. \norm (x - \) \ s; norm (x - \) \ r/2\ \ (g x) \ 0" by fastforce have "f x \ 0" if "x \ \" "norm (x - \) \ s" "norm (x - \) \ r/2" for x using bsums [of x] that gnz [of x] apply (auto simp: g_def) using r sums_iff by fastforce then show ?thesis apply (rule_tac s="min s (r/2)" in that) using \0 < r\ \0 < s\ by (auto simp: dist_commute dist_norm) qed subsection \Analytic continuation\ proposition isolated_zeros: assumes holf: "f holomorphic_on S" and "open S" "connected S" "\ \ S" "f \ = 0" "\ \ S" "f \ \ 0" - obtains r where "0 < r" and "ball \ r \ S" and + obtains r where "0 < r" and "ball \ r \ S" and "\z. z \ ball \ r - {\} \ f z \ 0" proof - obtain r where "0 < r" and r: "ball \ r \ S" using \open S\ \\ \ S\ open_contains_ball_eq by blast have powf: "((\n. (deriv ^^ n) f \ / (fact n) * (z - \)^n) sums f z)" if "z \ ball \ r" for z apply (rule holomorphic_power_series [OF _ that]) apply (rule holomorphic_on_subset [OF holf r]) done obtain m where m: "(deriv ^^ m) f \ / (fact m) \ 0" using holomorphic_fun_eq_0_on_connected [OF holf \open S\ \connected S\ _ \\ \ S\ \\ \ S\] \f \ \ 0\ by auto then have "m \ 0" using assms(5) funpow_0 by fastforce obtain s where "0 < s" and s: "\z. z \ cball \ s - {\} \ f z \ 0" apply (rule powser_0_nonzero [OF \0 < r\ powf \f \ = 0\ m]) using \m \ 0\ by (auto simp: dist_commute dist_norm) have "0 < min r s" by (simp add: \0 < r\ \0 < s\) then show ?thesis apply (rule that) using r s by auto qed proposition analytic_continuation: assumes holf: "f holomorphic_on S" and "open S" and "connected S" and "U \ S" and "\ \ S" and "\ islimpt U" and fU0 [simp]: "\z. z \ U \ f z = 0" and "w \ S" shows "f w = 0" proof - obtain e where "0 < e" and e: "cball \ e \ S" using \open S\ \\ \ S\ open_contains_cball_eq by blast define T where "T = cball \ e \ U" have contf: "continuous_on (closure T) f" by (metis T_def closed_cball closure_minimal e holf holomorphic_on_imp_continuous_on holomorphic_on_subset inf.cobounded1) have fT0 [simp]: "\x. x \ T \ f x = 0" by (simp add: T_def) have "\r. \\e>0. \x'\U. x' \ \ \ dist x' \ < e; 0 < r\ \ \x'\cball \ e \ U. x' \ \ \ dist x' \ < r" by (metis \0 < e\ IntI dist_commute less_eq_real_def mem_cball min_less_iff_conj) then have "\ islimpt T" using \\ islimpt U\ by (auto simp: T_def islimpt_approachable) then have "\ \ closure T" by (simp add: closure_def) then have "f \ = 0" by (auto simp: continuous_constant_on_closure [OF contf]) show ?thesis apply (rule ccontr) apply (rule isolated_zeros [OF holf \open S\ \connected S\ \\ \ S\ \f \ = 0\ \w \ S\], assumption) by (metis open_ball \\ islimpt T\ centre_in_ball fT0 insertE insert_Diff islimptE) qed corollary analytic_continuation_open: - assumes "open s" and "open s'" and "s \ {}" and "connected s'" + assumes "open s" and "open s'" and "s \ {}" and "connected s'" and "s \ s'" - assumes "f holomorphic_on s'" and "g holomorphic_on s'" + assumes "f holomorphic_on s'" and "g holomorphic_on s'" and "\z. z \ s \ f z = g z" assumes "z \ s'" shows "f z = g z" proof - from \s \ {}\ obtain \ where "\ \ s" by auto - with \open s\ have \: "\ islimpt s" + with \open s\ have \: "\ islimpt s" by (intro interior_limit_point) (auto simp: interior_open) have "f z - g z = 0" by (rule analytic_continuation[of "\z. f z - g z" s' s \]) (insert assms \\ \ s\ \, auto intro: holomorphic_intros) thus ?thesis by simp qed subsection\Open mapping theorem\ lemma holomorphic_contract_to_zero: assumes contf: "continuous_on (cball \ r) f" and holf: "f holomorphic_on ball \ r" and "0 < r" and norm_less: "\z. norm(\ - z) = r \ norm(f \) < norm(f z)" obtains z where "z \ ball \ r" "f z = 0" proof - { assume fnz: "\w. w \ ball \ r \ f w \ 0" then have "0 < norm (f \)" by (simp add: \0 < r\) have fnz': "\w. w \ cball \ r \ f w \ 0" by (metis norm_less dist_norm fnz less_eq_real_def mem_ball mem_cball norm_not_less_zero norm_zero) have "frontier(cball \ r) \ {}" using \0 < r\ by simp define g where [abs_def]: "g z = inverse (f z)" for z have contg: "continuous_on (cball \ r) g" unfolding g_def using contf continuous_on_inverse fnz' by blast have holg: "g holomorphic_on ball \ r" unfolding g_def using fnz holf holomorphic_on_inverse by blast have "frontier (cball \ r) \ cball \ r" by (simp add: subset_iff) then have contf': "continuous_on (frontier (cball \ r)) f" and contg': "continuous_on (frontier (cball \ r)) g" by (blast intro: contf contg continuous_on_subset)+ have froc: "frontier(cball \ r) \ {}" using \0 < r\ by simp moreover have "continuous_on (frontier (cball \ r)) (norm o f)" using contf' continuous_on_compose continuous_on_norm_id by blast ultimately obtain w where w: "w \ frontier(cball \ r)" and now: "\x. x \ frontier(cball \ r) \ norm (f w) \ norm (f x)" apply (rule bexE [OF continuous_attains_inf [OF compact_frontier [OF compact_cball]]]) apply simp done then have fw: "0 < norm (f w)" by (simp add: fnz') have "continuous_on (frontier (cball \ r)) (norm o g)" using contg' continuous_on_compose continuous_on_norm_id by blast then obtain v where v: "v \ frontier(cball \ r)" and nov: "\x. x \ frontier(cball \ r) \ norm (g v) \ norm (g x)" apply (rule bexE [OF continuous_attains_sup [OF compact_frontier [OF compact_cball] froc]]) apply simp done then have fv: "0 < norm (f v)" by (simp add: fnz') have "norm ((deriv ^^ 0) g \) \ fact 0 * norm (g v) / r ^ 0" by (rule Cauchy_inequality [OF holg contg \0 < r\]) (simp add: dist_norm nov) then have "cmod (g \) \ norm (g v)" by simp with w have wr: "norm (\ - w) = r" and nfw: "norm (f w) \ norm (f \)" apply (simp_all add: dist_norm) by (metis \0 < cmod (f \)\ g_def less_imp_inverse_less norm_inverse not_le now order_trans v) with fw have False using norm_less by force } with that show ?thesis by blast qed theorem open_mapping_thm: assumes holf: "f holomorphic_on S" and S: "open S" and "connected S" and "open U" and "U \ S" and fne: "\ f constant_on S" shows "open (f ` U)" proof - have *: "open (f ` U)" if "U \ {}" and U: "open U" "connected U" and "f holomorphic_on U" and fneU: "\x. \y \ U. f y \ x" for U proof (clarsimp simp: open_contains_ball) fix \ assume \: "\ \ U" show "\e>0. ball (f \) e \ f ` U" proof - have hol: "(\z. f z - f \) holomorphic_on U" by (rule holomorphic_intros that)+ obtain s where "0 < s" and sbU: "ball \ s \ U" and sne: "\z. z \ ball \ s - {\} \ (\z. f z - f \) z \ 0" using isolated_zeros [OF hol U \] by (metis fneU right_minus_eq) obtain r where "0 < r" and r: "cball \ r \ ball \ s" apply (rule_tac r="s/2" in that) using \0 < s\ by auto have "cball \ r \ U" using sbU r by blast then have frsbU: "frontier (cball \ r) \ U" using Diff_subset frontier_def order_trans by fastforce then have cof: "compact (frontier(cball \ r))" by blast have frne: "frontier (cball \ r) \ {}" using \0 < r\ by auto have contfr: "continuous_on (frontier (cball \ r)) (\z. norm (f z - f \))" apply (rule continuous_on_compose2 [OF Complex_Analysis_Basics.continuous_on_norm_id]) using hol frsbU holomorphic_on_imp_continuous_on holomorphic_on_subset by blast+ obtain w where "norm (\ - w) = r" and w: "(\z. norm (\ - z) = r \ norm (f w - f \) \ norm(f z - f \))" apply (rule bexE [OF continuous_attains_inf [OF cof frne contfr]]) apply (simp add: dist_norm) done moreover define \ where "\ \ norm (f w - f \) / 3" ultimately have "0 < \" using \0 < r\ dist_complex_def r sne by auto have "ball (f \) \ \ f ` U" proof fix \ assume \: "\ \ ball (f \) \" have *: "cmod (\ - f \) < cmod (\ - f z)" if "cmod (\ - z) = r" for z proof - have lt: "cmod (f w - f \) / 3 < cmod (\ - f z)" using w [OF that] \ using dist_triangle2 [of "f \" "\" "f z"] dist_triangle2 [of "f \" "f z" \] by (simp add: \_def dist_norm norm_minus_commute) show ?thesis by (metis \_def dist_commute dist_norm less_trans lt mem_ball \) qed have "continuous_on (cball \ r) (\z. \ - f z)" apply (rule continuous_intros)+ using \cball \ r \ U\ \f holomorphic_on U\ apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) done moreover have "(\z. \ - f z) holomorphic_on ball \ r" apply (rule holomorphic_intros)+ apply (metis \cball \ r \ U\ \f holomorphic_on U\ holomorphic_on_subset interior_cball interior_subset) done ultimately obtain z where "z \ ball \ r" "\ - f z = 0" apply (rule holomorphic_contract_to_zero) apply (blast intro!: \0 < r\ *)+ done then show "\ \ f ` U" using \cball \ r \ U\ by fastforce qed then show ?thesis using \0 < \\ by blast qed qed have "open (f ` X)" if "X \ components U" for X proof - have holfU: "f holomorphic_on U" using \U \ S\ holf holomorphic_on_subset by blast have "X \ {}" using that by (simp add: in_components_nonempty) moreover have "open X" using that \open U\ open_components by auto moreover have "connected X" using that in_components_maximal by blast moreover have "f holomorphic_on X" by (meson that holfU holomorphic_on_subset in_components_maximal) moreover have "\y\X. f y \ x" for x proof (rule ccontr) assume not: "\ (\y\X. f y \ x)" have "X \ S" using \U \ S\ in_components_subset that by blast obtain w where w: "w \ X" using \X \ {}\ by blast have wis: "w islimpt X" using w \open X\ interior_eq by auto have hol: "(\z. f z - x) holomorphic_on S" by (simp add: holf holomorphic_on_diff) - with fne [unfolded constant_on_def] + with fne [unfolded constant_on_def] analytic_continuation[OF hol S \connected S\ \X \ S\ _ wis] not \X \ S\ w show False by auto qed ultimately show ?thesis by (rule *) qed then have "open (f ` \(components U))" by (metis (no_types, lifting) imageE image_Union open_Union) then show ?thesis by force qed text\No need for \<^term>\S\ to be connected. But the nonconstant condition is stronger.\ corollary\<^marker>\tag unimportant\ open_mapping_thm2: assumes holf: "f holomorphic_on S" and S: "open S" and "open U" "U \ S" and fnc: "\X. \open X; X \ S; X \ {}\ \ \ f constant_on X" shows "open (f ` U)" proof - have "S = \(components S)" by simp with \U \ S\ have "U = (\C \ components S. C \ U)" by auto then have "f ` U = (\C \ components S. f ` (C \ U))" using image_UN by fastforce moreover { fix C assume "C \ components S" with S \C \ components S\ open_components in_components_connected have C: "open C" "connected C" by auto have "C \ S" by (metis \C \ components S\ in_components_maximal) have nf: "\ f constant_on C" apply (rule fnc) using C \C \ S\ \C \ components S\ in_components_nonempty by auto have "f holomorphic_on C" by (metis holf holomorphic_on_subset \C \ S\) then have "open (f ` (C \ U))" apply (rule open_mapping_thm [OF _ C _ _ nf]) apply (simp add: C \open U\ open_Int, blast) done } ultimately show ?thesis by force qed corollary\<^marker>\tag unimportant\ open_mapping_thm3: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" shows "open (f ` S)" apply (rule open_mapping_thm2 [OF holf]) using assms apply (simp_all add:) using injective_not_constant subset_inj_on by blast subsection\Maximum modulus principle\ text\If \<^term>\f\ is holomorphic, then its norm (modulus) cannot exhibit a true local maximum that is properly within the domain of \<^term>\f\.\ proposition maximum_modulus_principle: assumes holf: "f holomorphic_on S" and S: "open S" and "connected S" and "open U" and "U \ S" and "\ \ U" and no: "\z. z \ U \ norm(f z) \ norm(f \)" shows "f constant_on S" proof (rule ccontr) assume "\ f constant_on S" then have "open (f ` U)" using open_mapping_thm assms by blast moreover have "\ open (f ` U)" proof - have "\t. cmod (f \ - t) < e \ t \ f ` U" if "0 < e" for e apply (rule_tac x="if 0 < Re(f \) then f \ + (e/2) else f \ - (e/2)" in exI) using that apply (simp add: dist_norm) apply (fastforce simp: cmod_Re_le_iff dest!: no dest: sym) done then show ?thesis unfolding open_contains_ball by (metis \\ \ U\ contra_subsetD dist_norm imageI mem_ball) qed ultimately show False by blast qed proposition maximum_modulus_frontier: assumes holf: "f holomorphic_on (interior S)" and contf: "continuous_on (closure S) f" and bos: "bounded S" and leB: "\z. z \ frontier S \ norm(f z) \ B" and "\ \ S" shows "norm(f \) \ B" proof - have "compact (closure S)" using bos by (simp add: bounded_closure compact_eq_bounded_closed) moreover have "continuous_on (closure S) (cmod \ f)" using contf continuous_on_compose continuous_on_norm_id by blast ultimately obtain z where zin: "z \ closure S" and z: "\y. y \ closure S \ (cmod \ f) y \ (cmod \ f) z" using continuous_attains_sup [of "closure S" "norm o f"] \\ \ S\ by auto then consider "z \ frontier S" | "z \ interior S" using frontier_def by auto then have "norm(f z) \ B" proof cases case 1 then show ?thesis using leB by blast next case 2 have zin: "z \ connected_component_set (interior S) z" by (simp add: 2) have "f constant_on (connected_component_set (interior S) z)" apply (rule maximum_modulus_principle [OF _ _ _ _ _ zin]) apply (metis connected_component_subset holf holomorphic_on_subset) apply (simp_all add: open_connected_component) by (metis closure_subset comp_eq_dest_lhs interior_subset subsetCE z connected_component_in) then obtain c where c: "\w. w \ connected_component_set (interior S) z \ f w = c" by (auto simp: constant_on_def) have "f ` closure(connected_component_set (interior S) z) \ {c}" apply (rule image_closure_subset) apply (meson closure_mono connected_component_subset contf continuous_on_subset interior_subset) using c apply auto done then have cc: "\w. w \ closure(connected_component_set (interior S) z) \ f w = c" by blast have "frontier(connected_component_set (interior S) z) \ {}" apply (simp add: frontier_eq_empty) by (metis "2" bos bounded_interior connected_component_eq_UNIV connected_component_refl not_bounded_UNIV) then obtain w where w: "w \ frontier(connected_component_set (interior S) z)" by auto then have "norm (f z) = norm (f w)" by (simp add: "2" c cc frontier_def) also have "... \ B" apply (rule leB) using w using frontier_interior_subset frontier_of_connected_component_subset by blast finally show ?thesis . qed then show ?thesis using z \\ \ S\ closure_subset by fastforce qed corollary\<^marker>\tag unimportant\ maximum_real_frontier: assumes holf: "f holomorphic_on (interior S)" and contf: "continuous_on (closure S) f" and bos: "bounded S" and leB: "\z. z \ frontier S \ Re(f z) \ B" and "\ \ S" shows "Re(f \) \ B" using maximum_modulus_frontier [of "exp o f" S "exp B"] Transcendental.continuous_on_exp holomorphic_on_compose holomorphic_on_exp assms by auto subsection\<^marker>\tag unimportant\ \Factoring out a zero according to its order\ lemma holomorphic_factor_order_of_zero: assumes holf: "f holomorphic_on S" and os: "open S" and "\ \ S" "0 < n" and dnz: "(deriv ^^ n) f \ \ 0" and dfz: "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" obtains g r where "0 < r" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" "\w. w \ ball \ r \ g w \ 0" proof - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ r" using holf holomorphic_on_subset by blast define g where "g w = suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" for w have sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" and feq: "f w - f \ = (w - \)^n * g w" if w: "w \ ball \ r" for w proof - define powf where "powf = (\i. (deriv ^^ i) f \/(fact i) * (w - \)^i)" have sing: "{.. = 0 then {} else {0})" unfolding powf_def using \0 < n\ dfz by (auto simp: dfz; metis funpow_0 not_gr0) have "powf sums f w" unfolding powf_def by (rule holomorphic_power_series [OF holfb w]) moreover have "(\i" apply (subst Groups_Big.comm_monoid_add_class.sum.setdiff_irrelevant [symmetric]) apply simp apply (simp only: dfz sing) apply (simp add: powf_def) done ultimately have fsums: "(\i. powf (i+n)) sums (f w - f \)" using w sums_iff_shift' by metis then have *: "summable (\i. (w - \) ^ n * ((deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n)))" unfolding powf_def using sums_summable by (auto simp: power_add mult_ac) have "summable (\i. (deriv ^^ (i + n)) f \ * (w - \) ^ i / fact (i + n))" proof (cases "w=\") case False then show ?thesis using summable_mult [OF *, of "1 / (w - \) ^ n"] by simp next case True then show ?thesis by (auto simp: Power.semiring_1_class.power_0_left intro!: summable_finite [of "{0}"] split: if_split_asm) qed then show sumsg: "(\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i) sums g w" by (simp add: summable_sums_iff g_def) show "f w - f \ = (w - \)^n * g w" apply (rule sums_unique2) apply (rule fsums [unfolded powf_def]) using sums_mult [OF sumsg, of "(w - \) ^ n"] by (auto simp: power_add mult_ac) qed then have holg: "g holomorphic_on ball \ r" by (meson sumsg power_series_holomorphic) then have contg: "continuous_on (ball \ r) g" by (blast intro: holomorphic_on_imp_continuous_on) have "g \ \ 0" using dnz unfolding g_def by (subst suminf_finite [of "{0}"]) auto obtain d where "0 < d" and d: "\w. w \ ball \ d \ g w \ 0" apply (rule exE [OF continuous_on_avoid [OF contg _ \g \ \ 0\]]) using \0 < r\ apply force by (metis \0 < r\ less_trans mem_ball not_less_iff_gr_or_eq) show ?thesis apply (rule that [where g=g and r ="min r d"]) using \0 < r\ \0 < d\ holg apply (auto simp: feq holomorphic_on_subset subset_ball d) done qed lemma holomorphic_factor_order_of_zero_strong: assumes holf: "f holomorphic_on S" "open S" "\ \ S" "0 < n" and "(deriv ^^ n) f \ \ 0" and "\i. \0 < i; i < n\ \ (deriv ^^ i) f \ = 0" obtains g r where "0 < r" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w - f \ = ((w - \) * g w) ^ n" "\w. w \ ball \ r \ g w \ 0" proof - obtain g r where "0 < r" and holg: "g holomorphic_on ball \ r" and feq: "\w. w \ ball \ r \ f w - f \ = (w - \)^n * g w" and gne: "\w. w \ ball \ r \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF assms]) have con: "continuous_on (ball \ r) (\z. deriv g z / g z)" by (rule continuous_intros) (auto simp: gne holg holomorphic_deriv holomorphic_on_imp_continuous_on) have cd: "\x. dist \ x < r \ (\z. deriv g z / g z) field_differentiable at x" apply (rule derivative_intros)+ using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) apply (metis open_ball at_within_open holg holomorphic_on_def mem_ball) using gne mem_ball by blast obtain h where h: "\x. x \ ball \ r \ (h has_field_derivative deriv g x / g x) (at x)" apply (rule exE [OF holomorphic_convex_primitive [of "ball \ r" "{}" "\z. deriv g z / g z"]]) apply (auto simp: con cd) apply (metis open_ball at_within_open mem_ball) done then have "continuous_on (ball \ r) h" by (metis open_ball holomorphic_on_imp_continuous_on holomorphic_on_open) then have con: "continuous_on (ball \ r) (\x. exp (h x) / g x)" by (auto intro!: continuous_intros simp add: holg holomorphic_on_imp_continuous_on gne) have 0: "dist \ x < r \ ((\x. exp (h x) / g x) has_field_derivative 0) (at x)" for x apply (rule h derivative_eq_intros | simp)+ apply (rule DERIV_deriv_iff_field_differentiable [THEN iffD2]) using holg apply (auto simp: holomorphic_on_imp_differentiable_at gne h) done obtain c where c: "\x. x \ ball \ r \ exp (h x) / g x = c" by (rule DERIV_zero_connected_constant [of "ball \ r" "{}" "\x. exp(h x) / g x"]) (auto simp: con 0) have hol: "(\z. exp ((Ln (inverse c) + h z) / of_nat n)) holomorphic_on ball \ r" apply (rule holomorphic_on_compose [unfolded o_def, where g = exp]) apply (rule holomorphic_intros)+ using h holomorphic_on_open apply blast apply (rule holomorphic_intros)+ using \0 < n\ apply simp apply (rule holomorphic_intros)+ done show ?thesis apply (rule that [where g="\z. exp((Ln(inverse c) + h z)/n)" and r =r]) using \0 < r\ \0 < n\ apply (auto simp: feq power_mult_distrib exp_divide_power_eq c [symmetric]) apply (rule hol) apply (simp add: Transcendental.exp_add gne) done qed lemma fixes k :: "'a::wellorder" assumes a_def: "a == LEAST x. P x" and P: "P k" shows def_LeastI: "P a" and def_Least_le: "a \ k" unfolding a_def by (rule LeastI Least_le; rule P)+ lemma holomorphic_factor_zero_nonconstant: assumes holf: "f holomorphic_on S" and S: "open S" "connected S" and "\ \ S" "f \ = 0" and nonconst: "\ f constant_on S" obtains g r n where "0 < n" "0 < r" "ball \ r \ S" "g holomorphic_on ball \ r" "\w. w \ ball \ r \ f w = (w - \)^n * g w" "\w. w \ ball \ r \ g w \ 0" proof (cases "\n>0. (deriv ^^ n) f \ = 0") case True then show ?thesis using holomorphic_fun_eq_const_on_connected [OF holf S _ \\ \ S\] nonconst by (simp add: constant_on_def) next case False then obtain n0 where "n0 > 0" and n0: "(deriv ^^ n0) f \ \ 0" by blast obtain r0 where "r0 > 0" "ball \ r0 \ S" using S openE \\ \ S\ by auto define n where "n \ LEAST n. (deriv ^^ n) f \ \ 0" have n_ne: "(deriv ^^ n) f \ \ 0" by (rule def_LeastI [OF n_def]) (rule n0) then have "0 < n" using \f \ = 0\ using funpow_0 by fastforce have n_min: "\k. k < n \ (deriv ^^ k) f \ = 0" using def_Least_le [OF n_def] not_le by blast then obtain g r1 where "0 < r1" "g holomorphic_on ball \ r1" "\w. w \ ball \ r1 \ f w = (w - \) ^ n * g w" "\w. w \ ball \ r1 \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne] simp: \f \ = 0\) then show ?thesis apply (rule_tac g=g and r="min r0 r1" and n=n in that) using \0 < n\ \0 < r0\ \0 < r1\ \ball \ r0 \ S\ apply (auto simp: subset_ball intro: holomorphic_on_subset) done qed lemma holomorphic_lower_bound_difference: assumes holf: "f holomorphic_on S" and S: "open S" "connected S" and "\ \ S" and "\ \ S" and fne: "f \ \ f \" obtains k n r where "0 < k" "0 < r" "ball \ r \ S" "\w. w \ ball \ r \ k * norm(w - \)^n \ norm(f w - f \)" proof - define n where "n = (LEAST n. 0 < n \ (deriv ^^ n) f \ \ 0)" obtain n0 where "0 < n0" and n0: "(deriv ^^ n0) f \ \ 0" using fne holomorphic_fun_eq_const_on_connected [OF holf S] \\ \ S\ \\ \ S\ by blast then have "0 < n" and n_ne: "(deriv ^^ n) f \ \ 0" unfolding n_def by (metis (mono_tags, lifting) LeastI)+ have n_min: "\k. \0 < k; k < n\ \ (deriv ^^ k) f \ = 0" unfolding n_def by (blast dest: not_less_Least) then obtain g r where "0 < r" and holg: "g holomorphic_on ball \ r" and fne: "\w. w \ ball \ r \ f w - f \ = (w - \) ^ n * g w" and gnz: "\w. w \ ball \ r \ g w \ 0" by (auto intro: holomorphic_factor_order_of_zero [OF holf \open S\ \\ \ S\ \n > 0\ n_ne]) obtain e where "e>0" and e: "ball \ e \ S" using assms by (blast elim!: openE) then have holfb: "f holomorphic_on ball \ e" using holf holomorphic_on_subset by blast define d where "d = (min e r) / 2" have "0 < d" using \0 < r\ \0 < e\ by (simp add: d_def) have "d < r" using \0 < r\ by (auto simp: d_def) then have cbb: "cball \ d \ ball \ r" by (auto simp: cball_subset_ball_iff) then have "g holomorphic_on cball \ d" by (rule holomorphic_on_subset [OF holg]) then have "closed (g ` cball \ d)" by (simp add: compact_imp_closed compact_continuous_image holomorphic_on_imp_continuous_on) moreover have "g ` cball \ d \ {}" using \0 < d\ by auto ultimately obtain x where x: "x \ g ` cball \ d" and "\y. y \ g ` cball \ d \ dist 0 x \ dist 0 y" by (rule distance_attains_inf) blast then have leg: "\w. w \ cball \ d \ norm x \ norm (g w)" by auto have "ball \ d \ cball \ d" by auto also have "... \ ball \ e" using \0 < d\ d_def by auto also have "... \ S" by (rule e) finally have dS: "ball \ d \ S" . moreover have "x \ 0" using gnz x \d < r\ by auto ultimately show ?thesis apply (rule_tac k="norm x" and n=n and r=d in that) using \d < r\ leg apply (auto simp: \0 < d\ fne norm_mult norm_power algebra_simps mult_right_mono) done qed lemma assumes holf: "f holomorphic_on (S - {\})" and \: "\ \ interior S" shows holomorphic_on_extend_lim: "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ ((\z. (z - \) * f z) \ 0) (at \)" (is "?P = ?Q") and holomorphic_on_extend_bounded: "(\g. g holomorphic_on S \ (\z \ S - {\}. g z = f z)) \ (\B. eventually (\z. norm(f z) \ B) (at \))" (is "?P = ?R") proof - obtain \ where "0 < \" and \: "ball \ \ \ S" using \ mem_interior by blast have "?R" if holg: "g holomorphic_on S" and gf: "\z. z \ S - {\} \ g z = f z" for g proof - have *: "\\<^sub>F z in at \. dist (g z) (g \) < 1 \ cmod (f z) \ cmod (g \) + 1" apply (simp add: eventually_at) apply (rule_tac x="\" in exI) using \ \0 < \\ apply (clarsimp simp:) apply (drule_tac c=x in subsetD) apply (simp add: dist_commute) by (metis DiffI add.commute diff_le_eq dist_norm gf le_less_trans less_eq_real_def norm_triangle_ineq2 singletonD) have "continuous_on (interior S) g" by (meson continuous_on_subset holg holomorphic_on_imp_continuous_on interior_subset) then have "\x. x \ interior S \ (g \ g x) (at x)" using continuous_on_interior continuous_within holg holomorphic_on_imp_continuous_on by blast then have "(g \ g \) (at \)" by (simp add: \) then show ?thesis apply (rule_tac x="norm(g \) + 1" in exI) apply (rule eventually_mp [OF * tendstoD [where e=1]], auto) done qed moreover have "?Q" if "\\<^sub>F z in at \. cmod (f z) \ B" for B by (rule lim_null_mult_right_bounded [OF _ that]) (simp add: LIM_zero) moreover have "?P" if "(\z. (z - \) * f z) \\\ 0" proof - define h where [abs_def]: "h z = (z - \)^2 * f z" for z have h0: "(h has_field_derivative 0) (at \)" apply (simp add: h_def has_field_derivative_iff) apply (rule Lim_transform_within [OF that, of 1]) apply (auto simp: field_split_simps power2_eq_square) done have holh: "h holomorphic_on S" proof (simp add: holomorphic_on_def, clarify) fix z assume "z \ S" show "h field_differentiable at z within S" proof (cases "z = \") case True then show ?thesis using field_differentiable_at_within field_differentiable_def h0 by blast next case False then have "f field_differentiable at z within S" using holomorphic_onD [OF holf, of z] \z \ S\ unfolding field_differentiable_def has_field_derivative_iff by (force intro: exI [where x="dist \ z"] elim: Lim_transform_within_set [unfolded eventually_at]) then show ?thesis by (simp add: h_def power2_eq_square derivative_intros) qed qed define g where [abs_def]: "g z = (if z = \ then deriv h \ else (h z - h \) / (z - \))" for z have holg: "g holomorphic_on S" unfolding g_def by (rule pole_lemma [OF holh \]) show ?thesis apply (rule_tac x="\z. if z = \ then deriv g \ else (g z - g \)/(z - \)" in exI) apply (rule conjI) apply (rule pole_lemma [OF holg \]) apply (auto simp: g_def power2_eq_square divide_simps) using h0 apply (simp add: h0 DERIV_imp_deriv h_def power2_eq_square) done qed ultimately show "?P = ?Q" and "?P = ?R" by meson+ qed lemma pole_at_infinity: assumes holf: "f holomorphic_on UNIV" and lim: "((inverse o f) \ l) at_infinity" obtains a n where "\z. f z = (\i\n. a i * z^i)" proof (cases "l = 0") case False with tendsto_inverse [OF lim] show ?thesis apply (rule_tac a="(\n. inverse l)" and n=0 in that) apply (simp add: Liouville_weak [OF holf, of "inverse l"]) done next case True then have [simp]: "l = 0" . show ?thesis proof (cases "\r. 0 < r \ (\z \ ball 0 r - {0}. f(inverse z) \ 0)") case True then obtain r where "0 < r" and r: "\z. z \ ball 0 r - {0} \ f(inverse z) \ 0" by auto have 1: "inverse \ f \ inverse holomorphic_on ball 0 r - {0}" by (rule holomorphic_on_compose holomorphic_intros holomorphic_on_subset [OF holf] | force simp: r)+ have 2: "0 \ interior (ball 0 r)" using \0 < r\ by simp have "\B. 0 eventually (\z. cmod ((inverse \ f \ inverse) z) \ B) (at 0)" apply (rule exI [where x=1]) apply simp using tendstoD [OF lim [unfolded lim_at_infinity_0] zero_less_one] apply (rule eventually_mono) apply (simp add: dist_norm) done with holomorphic_on_extend_bounded [OF 1 2] obtain g where holg: "g holomorphic_on ball 0 r" and geq: "\z. z \ ball 0 r - {0} \ g z = (inverse \ f \ inverse) z" by meson have ifi0: "(inverse \ f \ inverse) \0\ 0" using \l = 0\ lim lim_at_infinity_0 by blast have g2g0: "g \0\ g 0" using \0 < r\ centre_in_ball continuous_at continuous_on_eq_continuous_at holg by (blast intro: holomorphic_on_imp_continuous_on) have g2g1: "g \0\ 0" apply (rule Lim_transform_within_open [OF ifi0 open_ball [of 0 r]]) using \0 < r\ by (auto simp: geq) have [simp]: "g 0 = 0" by (rule tendsto_unique [OF _ g2g0 g2g1]) simp have "ball 0 r - {0::complex} \ {}" using \0 < r\ apply (clarsimp simp: ball_def dist_norm) apply (drule_tac c="of_real r/2" in subsetD, auto) done then obtain w::complex where "w \ 0" and w: "norm w < r" by force then have "g w \ 0" by (simp add: geq r) obtain B n e where "0 < B" "0 < e" "e \ r" and leg: "\w. norm w < e \ B * cmod w ^ n \ cmod (g w)" apply (rule holomorphic_lower_bound_difference [OF holg open_ball connected_ball, of 0 w]) using \0 < r\ w \g w \ 0\ by (auto simp: ball_subset_ball_iff) have "cmod (f z) \ cmod z ^ n / B" if "2/e \ cmod z" for z proof - have ize: "inverse z \ ball 0 e - {0}" using that \0 < e\ by (auto simp: norm_divide field_split_simps algebra_simps) then have [simp]: "z \ 0" and izr: "inverse z \ ball 0 r - {0}" using \e \ r\ by auto then have [simp]: "f z \ 0" using r [of "inverse z"] by simp have [simp]: "f z = inverse (g (inverse z))" using izr geq [of "inverse z"] by simp show ?thesis using ize leg [of "inverse z"] \0 < B\ \0 < e\ by (simp add: field_split_simps norm_divide algebra_simps) qed then show ?thesis apply (rule_tac a = "\k. (deriv ^^ k) f 0 / (fact k)" and n=n in that) apply (rule_tac A = "2/e" and B = "1/B" in Liouville_polynomial [OF holf], simp) done next case False then have fi0: "\r. r > 0 \ \z\ball 0 r - {0}. f (inverse z) = 0" by simp have fz0: "f z = 0" if "0 < r" and lt1: "\x. x \ 0 \ cmod x < r \ inverse (cmod (f (inverse x))) < 1" for z r proof - have f0: "(f \ 0) at_infinity" proof - have DIM_complex[intro]: "2 \ DIM(complex)" \ \should not be necessary!\ by simp from lt1 have "f (inverse x) \ 0 \ x \ 0 \ cmod x < r \ 1 < cmod (f (inverse x))" for x using one_less_inverse by force then have **: "cmod (f (inverse x)) \ 1 \ x \ 0 \ cmod x < r \ f (inverse x) = 0" for x by force then have *: "(f \ inverse) ` (ball 0 r - {0}) \ {0} \ - ball 0 1" by force have "continuous_on (inverse ` (ball 0 r - {0})) f" using continuous_on_subset holf holomorphic_on_imp_continuous_on by blast then have "connected ((f \ inverse) ` (ball 0 r - {0}))" apply (intro connected_continuous_image continuous_intros) apply (force intro: connected_punctured_ball)+ done then have "{0} \ (f \ inverse) ` (ball 0 r - {0}) = {} \ - ball 0 1 \ (f \ inverse) ` (ball 0 r - {0}) = {}" by (rule connected_closedD) (use * in auto) then have "w \ 0 \ cmod w < r \ f (inverse w) = 0" for w using fi0 **[of w] \0 < r\ apply (auto simp add: inf.commute [of "- ball 0 1"] Diff_eq [symmetric] image_subset_iff dest: less_imp_le) apply fastforce apply (drule bspec [of _ _ w]) apply (auto dest: less_imp_le) done then show ?thesis apply (simp add: lim_at_infinity_0) apply (rule tendsto_eventually) apply (simp add: eventually_at) apply (rule_tac x=r in exI) apply (simp add: \0 < r\ dist_norm) done qed obtain w where "w \ ball 0 r - {0}" and "f (inverse w) = 0" using False \0 < r\ by blast then show ?thesis by (auto simp: f0 Liouville_weak [OF holf, of 0]) qed show ?thesis apply (rule that [of "\n. 0" 0]) using lim [unfolded lim_at_infinity_0] apply (simp add: Lim_at dist_norm norm_inverse) apply (drule_tac x=1 in spec) using fz0 apply auto done qed qed subsection\<^marker>\tag unimportant\ \Entire proper functions are precisely the non-trivial polynomials\ lemma proper_map_polyfun: fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" assumes "closed S" and "compact K" and c: "c i \ 0" "1 \ i" "i \ n" shows "compact (S \ {z. (\i\n. c i * z^i) \ K})" proof - obtain B where "B > 0" and B: "\x. x \ K \ norm x \ B" by (metis compact_imp_bounded \compact K\ bounded_pos) have *: "norm x \ b" if "\x. b \ norm x \ B + 1 \ norm (\i\n. c i * x ^ i)" "(\i\n. c i * x ^ i) \ K" for b x proof - have "norm (\i\n. c i * x ^ i) \ B" using B that by blast moreover have "\ B + 1 \ B" by simp ultimately show "norm x \ b" using that by (metis (no_types) less_eq_real_def not_less order_trans) qed have "bounded {z. (\i\n. c i * z ^ i) \ K}" using polyfun_extremal [where c=c and B="B+1", OF c] by (auto simp: bounded_pos eventually_at_infinity_pos *) moreover have "closed ((\z. (\i\n. c i * z ^ i)) -` K)" apply (intro allI continuous_closed_vimage continuous_intros) using \compact K\ compact_eq_bounded_closed by blast ultimately show ?thesis using closed_Int_compact [OF \closed S\] compact_eq_bounded_closed by (auto simp add: vimage_def) qed lemma proper_map_polyfun_univ: fixes c :: "nat \ 'a::{real_normed_div_algebra,heine_borel}" assumes "compact K" "c i \ 0" "1 \ i" "i \ n" shows "compact ({z. (\i\n. c i * z^i) \ K})" using proper_map_polyfun [of UNIV K c i n] assms by simp lemma proper_map_polyfun_eq: assumes "f holomorphic_on UNIV" shows "(\k. compact k \ compact {z. f z \ k}) \ (\c n. 0 < n \ (c n \ 0) \ f = (\z. \i\n. c i * z^i))" (is "?lhs = ?rhs") proof assume compf [rule_format]: ?lhs have 2: "\k. 0 < k \ a k \ 0 \ f = (\z. \i \ k. a i * z ^ i)" if "\z. f z = (\i\n. a i * z ^ i)" for a n proof (cases "\i\n. 0 a i = 0") case True then have [simp]: "\z. f z = a 0" by (simp add: that sum.atMost_shift) have False using compf [of "{a 0}"] by simp then show ?thesis .. next case False then obtain k where k: "0 < k" "k\n" "a k \ 0" by force define m where "m = (GREATEST k. k\n \ a k \ 0)" have m: "m\n \ a m \ 0" unfolding m_def apply (rule GreatestI_nat [where b = n]) using k apply auto done have [simp]: "a i = 0" if "m < i" "i \ n" for i using Greatest_le_nat [where b = "n" and P = "\k. k\n \ a k \ 0"] using m_def not_le that by auto have "k \ m" unfolding m_def apply (rule Greatest_le_nat [where b = "n"]) using k apply auto done with k m show ?thesis by (rule_tac x=m in exI) (auto simp: that comm_monoid_add_class.sum.mono_neutral_right) qed have "((inverse \ f) \ 0) at_infinity" proof (rule Lim_at_infinityI) fix e::real assume "0 < e" with compf [of "cball 0 (inverse e)"] show "\B. \x. B \ cmod x \ dist ((inverse \ f) x) 0 \ e" apply simp apply (clarsimp simp add: compact_eq_bounded_closed bounded_pos norm_inverse) apply (rule_tac x="b+1" in exI) apply (metis inverse_inverse_eq less_add_same_cancel2 less_imp_inverse_less add.commute not_le not_less_iff_gr_or_eq order_trans zero_less_one) done qed then show ?rhs apply (rule pole_at_infinity [OF assms]) using 2 apply blast done next assume ?rhs then obtain c n where "0 < n" "c n \ 0" "f = (\z. \i\n. c i * z ^ i)" by blast then have "compact {z. f z \ k}" if "compact k" for k by (auto intro: proper_map_polyfun_univ [OF that]) then show ?lhs by blast qed subsection \Relating invertibility and nonvanishing of derivative\ lemma has_complex_derivative_locally_injective: assumes holf: "f holomorphic_on S" and S: "\ \ S" "open S" and dnz: "deriv f \ \ 0" obtains r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" proof - have *: "\d>0. \x. dist \ x < d \ onorm (\v. deriv f x * v - deriv f \ * v) < e" if "e > 0" for e proof - have contdf: "continuous_on S (deriv f)" by (simp add: holf holomorphic_deriv holomorphic_on_imp_continuous_on \open S\) obtain \ where "\>0" and \: "\x. \x \ S; dist x \ \ \\ \ cmod (deriv f x - deriv f \) \ e/2" using continuous_onE [OF contdf \\ \ S\, of "e/2"] \0 < e\ by (metis dist_complex_def half_gt_zero less_imp_le) obtain \ where "\>0" "ball \ \ \ S" by (metis openE [OF \open S\ \\ \ S\]) with \\>0\ have "\\>0. \x. dist \ x < \ \ onorm (\v. deriv f x * v - deriv f \ * v) \ e/2" apply (rule_tac x="min \ \" in exI) apply (intro conjI allI impI Operator_Norm.onorm_le) apply simp apply (simp only: Rings.ring_class.left_diff_distrib [symmetric] norm_mult) apply (rule mult_right_mono [OF \]) apply (auto simp: dist_commute Rings.ordered_semiring_class.mult_right_mono \) done with \e>0\ show ?thesis by force qed have "inj ((*) (deriv f \))" using dnz by simp then obtain g' where g': "linear g'" "g' \ (*) (deriv f \) = id" using linear_injective_left_inverse [of "(*) (deriv f \)"] by (auto simp: linear_times) show ?thesis apply (rule has_derivative_locally_injective [OF S, where f=f and f' = "\z h. deriv f z * h" and g' = g']) using g' * apply (simp_all add: linear_conv_bounded_linear that) using DERIV_deriv_iff_field_differentiable has_field_derivative_imp_has_derivative holf holomorphic_on_imp_differentiable_at \open S\ apply blast done qed lemma has_complex_derivative_locally_invertible: assumes holf: "f holomorphic_on S" and S: "\ \ S" "open S" and dnz: "deriv f \ \ 0" obtains r where "r > 0" "ball \ r \ S" "open (f ` (ball \ r))" "inj_on f (ball \ r)" proof - obtain r where "r > 0" "ball \ r \ S" "inj_on f (ball \ r)" by (blast intro: that has_complex_derivative_locally_injective [OF assms]) then have \: "\ \ ball \ r" by simp then have nc: "\ f constant_on ball \ r" using \inj_on f (ball \ r)\ injective_not_constant by fastforce have holf': "f holomorphic_on ball \ r" using \ball \ r \ S\ holf holomorphic_on_subset by blast have "open (f ` ball \ r)" apply (rule open_mapping_thm [OF holf']) using nc apply auto done then show ?thesis using \0 < r\ \ball \ r \ S\ \inj_on f (ball \ r)\ that by blast qed lemma holomorphic_injective_imp_regular: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" and "\ \ S" shows "deriv f \ \ 0" proof - obtain r where "r>0" and r: "ball \ r \ S" using assms by (blast elim!: openE) have holf': "f holomorphic_on ball \ r" using \ball \ r \ S\ holf holomorphic_on_subset by blast show ?thesis proof (cases "\n>0. (deriv ^^ n) f \ = 0") case True have fcon: "f w = f \" if "w \ ball \ r" for w apply (rule holomorphic_fun_eq_const_on_connected [OF holf']) using True \0 < r\ that by auto have False using fcon [of "\ + r/2"] \0 < r\ r injf unfolding inj_on_def by (metis \\ \ S\ contra_subsetD dist_commute fcon mem_ball perfect_choose_dist) then show ?thesis .. next case False then obtain n0 where n0: "n0 > 0 \ (deriv ^^ n0) f \ \ 0" by blast define n where [abs_def]: "n = (LEAST n. n > 0 \ (deriv ^^ n) f \ \ 0)" have n_ne: "n > 0" "(deriv ^^ n) f \ \ 0" using def_LeastI [OF n_def n0] by auto have n_min: "\k. 0 < k \ k < n \ (deriv ^^ k) f \ = 0" using def_Least_le [OF n_def] not_le by auto obtain g \ where "0 < \" and holg: "g holomorphic_on ball \ \" and fd: "\w. w \ ball \ \ \ f w - f \ = ((w - \) * g w) ^ n" and gnz: "\w. w \ ball \ \ \ g w \ 0" apply (rule holomorphic_factor_order_of_zero_strong [OF holf \open S\ \\ \ S\ n_ne]) apply (blast intro: n_min)+ done show ?thesis proof (cases "n=1") case True with n_ne show ?thesis by auto next case False have holgw: "(\w. (w - \) * g w) holomorphic_on ball \ (min r \)" apply (rule holomorphic_intros)+ using holg by (simp add: holomorphic_on_subset subset_ball) have gd: "\w. dist \ w < \ \ (g has_field_derivative deriv g w) (at w)" using holg by (simp add: DERIV_deriv_iff_field_differentiable holomorphic_on_def at_within_open_NO_MATCH) have *: "\w. w \ ball \ (min r \) \ ((\w. (w - \) * g w) has_field_derivative ((w - \) * deriv g w + g w)) (at w)" by (rule gd derivative_eq_intros | simp)+ have [simp]: "deriv (\w. (w - \) * g w) \ \ 0" using * [of \] \0 < \\ \0 < r\ by (simp add: DERIV_imp_deriv gnz) obtain T where "\ \ T" "open T" and Tsb: "T \ ball \ (min r \)" and oimT: "open ((\w. (w - \) * g w) ` T)" apply (rule has_complex_derivative_locally_invertible [OF holgw, of \]) using \0 < r\ \0 < \\ apply (simp_all add:) by (meson open_ball centre_in_ball) define U where "U = (\w. (w - \) * g w) ` T" have "open U" by (metis oimT U_def) have "0 \ U" apply (auto simp: U_def) apply (rule image_eqI [where x = \]) apply (auto simp: \\ \ T\) done then obtain \ where "\>0" and \: "cball 0 \ \ U" using \open U\ open_contains_cball by blast then have "\ * exp(2 * of_real pi * \ * (0/n)) \ cball 0 \" "\ * exp(2 * of_real pi * \ * (1/n)) \ cball 0 \" by (auto simp: norm_mult) with \ have "\ * exp(2 * of_real pi * \ * (0/n)) \ U" "\ * exp(2 * of_real pi * \ * (1/n)) \ U" by blast+ then obtain y0 y1 where "y0 \ T" and y0: "(y0 - \) * g y0 = \ * exp(2 * of_real pi * \ * (0/n))" and "y1 \ T" and y1: "(y1 - \) * g y1 = \ * exp(2 * of_real pi * \ * (1/n))" by (auto simp: U_def) then have "y0 \ ball \ \" "y1 \ ball \ \" using Tsb by auto moreover have "y0 \ y1" using y0 y1 \\ > 0\ complex_root_unity_eq_1 [of n 1] \n > 0\ False by auto moreover have "T \ S" by (meson Tsb min.cobounded1 order_trans r subset_ball) ultimately have False using inj_onD [OF injf, of y0 y1] \y0 \ T\ \y1 \ T\ using fd [of y0] fd [of y1] complex_root_unity [of n 1] n_ne apply (simp add: y0 y1 power_mult_distrib) apply (force simp: algebra_simps) done then show ?thesis .. qed qed qed text\Hence a nice clean inverse function theorem\ +lemma has_field_derivative_inverse_strong: + fixes f :: "'a::{euclidean_space,real_normed_field} ⇒ 'a" + shows "DERIV f x :> f' ⟹ + f' ≠ 0 ⟹ + open S ⟹ + x ∈ S ⟹ + continuous_on S f ⟹ + (⋀z. z ∈ S ⟹ g (f z) = z) + ⟹ DERIV g (f x) :> inverse (f')" + unfolding has_field_derivative_def + apply (rule has_derivative_inverse_strong [of S x f g ]) + by auto + +lemma has_field_derivative_inverse_strong_x: + fixes f :: "'a::{euclidean_space,real_normed_field} ⇒ 'a" + shows "DERIV f (g y) :> f' ⟹ + f' ≠ 0 ⟹ + open S ⟹ + continuous_on S f ⟹ + g y ∈ S ⟹ f(g y) = y ⟹ + (⋀z. z ∈ S ⟹ g (f z) = z) + ⟹ DERIV g y :> inverse (f')" + unfolding has_field_derivative_def + apply (rule has_derivative_inverse_strong_x [of S g y f]) + by auto + proposition holomorphic_has_inverse: assumes holf: "f holomorphic_on S" and "open S" and injf: "inj_on f S" obtains g where "g holomorphic_on (f ` S)" "\z. z \ S \ deriv f z * deriv g (f z) = 1" "\z. z \ S \ g(f z) = z" proof - have ofs: "open (f ` S)" by (rule open_mapping_thm3 [OF assms]) have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) have *: "(the_inv_into S f has_field_derivative inverse (deriv f z)) (at (f z))" if "z \ S" for z proof - have 1: "(f has_field_derivative deriv f z) (at z)" using DERIV_deriv_iff_field_differentiable \z \ S\ \open S\ holf holomorphic_on_imp_differentiable_at by blast have 2: "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast show ?thesis apply (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) apply (simp add: holf holomorphic_on_imp_continuous_on) by (simp add: injf the_inv_into_f_f) qed show ?thesis proof show "the_inv_into S f holomorphic_on f ` S" by (simp add: holomorphic_on_open ofs) (blast intro: *) next fix z assume "z \ S" have "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast then show "deriv f z * deriv (the_inv_into S f) (f z) = 1" using * [OF \z \ S\] by (simp add: DERIV_imp_deriv) next fix z assume "z \ S" show "the_inv_into S f (f z) = z" by (simp add: \z \ S\ injf the_inv_into_f_f) qed qed subsection\The Schwarz Lemma\ lemma Schwarz1: assumes holf: "f holomorphic_on S" and contf: "continuous_on (closure S) f" and S: "open S" "connected S" and boS: "bounded S" and "S \ {}" obtains w where "w \ frontier S" "\z. z \ closure S \ norm (f z) \ norm (f w)" proof - have connf: "continuous_on (closure S) (norm o f)" using contf continuous_on_compose continuous_on_norm_id by blast have coc: "compact (closure S)" by (simp add: \bounded S\ bounded_closure compact_eq_bounded_closed) then obtain x where x: "x \ closure S" and xmax: "\z. z \ closure S \ norm(f z) \ norm(f x)" apply (rule bexE [OF continuous_attains_sup [OF _ _ connf]]) using \S \ {}\ apply auto done then show ?thesis proof (cases "x \ frontier S") case True then show ?thesis using that xmax by blast next case False then have "x \ S" using \open S\ frontier_def interior_eq x by auto then have "f constant_on S" apply (rule maximum_modulus_principle [OF holf S \open S\ order_refl]) using closure_subset apply (blast intro: xmax) done then have "f constant_on (closure S)" by (rule constant_on_closureI [OF _ contf]) then obtain c where c: "\x. x \ closure S \ f x = c" by (meson constant_on_def) obtain w where "w \ frontier S" by (metis coc all_not_in_conv assms(6) closure_UNIV frontier_eq_empty not_compact_UNIV) then show ?thesis by (simp add: c frontier_def that) qed qed lemma Schwarz2: "\f holomorphic_on ball 0 r; 0 < s; ball w s \ ball 0 r; \z. norm (w-z) < s \ norm(f z) \ norm(f w)\ \ f constant_on ball 0 r" by (rule maximum_modulus_principle [where U = "ball w s" and \ = w]) (simp_all add: dist_norm) lemma Schwarz3: assumes holf: "f holomorphic_on (ball 0 r)" and [simp]: "f 0 = 0" obtains h where "h holomorphic_on (ball 0 r)" and "\z. norm z < r \ f z = z * (h z)" and "deriv f 0 = h 0" proof - define h where "h z = (if z = 0 then deriv f 0 else f z / z)" for z have d0: "deriv f 0 = h 0" by (simp add: h_def) moreover have "h holomorphic_on (ball 0 r)" by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) moreover have "norm z < r \ f z = z * h z" for z by (simp add: h_def) ultimately show ?thesis using that by blast qed proposition Schwarz_Lemma: assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" and \: "norm \ < 1" shows "norm (f \) \ norm \" and "norm(deriv f 0) \ 1" - and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) + and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) - \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" + \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" (is "?P \ ?Q") proof - obtain h where holh: "h holomorphic_on (ball 0 1)" and fz_eq: "\z. norm z < 1 \ f z = z * (h z)" and df0: "deriv f 0 = h 0" by (rule Schwarz3 [OF holf]) auto have noh_le: "norm (h z) \ 1" if z: "norm z < 1" for z proof - have "norm (h z) < a" if a: "1 < a" for a proof - have "max (inverse a) (norm z) < 1" using z a by (simp_all add: inverse_less_1_iff) then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" using Rats_dense_in_real by blast then have nzr: "norm z < r" and ira: "inverse r < a" using z a less_imp_inverse_less by force+ then have "0 < r" by (meson norm_not_less_zero not_le order.strict_trans2) have holh': "h holomorphic_on ball 0 r" by (meson holh \r < 1\ holomorphic_on_subset less_eq_real_def subset_ball) have conth': "continuous_on (cball 0 r) h" by (meson \r < 1\ dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) obtain w where w: "norm w = r" and lenw: "\z. norm z < r \ norm(h z) \ norm(h w)" apply (rule Schwarz1 [OF holh']) using conth' \0 < r\ by auto have "h w = f w / w" using fz_eq \r < 1\ nzr w by auto then have "cmod (h z) < inverse r" by (metis \0 < r\ \r < 1\ divide_strict_right_mono inverse_eq_divide le_less_trans lenw no norm_divide nzr w) then show ?thesis using ira by linarith qed then show "norm (h z) \ 1" using not_le by blast qed show "cmod (f \) \ cmod \" proof (cases "\ = 0") case True then show ?thesis by auto next case False then show ?thesis by (simp add: noh_le fz_eq \ mult_left_le norm_mult) qed show no_df0: "norm(deriv f 0) \ 1" by (simp add: \\z. cmod z < 1 \ cmod (h z) \ 1\ df0) show "?Q" if "?P" using that proof assume "\z. cmod z < 1 \ z \ 0 \ cmod (f z) = cmod z" then obtain \ where \: "cmod \ < 1" "\ \ 0" "cmod (f \) = cmod \" by blast then have [simp]: "norm (h \) = 1" by (simp add: fz_eq norm_mult) have "ball \ (1 - cmod \) \ ball 0 1" by (simp add: ball_subset_ball_iff) moreover have "\z. cmod (\ - z) < 1 - cmod \ \ cmod (h z) \ cmod (h \)" apply (simp add: algebra_simps) by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) ultimately obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh, of "1 - norm \" \, unfolded constant_on_def] \ by auto then have "norm c = 1" using \ by force with c show ?thesis using fz_eq by auto next assume [simp]: "cmod (deriv f 0) = 1" then obtain c where c: "\z. norm z < 1 \ h z = c" using Schwarz2 [OF holh zero_less_one, of 0, unfolded constant_on_def] df0 noh_le by auto moreover have "norm c = 1" using df0 c by auto ultimately show ?thesis using fz_eq by auto qed qed corollary Schwarz_Lemma': assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" and no: "\z. norm z < 1 \ norm (f z) < 1" - shows "((\\. norm \ < 1 \ norm (f \) \ norm \) - \ norm(deriv f 0) \ 1) - \ (((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) + shows "((\\. norm \ < 1 \ norm (f \) \ norm \) + \ norm(deriv f 0) \ 1) + \ (((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) \ (\\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1))" using Schwarz_Lemma [OF assms] by (metis (no_types) norm_eq_zero zero_less_one) subsection\The Schwarz reflection principle\ lemma hol_pal_lem0: assumes "d \ a \ k" "k \ d \ b" obtains c where "c \ closed_segment a b" "d \ c = k" "\z. z \ closed_segment a c \ d \ z \ k" "\z. z \ closed_segment c b \ k \ d \ z" proof - obtain c where cin: "c \ closed_segment a b" and keq: "k = d \ c" using connected_ivt_hyperplane [of "closed_segment a b" a b d k] by (auto simp: assms) have "closed_segment a c \ {z. d \ z \ k}" "closed_segment c b \ {z. k \ d \ z}" unfolding segment_convex_hull using assms keq by (auto simp: convex_halfspace_le convex_halfspace_ge hull_minimal) then show ?thesis using cin that by fastforce qed lemma hol_pal_lem1: assumes "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" "d \ 0" and lek: "d \ a \ k" "d \ b \ k" "d \ c \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof - have "interior (convex hull {a, b, c}) \ interior(S \ {x. d \ x \ k})" apply (rule interior_mono) apply (rule hull_minimal) apply (simp add: abc lek) apply (rule convex_Int [OF \convex S\ convex_halfspace_le]) done also have "... \ {z \ S. d \ z < k}" by (force simp: interior_open [OF \open S\] \d \ 0\) finally have *: "interior (convex hull {a, b, c}) \ {z \ S. d \ z < k}" . have "continuous_on (convex hull {a,b,c}) f" using \convex S\ contf abc continuous_on_subset subset_hull by fastforce moreover have "f holomorphic_on interior (convex hull {a,b,c})" by (rule holomorphic_on_subset [OF holf1 *]) ultimately show ?thesis using Cauchy_theorem_triangle_interior has_chain_integral_chain_integral3 by blast qed lemma hol_pal_lem2: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" "d \ b \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ c \ k") case True show ?thesis by (rule hol_pal_lem1 [OF S abc \d \ 0\ lek True holf1 contf]) next case False then have "d \ c > k" by force obtain a' where a': "a' \ closed_segment b c" and "d \ a' = k" and ba': "\z. z \ closed_segment b a' \ d \ z \ k" and a'c: "\z. z \ closed_segment a' c \ k \ d \ z" apply (rule hol_pal_lem0 [of d b k c, OF \d \ b \ k\]) using False by auto obtain b' where b': "b' \ closed_segment a c" and "d \ b' = k" and ab': "\z. z \ closed_segment a b' \ d \ z \ k" and b'c: "\z. z \ closed_segment b' c \ k \ d \ z" apply (rule hol_pal_lem0 [of d a k c, OF \d \ a \ k\]) using False by auto have a'b': "a' \ S \ b' \ S" using a' abc b' convex_contains_segment \convex S\ by auto have "continuous_on (closed_segment c a) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 1: "contour_integral (linepath c a) f = contour_integral (linepath c b') f + contour_integral (linepath b' a) f" apply (rule contour_integral_split_linepath) using b' by (simp add: closed_segment_commute) have "continuous_on (closed_segment b c) f" by (meson abc contf continuous_on_subset convex_contains_segment \convex S\) then have 2: "contour_integral (linepath b c) f = contour_integral (linepath b a') f + contour_integral (linepath a' c) f" by (rule contour_integral_split_linepath [OF _ a']) have 3: "contour_integral (reversepath (linepath b' a')) f = - contour_integral (linepath b' a') f" by (rule contour_integral_reversepath [OF valid_path_linepath]) have fcd_le: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. d \ x \ k}" for x proof - have "f holomorphic_on S \ {c. d \ c < k}" by (metis (no_types) Collect_conj_eq Collect_mem_eq holf1) then have "\C D. x \ interior C \ interior D \ f holomorphic_on interior C \ interior D" using that by (metis Collect_mem_eq Int_Collect \d \ 0\ interior_halfspace_le interior_open \open S\) then show "f field_differentiable at x" by (metis at_within_interior holomorphic_on_def interior_Int interior_interior) qed have ab_le: "\x. x \ closed_segment a b \ d \ x \ k" proof - fix x :: complex assume "x \ closed_segment a b" then have "\C. x \ C \ b \ C \ a \ C \ \ convex C" by (meson contra_subsetD convex_contains_segment) then show "d \ x \ k" by (metis lek convex_halfspace_le mem_Collect_eq) qed have "continuous_on (S \ {x. d \ x \ k}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a b +++ linepath b a' +++ linepath a' b' +++ linepath b' a)" apply (rule Cauchy_theorem_convex [where K = "{}"]) apply (simp_all add: path_image_join convex_Int convex_halfspace_le \convex S\ fcd_le ab_le closed_segment_subset abc a'b' ba') by (metis \d \ a' = k\ \d \ b' = k\ convex_contains_segment convex_halfspace_le lek(1) mem_Collect_eq order_refl) then have 4: "contour_integral (linepath a b) f + contour_integral (linepath b a') f + contour_integral (linepath a' b') f + contour_integral (linepath b' a) f = 0" by (rule has_chain_integral_chain_integral4) have fcd_ge: "f field_differentiable at x" if "x \ interior S \ x \ interior {x. k \ d \ x}" for x proof - have f2: "f holomorphic_on S \ {c. k < d \ c}" by (metis (full_types) Collect_conj_eq Collect_mem_eq holf2) have f3: "interior S = S" by (simp add: interior_open \open S\) then have "x \ S \ interior {c. k \ d \ c}" using that by simp then show "f field_differentiable at x" using f3 f2 unfolding holomorphic_on_def by (metis (no_types) \d \ 0\ at_within_interior interior_Int interior_halfspace_ge interior_interior) qed have "continuous_on (S \ {x. k \ d \ x}) f" using contf by (simp add: continuous_on_subset) then have "(f has_contour_integral 0) (linepath a' c +++ linepath c b' +++ linepath b' a')" apply (rule Cauchy_theorem_convex [where K = "{}"]) apply (simp_all add: path_image_join convex_Int convex_halfspace_ge \convex S\ fcd_ge closed_segment_subset abc a'b' a'c) by (metis \d \ a' = k\ b'c closed_segment_commute convex_contains_segment convex_halfspace_ge ends_in_segment(2) mem_Collect_eq order_refl) then have 5: "contour_integral (linepath a' c) f + contour_integral (linepath c b') f + contour_integral (linepath b' a') f = 0" by (rule has_chain_integral_chain_integral3) show ?thesis using 1 2 3 4 5 by (metis add.assoc eq_neg_iff_add_eq_0 reversepath_linepath) qed lemma hol_pal_lem3: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and lek: "d \ a \ k" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ b \ k") case True show ?thesis by (rule hol_pal_lem2 [OF S abc \d \ 0\ lek True holf1 holf2 contf]) next case False show ?thesis proof (cases "d \ c \ k") case True have "contour_integral (linepath c a) f + contour_integral (linepath a b) f + contour_integral (linepath b c) f = 0" by (rule hol_pal_lem2 [OF S \c \ S\ \a \ S\ \b \ S\ \d \ 0\ \d \ c \ k\ lek holf1 holf2 contf]) then show ?thesis by (simp add: algebra_simps) next case False have "contour_integral (linepath b c) f + contour_integral (linepath c a) f + contour_integral (linepath a b) f = 0" apply (rule hol_pal_lem2 [OF S \b \ S\ \c \ S\ \a \ S\, of "-d" "-k"]) using \d \ 0\ \\ d \ b \ k\ False by (simp_all add: holf1 holf2 contf) then show ?thesis by (simp add: algebra_simps) qed qed lemma hol_pal_lem4: assumes S: "convex S" "open S" and abc: "a \ S" "b \ S" "c \ S" and "d \ 0" and holf1: "f holomorphic_on {z. z \ S \ d \ z < k}" and holf2: "f holomorphic_on {z. z \ S \ k < d \ z}" and contf: "continuous_on S f" shows "contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0" proof (cases "d \ a \ k") case True show ?thesis by (rule hol_pal_lem3 [OF S abc \d \ 0\ True holf1 holf2 contf]) next case False show ?thesis apply (rule hol_pal_lem3 [OF S abc, of "-d" "-k"]) using \d \ 0\ False by (simp_all add: holf1 holf2 contf) qed lemma holomorphic_on_paste_across_line: assumes S: "open S" and "d \ 0" and holf1: "f holomorphic_on (S \ {z. d \ z < k})" and holf2: "f holomorphic_on (S \ {z. k < d \ z})" and contf: "continuous_on S f" shows "f holomorphic_on S" proof - have *: "\t. open t \ p \ t \ continuous_on t f \ (\a b c. convex hull {a, b, c} \ t \ contour_integral (linepath a b) f + contour_integral (linepath b c) f + contour_integral (linepath c a) f = 0)" if "p \ S" for p proof - obtain e where "e>0" and e: "ball p e \ S" using \p \ S\ openE S by blast then have "continuous_on (ball p e) f" using contf continuous_on_subset by blast moreover have "f holomorphic_on {z. dist p z < e \ d \ z < k}" apply (rule holomorphic_on_subset [OF holf1]) using e by auto moreover have "f holomorphic_on {z. dist p z < e \ k < d \ z}" apply (rule holomorphic_on_subset [OF holf2]) using e by auto ultimately show ?thesis apply (rule_tac x="ball p e" in exI) using \e > 0\ e \d \ 0\ apply (simp add:, clarify) apply (rule hol_pal_lem4 [of "ball p e" _ _ _ d _ k]) apply (auto simp: subset_hull) done qed show ?thesis by (blast intro: * Morera_local_triangle analytic_imp_holomorphic) qed proposition Schwarz_reflection: assumes "open S" and cnjs: "cnj ` S \ S" and holf: "f holomorphic_on (S \ {z. 0 < Im z})" and contf: "continuous_on (S \ {z. 0 \ Im z}) f" and f: "\z. \z \ S; z \ \\ \ (f z) \ \" shows "(\z. if 0 \ Im z then f z else cnj(f(cnj z))) holomorphic_on S" proof - have 1: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. 0 < Im z})" by (force intro: iffD1 [OF holomorphic_cong [OF refl] holf]) have cont_cfc: "continuous_on (S \ {z. Im z \ 0}) (cnj o f o cnj)" apply (intro continuous_intros continuous_on_compose continuous_on_subset [OF contf]) using cnjs apply auto done have "cnj \ f \ cnj field_differentiable at x within S \ {z. Im z < 0}" if "x \ S" "Im x < 0" "f field_differentiable at (cnj x) within S \ {z. 0 < Im z}" for x using that apply (simp add: field_differentiable_def has_field_derivative_iff Lim_within dist_norm, clarify) apply (rule_tac x="cnj f'" in exI) apply (elim all_forward ex_forward conj_forward imp_forward asm_rl, clarify) apply (drule_tac x="cnj xa" in bspec) using cnjs apply force apply (metis complex_cnj_cnj complex_cnj_diff complex_cnj_divide complex_mod_cnj) done then have hol_cfc: "(cnj o f o cnj) holomorphic_on (S \ {z. Im z < 0})" using holf cnjs by (force simp: holomorphic_on_def) have 2: "(\z. if 0 \ Im z then f z else cnj (f (cnj z))) holomorphic_on (S \ {z. Im z < 0})" apply (rule iffD1 [OF holomorphic_cong [OF refl]]) using hol_cfc by auto have [simp]: "(S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0}) = S" by force have "continuous_on ((S \ {z. 0 \ Im z}) \ (S \ {z. Im z \ 0})) (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" apply (rule continuous_on_cases_local) using cont_cfc contf apply (simp_all add: closedin_closed_Int closed_halfspace_Im_le closed_halfspace_Im_ge) using f Reals_cnj_iff complex_is_Real_iff apply auto done then have 3: "continuous_on S (\z. if 0 \ Im z then f z else cnj (f (cnj z)))" by force show ?thesis apply (rule holomorphic_on_paste_across_line [OF \open S\, of "- \" _ 0]) using 1 2 3 apply auto done qed subsection\Bloch's theorem\ lemma Bloch_lemma_0: assumes holf: "f holomorphic_on cball 0 r" and "0 < r" and [simp]: "f 0 = 0" and le: "\z. norm z < r \ norm(deriv f z) \ 2 * norm(deriv f 0)" shows "ball 0 ((3 - 2 * sqrt 2) * r * norm(deriv f 0)) \ f ` ball 0 r" proof - have "sqrt 2 < 3/2" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp show ?thesis proof (cases "deriv f 0 = 0") case True then show ?thesis by simp next case False define C where "C = 2 * norm(deriv f 0)" have "0 < C" using False by (simp add: C_def) have holf': "f holomorphic_on ball 0 r" using holf using ball_subset_cball holomorphic_on_subset by blast then have holdf': "deriv f holomorphic_on ball 0 r" by (rule holomorphic_deriv [OF _ open_ball]) have "Le1": "norm(deriv f z - deriv f 0) \ norm z / (r - norm z) * C" if "norm z < r" for z proof - have T1: "norm(deriv f z - deriv f 0) \ norm z / (R - norm z) * C" if R: "norm z < R" "R < r" for R proof - have "0 < R" using R by (metis less_trans norm_zero zero_less_norm_iff) have df_le: "\x. norm x < r \ norm (deriv f x) \ C" using le by (simp add: C_def) have hol_df: "deriv f holomorphic_on cball 0 R" apply (rule holomorphic_on_subset) using R holdf' by auto have *: "((\w. deriv f w / (w - z)) has_contour_integral 2 * pi * \ * deriv f z) (circlepath 0 R)" if "norm z < R" for z using \0 < R\ that Cauchy_integral_formula_convex_simple [OF convex_cball hol_df, of _ "circlepath 0 R"] by (force simp: winding_number_circlepath) have **: "((\x. deriv f x / (x - z) - deriv f x / x) has_contour_integral of_real (2 * pi) * \ * (deriv f z - deriv f 0)) (circlepath 0 R)" using has_contour_integral_diff [OF * [of z] * [of 0]] \0 < R\ that by (simp add: algebra_simps) have [simp]: "\x. norm x = R \ x \ z" using that(1) by blast have "norm (deriv f x / (x - z) - deriv f x / x) \ C * norm z / (R * (R - norm z))" if "norm x = R" for x proof - have [simp]: "norm (deriv f x * x - deriv f x * (x - z)) = norm (deriv f x) * norm z" by (simp add: norm_mult right_diff_distrib') show ?thesis using \0 < R\ \0 < C\ R that apply (simp add: norm_mult norm_divide divide_simps) using df_le norm_triangle_ineq2 \0 < C\ apply (auto intro!: mult_mono) done qed then show ?thesis using has_contour_integral_bound_circlepath [OF **, of "C * norm z/(R*(R - norm z))"] \0 < R\ \0 < C\ R apply (simp add: norm_mult norm_divide) apply (simp add: divide_simps mult.commute) done qed obtain r' where r': "norm z < r'" "r' < r" using Rats_dense_in_real [of "norm z" r] \norm z < r\ by blast then have [simp]: "closure {r'<.. norm(f z)" if r: "norm z < r" for z proof - have 1: "\x. x \ ball 0 r \ ((\z. f z - deriv f 0 * z) has_field_derivative deriv f x - deriv f 0) (at x within ball 0 r)" by (rule derivative_eq_intros holomorphic_derivI holf' | simp)+ have 2: "closed_segment 0 z \ ball 0 r" by (metis \0 < r\ convex_ball convex_contains_segment dist_self mem_ball mem_ball_0 that) have 3: "(\t. (norm z)\<^sup>2 * t / (r - norm z) * C) integrable_on {0..1}" apply (rule integrable_on_cmult_right [where 'b=real, simplified]) apply (rule integrable_on_cdivide [where 'b=real, simplified]) apply (rule integrable_on_cmult_left [where 'b=real, simplified]) apply (rule ident_integrable_on) done have 4: "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm z * norm z * x * C / (r - norm z)" if x: "0 \ x" "x \ 1" for x proof - have [simp]: "x * norm z < r" using r x by (meson le_less_trans mult_le_cancel_right2 norm_not_less_zero) have "norm (deriv f (x *\<^sub>R z) - deriv f 0) \ norm (x *\<^sub>R z) / (r - norm (x *\<^sub>R z)) * C" apply (rule Le1) using r x \0 < r\ by simp also have "... \ norm (x *\<^sub>R z) / (r - norm z) * C" using r x \0 < r\ apply (simp add: field_split_simps) by (simp add: \0 < C\ mult.assoc mult_left_le_one_le ordered_comm_semiring_class.comm_mult_left_mono) finally have "norm (deriv f (x *\<^sub>R z) - deriv f 0) * norm z \ norm (x *\<^sub>R z) / (r - norm z) * C * norm z" by (rule mult_right_mono) simp with x show ?thesis by (simp add: algebra_simps) qed have le_norm: "abc \ norm d - e \ norm(f - d) \ e \ abc \ norm f" for abc d e and f::complex by (metis add_diff_cancel_left' add_diff_eq diff_left_mono norm_diff_ineq order_trans) have "norm (integral {0..1} (\x. (deriv f (x *\<^sub>R z) - deriv f 0) * z)) \ integral {0..1} (\t. (norm z)\<^sup>2 * t / (r - norm z) * C)" apply (rule integral_norm_bound_integral) using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 apply (simp add: has_contour_integral_linepath has_integral_integrable_integral) apply (rule 3) apply (simp add: norm_mult power2_eq_square 4) done then have int_le: "norm (f z - deriv f 0 * z) \ (norm z)\<^sup>2 * norm(deriv f 0) / ((r - norm z))" using contour_integral_primitive [OF 1, of "linepath 0 z"] 2 apply (simp add: has_contour_integral_linepath has_integral_integrable_integral C_def) done show ?thesis apply (rule le_norm [OF _ int_le]) using \norm z < r\ apply (simp add: power2_eq_square divide_simps C_def norm_mult) proof - have "norm z * (norm (deriv f 0) * (r - norm z - norm z)) \ norm z * (norm (deriv f 0) * (r - norm z) - norm (deriv f 0) * norm z)" by (simp add: algebra_simps) then show "(norm z * (r - norm z) - norm z * norm z) * norm (deriv f 0) \ norm (deriv f 0) * norm z * (r - norm z) - norm z * norm z * norm (deriv f 0)" by (simp add: algebra_simps) qed qed have sq201 [simp]: "0 < (1 - sqrt 2 / 2)" "(1 - sqrt 2 / 2) < 1" by (auto simp: sqrt2_less_2) have 1: "continuous_on (closure (ball 0 ((1 - sqrt 2 / 2) * r))) f" apply (rule continuous_on_subset [OF holomorphic_on_imp_continuous_on [OF holf]]) apply (subst closure_ball) using \0 < r\ mult_pos_pos sq201 apply (auto simp: cball_subset_cball_iff) done have 2: "open (f ` interior (ball 0 ((1 - sqrt 2 / 2) * r)))" apply (rule open_mapping_thm [OF holf' open_ball connected_ball], force) using \0 < r\ mult_pos_pos sq201 apply (simp add: ball_subset_ball_iff) using False \0 < r\ centre_in_ball holf' holomorphic_nonconstant by blast have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv f 0)) = ball (f 0) ((3 - 2 * sqrt 2) * r * norm (deriv f 0))" by simp also have "... \ f ` ball 0 ((1 - sqrt 2 / 2) * r)" proof - have 3: "(3 - 2 * sqrt 2) * r * norm (deriv f 0) \ norm (f z)" if "norm z = (1 - sqrt 2 / 2) * r" for z apply (rule order_trans [OF _ *]) using \0 < r\ apply (simp_all add: field_simps power2_eq_square that) apply (simp add: mult.assoc [symmetric]) done show ?thesis apply (rule ball_subset_open_map_image [OF 1 2 _ bounded_ball]) using \0 < r\ sq201 3 apply simp_all using C_def \0 < C\ sq3 apply force done qed also have "... \ f ` ball 0 r" apply (rule image_subsetI [OF imageI], simp) apply (erule less_le_trans) using \0 < r\ apply (auto simp: field_simps) done finally show ?thesis . qed qed lemma Bloch_lemma: assumes holf: "f holomorphic_on cball a r" and "0 < r" and le: "\z. z \ ball a r \ norm(deriv f z) \ 2 * norm(deriv f a)" shows "ball (f a) ((3 - 2 * sqrt 2) * r * norm(deriv f a)) \ f ` ball a r" proof - have fz: "(\z. f (a + z)) = f o (\z. (a + z))" by (simp add: o_def) have hol0: "(\z. f (a + z)) holomorphic_on cball 0 r" unfolding fz by (intro holomorphic_intros holf holomorphic_on_compose | simp)+ then have [simp]: "\x. norm x < r \ (\z. f (a + z)) field_differentiable at x" by (metis open_ball at_within_open ball_subset_cball diff_0 dist_norm holomorphic_on_def holomorphic_on_subset mem_ball norm_minus_cancel) have [simp]: "\z. norm z < r \ f field_differentiable at (a + z)" by (metis holf open_ball add_diff_cancel_left' dist_complex_def holomorphic_on_imp_differentiable_at holomorphic_on_subset interior_cball interior_subset mem_ball norm_minus_commute) then have [simp]: "f field_differentiable at a" by (metis add.comm_neutral \0 < r\ norm_eq_zero) have hol1: "(\z. f (a + z) - f a) holomorphic_on cball 0 r" by (intro holomorphic_intros hol0) then have "ball 0 ((3 - 2 * sqrt 2) * r * norm (deriv (\z. f (a + z) - f a) 0)) \ (\z. f (a + z) - f a) ` ball 0 r" apply (rule Bloch_lemma_0) apply (simp_all add: \0 < r\) apply (simp add: fz complex_derivative_chain) apply (simp add: dist_norm le) done then show ?thesis apply clarify apply (drule_tac c="x - f a" in subsetD) apply (force simp: fz \0 < r\ dist_norm complex_derivative_chain field_differentiable_compose)+ done qed proposition Bloch_unit: assumes holf: "f holomorphic_on ball a 1" and [simp]: "deriv f a = 1" obtains b r where "1/12 < r" and "ball b r \ f ` (ball a 1)" proof - define r :: real where "r = 249/256" have "0 < r" "r < 1" by (auto simp: r_def) define g where "g z = deriv f z * of_real(r - norm(z - a))" for z have "deriv f holomorphic_on ball a 1" by (rule holomorphic_deriv [OF holf open_ball]) then have "continuous_on (ball a 1) (deriv f)" using holomorphic_on_imp_continuous_on by blast then have "continuous_on (cball a r) (deriv f)" by (rule continuous_on_subset) (simp add: cball_subset_ball_iff \r < 1\) then have "continuous_on (cball a r) g" by (simp add: g_def continuous_intros) then have 1: "compact (g ` cball a r)" by (rule compact_continuous_image [OF _ compact_cball]) have 2: "g ` cball a r \ {}" using \r > 0\ by auto obtain p where pr: "p \ cball a r" and pge: "\y. y \ cball a r \ norm (g y) \ norm (g p)" using distance_attains_sup [OF 1 2, of 0] by force define t where "t = (r - norm(p - a)) / 2" have "norm (p - a) \ r" using pge [of a] \r > 0\ by (auto simp: g_def norm_mult) then have "norm (p - a) < r" using pr by (simp add: norm_minus_commute dist_norm) then have "0 < t" by (simp add: t_def) have cpt: "cball p t \ ball a r" using \0 < t\ by (simp add: cball_subset_ball_iff dist_norm t_def field_simps) have gen_le_dfp: "norm (deriv f y) * (r - norm (y - a)) / (r - norm (p - a)) \ norm (deriv f p)" if "y \ cball a r" for y proof - have [simp]: "norm (y - a) \ r" using that by (simp add: dist_norm norm_minus_commute) have "norm (g y) \ norm (g p)" using pge [OF that] by simp then have "norm (deriv f y) * abs (r - norm (y - a)) \ norm (deriv f p) * abs (r - norm (p - a))" by (simp only: dist_norm g_def norm_mult norm_of_real) with that \norm (p - a) < r\ show ?thesis by (simp add: dist_norm field_split_simps) qed have le_norm_dfp: "r / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [of a] \r > 0\ by auto have 1: "f holomorphic_on cball p t" apply (rule holomorphic_on_subset [OF holf]) using cpt \r < 1\ order_subst1 subset_ball by auto have 2: "norm (deriv f z) \ 2 * norm (deriv f p)" if "z \ ball p t" for z proof - have z: "z \ cball a r" by (meson ball_subset_cball subsetD cpt that) then have "norm(z - a) < r" by (metis ball_subset_cball contra_subsetD cpt dist_norm mem_ball norm_minus_commute that) have "norm (deriv f z) * (r - norm (z - a)) / (r - norm (p - a)) \ norm (deriv f p)" using gen_le_dfp [OF z] by simp with \norm (z - a) < r\ \norm (p - a) < r\ have "norm (deriv f z) \ (r - norm (p - a)) / (r - norm (z - a)) * norm (deriv f p)" by (simp add: field_simps) also have "... \ 2 * norm (deriv f p)" apply (rule mult_right_mono) using that \norm (p - a) < r\ \norm(z - a) < r\ apply (simp_all add: field_simps t_def dist_norm [symmetric]) using dist_triangle3 [of z a p] by linarith finally show ?thesis . qed have sqrt2: "sqrt 2 < 2113/1494" by (rule real_less_lsqrt) (auto simp: power2_eq_square) then have sq3: "0 < 3 - 2 * sqrt 2" by simp have "1 / 12 / ((3 - 2 * sqrt 2) / 2) < r" using sq3 sqrt2 by (auto simp: field_simps r_def) also have "... \ cmod (deriv f p) * (r - cmod (p - a))" using \norm (p - a) < r\ le_norm_dfp by (simp add: pos_divide_le_eq) finally have "1 / 12 < cmod (deriv f p) * (r - cmod (p - a)) * ((3 - 2 * sqrt 2) / 2)" using pos_divide_less_eq half_gt_zero_iff sq3 by blast then have **: "1 / 12 < (3 - 2 * sqrt 2) * t * norm (deriv f p)" using sq3 by (simp add: mult.commute t_def) have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball p t" by (rule Bloch_lemma [OF 1 \0 < t\ 2]) also have "... \ f ` ball a 1" apply (rule image_mono) apply (rule order_trans [OF ball_subset_cball]) apply (rule order_trans [OF cpt]) using \0 < t\ \r < 1\ apply (simp add: ball_subset_ball_iff dist_norm) done finally have "ball (f p) ((3 - 2 * sqrt 2) * t * norm (deriv f p)) \ f ` ball a 1" . with ** show ?thesis by (rule that) qed theorem Bloch: assumes holf: "f holomorphic_on ball a r" and "0 < r" and r': "r' \ r * norm (deriv f a) / 12" obtains b where "ball b r' \ f ` (ball a r)" proof (cases "deriv f a = 0") case True with r' show ?thesis using ball_eq_empty that by fastforce next case False define C where "C = deriv f a" have "0 < norm C" using False by (simp add: C_def) have dfa: "f field_differentiable at a" apply (rule holomorphic_on_imp_differentiable_at [OF holf]) using \0 < r\ by auto have fo: "(\z. f (a + of_real r * z)) = f o (\z. (a + of_real r * z))" by (simp add: o_def) have holf': "f holomorphic_on (\z. a + complex_of_real r * z) ` ball 0 1" apply (rule holomorphic_on_subset [OF holf]) using \0 < r\ apply (force simp: dist_norm norm_mult) done have 1: "(\z. f (a + r * z) / (C * r)) holomorphic_on ball 0 1" apply (rule holomorphic_intros holomorphic_on_compose holf' | simp add: fo)+ using \0 < r\ by (simp add: C_def False) have "((\z. f (a + of_real r * z) / (C * of_real r)) has_field_derivative (deriv f (a + of_real r * z) / C)) (at z)" if "norm z < 1" for z proof - have *: "((\x. f (a + of_real r * x)) has_field_derivative (deriv f (a + of_real r * z) * of_real r)) (at z)" apply (simp add: fo) apply (rule DERIV_chain [OF field_differentiable_derivI]) apply (rule holomorphic_on_imp_differentiable_at [OF holf], simp) using \0 < r\ apply (simp add: dist_norm norm_mult that) apply (rule derivative_eq_intros | simp)+ done show ?thesis apply (rule derivative_eq_intros * | simp)+ using \0 < r\ by (auto simp: C_def False) qed have 2: "deriv (\z. f (a + of_real r * z) / (C * of_real r)) 0 = 1" apply (subst deriv_cdivide_right) apply (simp add: field_differentiable_def fo) apply (rule exI) apply (rule DERIV_chain [OF field_differentiable_derivI]) apply (simp add: dfa) apply (rule derivative_eq_intros | simp add: C_def False fo)+ using \0 < r\ apply (simp add: C_def False fo) apply (simp add: derivative_intros dfa complex_derivative_chain) done have sb1: "(*) (C * r) ` (\z. f (a + of_real r * z) / (C * r)) ` ball 0 1 \ f ` ball a r" using \0 < r\ by (auto simp: dist_norm norm_mult C_def False) have sb2: "ball (C * r * b) r' \ (*) (C * r) ` ball b t" if "1 / 12 < t" for b t proof - have *: "r * cmod (deriv f a) / 12 \ r * (t * cmod (deriv f a))" using that \0 < r\ less_eq_real_def mult.commute mult.right_neutral mult_left_mono norm_ge_zero times_divide_eq_right by auto show ?thesis apply clarify apply (rule_tac x="x / (C * r)" in image_eqI) using \0 < r\ apply (simp_all add: dist_norm norm_mult norm_divide C_def False field_simps) apply (erule less_le_trans) apply (rule order_trans [OF r' *]) done qed show ?thesis apply (rule Bloch_unit [OF 1 2]) apply (rename_tac t) apply (rule_tac b="(C * of_real r) * b" in that) apply (drule image_mono [where f = "\z. (C * of_real r) * z"]) using sb1 sb2 apply force done qed corollary Bloch_general: assumes holf: "f holomorphic_on s" and "a \ s" and tle: "\z. z \ frontier s \ t \ dist a z" and rle: "r \ t * norm(deriv f a) / 12" obtains b where "ball b r \ f ` s" proof - consider "r \ 0" | "0 < t * norm(deriv f a) / 12" using rle by force then show ?thesis proof cases case 1 then show ?thesis by (simp add: ball_empty that) next case 2 show ?thesis proof (cases "deriv f a = 0") case True then show ?thesis using rle by (simp add: ball_empty that) next case False then have "t > 0" using 2 by (force simp: zero_less_mult_iff) have "\ ball a t \ s \ ball a t \ frontier s \ {}" apply (rule connected_Int_frontier [of "ball a t" s], simp_all) using \0 < t\ \a \ s\ centre_in_ball apply blast done with tle have *: "ball a t \ s" by fastforce then have 1: "f holomorphic_on ball a t" using holf using holomorphic_on_subset by blast show ?thesis apply (rule Bloch [OF 1 \t > 0\ rle]) apply (rule_tac b=b in that) using * apply force done qed qed qed subsection \Cauchy's residue theorem\ text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. Interactive Theorem Proving\ definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" lemma Eps_cong: assumes "\x. P x = Q x" shows "Eps P = Eps Q" using ext[of P Q, OF assms] by simp lemma residue_cong: assumes eq: "eventually (\z. f z = g z) (at z)" and "z = z'" shows "residue f z = residue g z'" proof - from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) let ?P = "\f c e. (\\>0. \ < e \ (f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \))" have "residue f z = residue g z" unfolding residue_def proof (rule Eps_cong) fix c :: complex - have "\e>0. ?P g c e" - if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g + have "\e>0. ?P g c e" + if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g proof - from that(1) obtain e where e: "e > 0" "?P f c e" by blast from that(2) obtain e' where e': "e' > 0" "\z'. z' \ z \ dist z' z < e' \ f z' = g z'" unfolding eventually_at by blast have "?P g c (min e e')" proof (intro allI exI impI, goal_cases) case (1 \) - hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" + hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" using e(2) by auto thus ?case proof (rule has_contour_integral_eq) fix z' assume "z' \ path_image (circlepath z \)" hence "dist z' z < e'" and "z' \ z" using 1 by (auto simp: dist_commute) with e'(2)[of z'] show "f z' = g z'" by simp qed qed moreover from e and e' have "min e e' > 0" by auto ultimately show ?thesis by blast qed from this[OF _ eq] and this[OF _ eq'] show "(\e>0. ?P f c e) \ (\e>0. ?P g c e)" by blast qed with assms show ?thesis by simp qed lemma contour_integral_circlepath_eq: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" and e2_cball:"cball z e2 \ s" shows "f contour_integrable_on circlepath z e1" "f contour_integrable_on circlepath z e2" "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" proof - define l where "l \ linepath (z+e2) (z+e1)" have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto have "e2>0" using \e1>0\ \e1\e2\ by auto have zl_img:"z\path_image l" proof assume "z \ path_image l" then have "e2 \ cmod (e2 - e1)" using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def by (auto simp add:closed_segment_commute) thus False using \e2>0\ \e1>0\ \e1\e2\ apply (subst (asm) norm_of_real) by auto qed define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" proof - show "f contour_integrable_on circlepath z e2" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e2>0\ e2_cball by auto show "f contour_integrable_on (circlepath z e1)" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e1>0\ \e1\e2\ e2_cball by auto qed have [simp]:"f contour_integrable_on l" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def by auto then show "f contour_integrable_on l" unfolding l_def apply (intro contour_integrable_continuous_linepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) by auto qed let ?ig="\g. contour_integral g f" have "(f has_contour_integral 0) g" proof (rule Cauchy_theorem_global[OF _ f_holo]) show "open (s - {z})" using \open s\ by auto show "valid_path g" unfolding g_def l_def by auto show "pathfinish g = pathstart g" unfolding g_def l_def by auto next have path_img:"path_image g \ cball z e2" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto ultimately show ?thesis unfolding g_def l_def using \e2>0\ by (simp add: path_image_join closed_segment_commute) qed show "path_image g \ s - {z}" proof - have "z\path_image g" using zl_img unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) moreover note \cball z e2 \ s\ and path_img ultimately show ?thesis by auto qed show "winding_number g w = 0" when"w \ s - {z}" for w proof - have "winding_number g w = 0" when "w\s" using that e2_cball apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) by (auto simp add:g_def l_def) moreover have "winding_number g z=0" proof - let ?Wz="\g. winding_number g z" have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + ?Wz (reversepath l)" using \e2>0\ \e1>0\ zl_img unfolding g_def l_def by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" using zl_img apply (subst (2) winding_number_reversepath) by (auto simp add:l_def closed_segment_commute) also have "... = 0" proof - have "?Wz (circlepath z e2) = 1" using \e2>0\ by (auto intro: winding_number_circlepath_centre) moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \e1>0\ apply (subst winding_number_reversepath) by (auto intro: winding_number_circlepath_centre) ultimately show ?thesis by auto qed finally show ?thesis . qed ultimately show ?thesis using that by auto qed qed then have "0 = ?ig g" using contour_integral_unique by simp also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + ?ig (reversepath l)" unfolding g_def by (auto simp add:contour_integrable_reversepath_eq) also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" by (auto simp add:contour_integral_reversepath) finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" by simp qed lemma base_residue: assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" and r_cball:"cball z r \ s" shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" proof - obtain e where "e>0" and e_cball:"cball z e \ s" using open_contains_cball[of s] \open s\ \z\s\ by auto define c where "c \ 2 * pi * \" define i where "i \ contour_integral (circlepath z e) f / c" have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ proof - have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" "f contour_integrable_on circlepath z \" "f contour_integrable_on circlepath z e" using \\ by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ then show ?thesis unfolding i_def c_def by (auto intro:has_contour_integral_integral) qed then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" unfolding residue_def c_def apply (rule_tac someI[of _ i],intro exI[where x=e]) by (auto simp add:\e>0\ c_def) then obtain e' where "e'>0" and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" by auto let ?int="\e. contour_integral (circlepath z e) f" define \ where "\ \ Min {r,e'} / 2" have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto have "(f has_contour_integral c * (residue f z)) (circlepath z \)" using e'_def[rule_format,OF \\>0\ \\] . then show ?thesis unfolding c_def using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) qed lemma residue_holo: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s" shows "residue f z = 0" proof - define c where "c \ 2 * pi * \" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f has_contour_integral c*residue f z) (circlepath z e)" using f_holo by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) moreover have "(f has_contour_integral 0) (circlepath z e)" using f_holo e_cball \e>0\ by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) ultimately have "c*residue f z =0" using has_contour_integral_unique by blast thus ?thesis unfolding c_def by auto qed lemma residue_const:"residue (\_. c) z = 0" by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) lemma residue_add: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z + g z) z= residue f z + residue g z" proof - define c where "c \ 2 * pi * \" define fg where "fg \ (\z. f z+g z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(fg has_contour_integral c * residue fg z) (circlepath z e)" unfolding fg_def using f_holo g_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro:holomorphic_intros) moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" unfolding fg_def using f_holo g_holo by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) ultimately have "c*(residue f z + residue g z) = c * residue fg z" using has_contour_integral_unique by (auto simp add:distrib_left) thus ?thesis unfolding fg_def by (auto simp add:c_def) qed lemma residue_lmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. c * (f z)) z= c * residue f z" proof (cases "c=0") case True thus ?thesis using residue_const by auto next case False define c' where "c' \ 2 * pi * \" define f' where "f' \ (\z. c * (f z))" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) by (auto intro:holomorphic_intros) moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" unfolding f'_def using f_holo by (auto intro: has_contour_integral_lmul base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) ultimately have "c' * residue f' z = c * (c' * residue f z)" using has_contour_integral_unique by auto thus ?thesis unfolding f'_def c'_def using False by (auto simp add:field_simps) qed lemma residue_rmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) * c) z= residue f z * c" using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) lemma residue_div: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) / c) z= residue f z / c " using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) lemma residue_neg: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. - (f z)) z= - residue f z" using residue_lmul[OF assms,of "-1"] by auto lemma residue_diff: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z - g z) z= residue f z - residue g z" using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] by (auto intro:holomorphic_intros g_holo) lemma residue_simple: assumes "open s" "z\s" and f_holo:"f holomorphic_on s" shows "residue (\w. f w / (w - z)) z = f z" proof - define c where "c \ 2 * pi * \" define f' where "f' \ \w. f w / (w - z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c * f z) (circlepath z e)" unfolding f'_def c_def using \e>0\ f_holo e_cball by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro!:holomorphic_intros) ultimately have "c * f z = c * residue f' z" using has_contour_integral_unique by blast thus ?thesis unfolding c_def f'_def by auto qed lemma residue_simple': - assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" + assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" and lim: "((\w. f w * (w - z)) \ c) (at z)" shows "residue f z = c" proof - define g where "g = (\w. if w = z then c else f w * (w - z))" from holo have "(\w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P") by (force intro: holomorphic_intros) also have "?P \ g holomorphic_on (s - {z})" by (intro holomorphic_cong refl) (simp_all add: g_def) finally have *: "g holomorphic_on (s - {z})" . note lim also have "(\w. f w * (w - z)) \z\ c \ g \z\ g z" by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter) finally have **: "g \z\ g z" . have g_holo: "g holomorphic_on s" by (rule no_isolated_singularity'[where K = "{z}"]) (insert assms * **, simp_all add: at_within_open_NO_MATCH) from s and this have "residue (\w. g w / (w - z)) z = g z" by (rule residue_simple) also have "\\<^sub>F za in at z. g za / (za - z) = f za" unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def) hence "residue (\w. g w / (w - z)) z = residue f z" by (intro residue_cong refl) finally show ?thesis by (simp add: g_def) qed lemma residue_holomorphic_over_power: assumes "open A" "z0 \ A" "f holomorphic_on A" shows "residue (\z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n" proof - let ?f = "\z. f z / (z - z0) ^ Suc n" from assms(1,2) obtain r where r: "r > 0" "cball z0 r \ A" by (auto simp: open_contains_cball) have "(?f has_contour_integral 2 * pi * \ * residue ?f z0) (circlepath z0 r)" using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros) moreover have "(?f has_contour_integral 2 * pi * \ / fact n * (deriv ^^ n) f z0) (circlepath z0 r)" using assms r by (intro Cauchy_has_contour_integral_higher_derivative_circlepath) (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on) - ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" + ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" by (rule has_contour_integral_unique) thus ?thesis by (simp add: field_simps) qed lemma residue_holomorphic_over_power': assumes "open A" "0 \ A" "f holomorphic_on A" shows "residue (\z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" using residue_holomorphic_over_power[OF assms] by simp lemma get_integrable_path: assumes "open s" "connected (s-pts)" "finite pts" "f holomorphic_on (s-pts) " "a\s-pts" "b\s-pts" obtains g where "valid_path g" "pathstart g = a" "pathfinish g = b" "path_image g \ s-pts" "f contour_integrable_on g" using assms proof (induct arbitrary:s thesis a rule:finite_induct[OF \finite pts\]) case 1 obtain g where "valid_path g" "path_image g \ s" "pathstart g = a" "pathfinish g = b" using connected_open_polynomial_connected[OF \open s\,of a b ] \connected (s - {})\ valid_path_polynomial_function "1.prems"(6) "1.prems"(7) by auto moreover have "f contour_integrable_on g" using contour_integrable_holomorphic_simple[OF _ \open s\ \valid_path g\ \path_image g \ s\,of f] \f holomorphic_on s - {}\ by auto ultimately show ?case using "1"(1)[of g] by auto next case idt:(2 p pts) obtain e where "e>0" and e:"\w\ball a e. w \ s \ (w \ a \ w \ insert p pts)" using finite_ball_avoid[OF \open s\ \finite (insert p pts)\, of a] \a \ s - insert p pts\ by auto define a' where "a' \ a+e/2" have "a'\s-{p} -pts" using e[rule_format,of "a+e/2"] \e>0\ by (auto simp add:dist_complex_def a'_def) then obtain g' where g'[simp]:"valid_path g'" "pathstart g' = a'" "pathfinish g' = b" "path_image g' \ s - {p} - pts" "f contour_integrable_on g'" using idt.hyps(3)[of a' "s-{p}"] idt.prems idt.hyps(1) by (metis Diff_insert2 open_delete) define g where "g \ linepath a a' +++ g'" have "valid_path g" unfolding g_def by (auto intro: valid_path_join) moreover have "pathstart g = a" and "pathfinish g = b" unfolding g_def by auto moreover have "path_image g \ s - insert p pts" unfolding g_def proof (rule subset_path_image_join) have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) then show "path_image (linepath a a') \ s - insert p pts" using e idt(9) by auto next show "path_image g' \ s - insert p pts" using g'(4) by blast qed moreover have "f contour_integrable_on g" proof - have "closed_segment a a' \ ball a e" using \e>0\ by (auto dest!:segment_bound1 simp:a'_def dist_complex_def norm_minus_commute) then have "continuous_on (closed_segment a a') f" using e idt.prems(6) holomorphic_on_imp_continuous_on[OF idt.prems(5)] apply (elim continuous_on_subset) by auto then have "f contour_integrable_on linepath a a'" using contour_integrable_continuous_linepath by auto then show ?thesis unfolding g_def apply (rule contour_integrable_joinI) by (auto simp add: \e>0\) qed ultimately show ?case using idt.prems(1)[of g] by auto qed lemma Cauchy_theorem_aux: assumes "open s" "connected (s-pts)" "finite pts" "pts \ s" "f holomorphic_on s-pts" "valid_path g" "pathfinish g = pathstart g" "path_image g \ s-pts" "\z. (z \ s) \ winding_number g z = 0" "\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using assms proof (induct arbitrary:s g rule:finite_induct[OF \finite pts\]) case 1 then show ?case by (simp add: Cauchy_theorem_global contour_integral_unique) next case (2 p pts) note fin[simp] = \finite (insert p pts)\ and connected = \connected (s - insert p pts)\ and valid[simp] = \valid_path g\ and g_loop[simp] = \pathfinish g = pathstart g\ and holo[simp]= \f holomorphic_on s - insert p pts\ and path_img = \path_image g \ s - insert p pts\ and winding = \\z. z \ s \ winding_number g z = 0\ and h = \\pa\s. 0 < h pa \ (\w\cball pa (h pa). w \ s \ (w \ pa \ w \ insert p pts))\ have "h p>0" and "p\s" and h_p: "\w\cball p (h p). w \ s \ (w \ p \ w \ insert p pts)" using h \insert p pts \ s\ by auto obtain pg where pg[simp]: "valid_path pg" "pathstart pg = pathstart g" "pathfinish pg=p+h p" "path_image pg \ s-insert p pts" "f contour_integrable_on pg" proof - have "p + h p\cball p (h p)" using h[rule_format,of p] by (simp add: \p \ s\ dist_norm) then have "p + h p \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by fastforce moreover have "pathstart g \ s - insert p pts " using path_img by auto ultimately show ?thesis using get_integrable_path[OF \open s\ connected fin holo,of "pathstart g" "p+h p"] that by blast qed obtain n::int where "n=winding_number g p" using integer_winding_number[OF _ g_loop,of p] valid path_img by (metis DiffD2 Ints_cases insertI1 subset_eq valid_path_imp_path) define p_circ where "p_circ \ circlepath p (h p)" define p_circ_pt where "p_circ_pt \ linepath (p+h p) (p+h p)" define n_circ where "n_circ \ \n. ((+++) p_circ ^^ n) p_circ_pt" define cp where "cp \ if n\0 then reversepath (n_circ (nat n)) else n_circ (nat (- n))" have n_circ:"valid_path (n_circ k)" "winding_number (n_circ k) p = k" "pathstart (n_circ k) = p + h p" "pathfinish (n_circ k) = p + h p" "path_image (n_circ k) = (if k=0 then {p + h p} else sphere p (h p))" "p \ path_image (n_circ k)" "\p'. p'\s - pts \ winding_number (n_circ k) p'=0 \ p'\path_image (n_circ k)" "f contour_integrable_on (n_circ k)" "contour_integral (n_circ k) f = k * contour_integral p_circ f" for k proof (induct k) case 0 show "valid_path (n_circ 0)" and "path_image (n_circ 0) = (if 0=0 then {p + h p} else sphere p (h p))" and "winding_number (n_circ 0) p = of_nat 0" and "pathstart (n_circ 0) = p + h p" and "pathfinish (n_circ 0) = p + h p" and "p \ path_image (n_circ 0)" unfolding n_circ_def p_circ_pt_def using \h p > 0\ by (auto simp add: dist_norm) show "winding_number (n_circ 0) p'=0 \ p'\path_image (n_circ 0)" when "p'\s- pts" for p' unfolding n_circ_def p_circ_pt_def apply (auto intro!:winding_number_trivial) by (metis Diff_iff pathfinish_in_path_image pg(3) pg(4) subsetCE subset_insertI that)+ show "f contour_integrable_on (n_circ 0)" unfolding n_circ_def p_circ_pt_def by (auto intro!:contour_integrable_continuous_linepath simp add:continuous_on_sing) show "contour_integral (n_circ 0) f = of_nat 0 * contour_integral p_circ f" unfolding n_circ_def p_circ_pt_def by auto next case (Suc k) have n_Suc:"n_circ (Suc k) = p_circ +++ n_circ k" unfolding n_circ_def by auto have pcirc:"p \ path_image p_circ" "valid_path p_circ" "pathfinish p_circ = pathstart (n_circ k)" using Suc(3) unfolding p_circ_def using \h p > 0\ by (auto simp add: p_circ_def) have pcirc_image:"path_image p_circ \ s - insert p pts" proof - have "path_image p_circ \ cball p (h p)" using \0 < h p\ p_circ_def by auto then show ?thesis using h_p pcirc(1) by auto qed have pcirc_integrable:"f contour_integrable_on p_circ" by (auto simp add:p_circ_def intro!: pcirc_image[unfolded p_circ_def] contour_integrable_continuous_circlepath holomorphic_on_imp_continuous_on holomorphic_on_subset[OF holo]) show "valid_path (n_circ (Suc k))" using valid_path_join[OF pcirc(2) Suc(1) pcirc(3)] unfolding n_circ_def by auto show "path_image (n_circ (Suc k)) = (if Suc k = 0 then {p + complex_of_real (h p)} else sphere p (h p))" proof - have "path_image p_circ = sphere p (h p)" unfolding p_circ_def using \0 < h p\ by auto then show ?thesis unfolding n_Suc using Suc.hyps(5) \h p>0\ by (auto simp add: path_image_join[OF pcirc(3)] dist_norm) qed then show "p \ path_image (n_circ (Suc k))" using \h p>0\ by auto show "winding_number (n_circ (Suc k)) p = of_nat (Suc k)" proof - have "winding_number p_circ p = 1" by (simp add: \h p > 0\ p_circ_def winding_number_circlepath_centre) moreover have "p \ path_image (n_circ k)" using Suc(5) \h p>0\ by auto then have "winding_number (p_circ +++ n_circ k) p = winding_number p_circ p + winding_number (n_circ k) p" using valid_path_imp_path Suc.hyps(1) Suc.hyps(2) pcirc apply (intro winding_number_join) by auto ultimately show ?thesis using Suc(2) unfolding n_circ_def by auto qed show "pathstart (n_circ (Suc k)) = p + h p" by (simp add: n_circ_def p_circ_def) show "pathfinish (n_circ (Suc k)) = p + h p" using Suc(4) unfolding n_circ_def by auto show "winding_number (n_circ (Suc k)) p'=0 \ p'\path_image (n_circ (Suc k))" when "p'\s-pts" for p' proof - have " p' \ path_image p_circ" using \p \ s\ h p_circ_def that using pcirc_image by blast moreover have "p' \ path_image (n_circ k)" using Suc.hyps(7) that by blast moreover have "winding_number p_circ p' = 0" proof - have "path_image p_circ \ cball p (h p)" using h unfolding p_circ_def using \p \ s\ by fastforce moreover have "p'\cball p (h p)" using \p \ s\ h that "2.hyps"(2) by fastforce ultimately show ?thesis unfolding p_circ_def apply (intro winding_number_zero_outside) by auto qed ultimately show ?thesis unfolding n_Suc apply (subst winding_number_join) by (auto simp: valid_path_imp_path pcirc Suc that not_in_path_image_join Suc.hyps(7)[OF that]) qed show "f contour_integrable_on (n_circ (Suc k))" unfolding n_Suc by (rule contour_integrable_joinI[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)]) show "contour_integral (n_circ (Suc k)) f = (Suc k) * contour_integral p_circ f" unfolding n_Suc by (auto simp add:contour_integral_join[OF pcirc_integrable Suc(8) pcirc(2) Suc(1)] Suc(9) algebra_simps) qed have cp[simp]:"pathstart cp = p + h p" "pathfinish cp = p + h p" "valid_path cp" "path_image cp \ s - insert p pts" "winding_number cp p = - n" "\p'. p'\s - pts \ winding_number cp p'=0 \ p' \ path_image cp" "f contour_integrable_on cp" "contour_integral cp f = - n * contour_integral p_circ f" proof - show "pathstart cp = p + h p" and "pathfinish cp = p + h p" and "valid_path cp" using n_circ unfolding cp_def by auto next have "sphere p (h p) \ s - insert p pts" using h[rule_format,of p] \insert p pts \ s\ by force moreover have "p + complex_of_real (h p) \ s - insert p pts" using pg(3) pg(4) by (metis pathfinish_in_path_image subsetCE) ultimately show "path_image cp \ s - insert p pts" unfolding cp_def using n_circ(5) by auto next show "winding_number cp p = - n" unfolding cp_def using winding_number_reversepath n_circ \h p>0\ by (auto simp: valid_path_imp_path) next show "winding_number cp p'=0 \ p' \ path_image cp" when "p'\s - pts" for p' unfolding cp_def apply (auto) apply (subst winding_number_reversepath) by (auto simp add: valid_path_imp_path n_circ(7)[OF that] n_circ(1)) next show "f contour_integrable_on cp" unfolding cp_def using contour_integrable_reversepath_eq n_circ(1,8) by auto next show "contour_integral cp f = - n * contour_integral p_circ f" unfolding cp_def using contour_integral_reversepath[OF n_circ(1)] n_circ(9) by auto qed define g' where "g' \ g +++ pg +++ cp +++ (reversepath pg)" have "contour_integral g' f = (\p\pts. winding_number g' p * contour_integral (circlepath p (h p)) f)" proof (rule "2.hyps"(3)[of "s-{p}" "g'",OF _ _ \finite pts\ ]) show "connected (s - {p} - pts)" using connected by (metis Diff_insert2) show "open (s - {p})" using \open s\ by auto show " pts \ s - {p}" using \insert p pts \ s\ \ p \ pts\ by blast show "f holomorphic_on s - {p} - pts" using holo \p \ pts\ by (metis Diff_insert2) show "valid_path g'" unfolding g'_def cp_def using n_circ valid pg g_loop by (auto intro!:valid_path_join ) show "pathfinish g' = pathstart g'" unfolding g'_def cp_def using pg(2) by simp show "path_image g' \ s - {p} - pts" proof - define s' where "s' \ s - {p} - pts" have s':"s' = s-insert p pts " unfolding s'_def by auto then show ?thesis using path_img pg(4) cp(4) unfolding g'_def apply (fold s'_def s') apply (intro subset_path_image_join) by auto qed note path_join_imp[simp] show "\z. z \ s - {p} \ winding_number g' z = 0" proof clarify fix z assume z:"z\s - {p}" have "winding_number (g +++ pg +++ cp +++ reversepath pg) z = winding_number g z + winding_number (pg +++ cp +++ (reversepath pg)) z" proof (rule winding_number_join) show "path g" using \valid_path g\ by (simp add: valid_path_imp_path) show "z \ path_image g" using z path_img by auto show "path (pg +++ cp +++ reversepath pg)" using pg(3) cp by (simp add: valid_path_imp_path) next have "path_image (pg +++ cp +++ reversepath pg) \ s - insert p pts" using pg(4) cp(4) by (auto simp:subset_path_image_join) then show "z \ path_image (pg +++ cp +++ reversepath pg)" using z by auto next show "pathfinish g = pathstart (pg +++ cp +++ reversepath pg)" using g_loop by auto qed also have "... = winding_number g z + (winding_number pg z + winding_number (cp +++ (reversepath pg)) z)" proof (subst add_left_cancel,rule winding_number_join) show "path pg" and "path (cp +++ reversepath pg)" and "pathfinish pg = pathstart (cp +++ reversepath pg)" by (auto simp add: valid_path_imp_path) show "z \ path_image pg" using pg(4) z by blast show "z \ path_image (cp +++ reversepath pg)" using z by (metis Diff_iff \z \ path_image pg\ contra_subsetD cp(4) insertI1 not_in_path_image_join path_image_reversepath singletonD) qed also have "... = winding_number g z + (winding_number pg z + (winding_number cp z + winding_number (reversepath pg) z))" apply (auto intro!:winding_number_join simp: valid_path_imp_path) apply (metis Diff_iff contra_subsetD cp(4) insertI1 singletonD z) by (metis Diff_insert2 Diff_subset contra_subsetD pg(4) z) also have "... = winding_number g z + winding_number cp z" apply (subst winding_number_reversepath) apply (auto simp: valid_path_imp_path) by (metis Diff_iff contra_subsetD insertI1 pg(4) singletonD z) finally have "winding_number g' z = winding_number g z + winding_number cp z" unfolding g'_def . moreover have "winding_number g z + winding_number cp z = 0" using winding z \n=winding_number g p\ by auto ultimately show "winding_number g' z = 0" unfolding g'_def by auto qed show "\pa\s - {p}. 0 < h pa \ (\w\cball pa (h pa). w \ s - {p} \ (w \ pa \ w \ pts))" using h by fastforce qed moreover have "contour_integral g' f = contour_integral g f - winding_number g p * contour_integral p_circ f" proof - have "contour_integral g' f = contour_integral g f + contour_integral (pg +++ cp +++ reversepath pg) f" unfolding g'_def apply (subst contour_integral_join) by (auto simp add:open_Diff[OF \open s\,OF finite_imp_closed[OF fin]] intro!: contour_integrable_holomorphic_simple[OF holo _ _ path_img] contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral pg f + contour_integral (cp +++ reversepath pg) f" apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral pg f + contour_integral cp f + contour_integral (reversepath pg) f" apply (subst contour_integral_join) by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f + contour_integral cp f" using contour_integral_reversepath by (auto simp add:contour_integrable_reversepath) also have "... = contour_integral g f - winding_number g p * contour_integral p_circ f" using \n=winding_number g p\ by auto finally show ?thesis . qed moreover have "winding_number g' p' = winding_number g p'" when "p'\pts" for p' proof - have [simp]: "p' \ path_image g" "p' \ path_image pg" "p'\path_image cp" using "2.prems"(8) that apply blast apply (metis Diff_iff Diff_insert2 contra_subsetD pg(4) that) by (meson DiffD2 cp(4) rev_subsetD subset_insertI that) have "winding_number g' p' = winding_number g p' + winding_number (pg +++ cp +++ reversepath pg) p'" unfolding g'_def apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) by auto also have "... = winding_number g p' + winding_number pg p' + winding_number (cp +++ reversepath pg) p'" apply (subst winding_number_join) apply (simp_all add: valid_path_imp_path) apply (intro not_in_path_image_join) by auto also have "... = winding_number g p' + winding_number pg p'+ winding_number cp p' + winding_number (reversepath pg) p'" apply (subst winding_number_join) by (simp_all add: valid_path_imp_path) also have "... = winding_number g p' + winding_number cp p'" apply (subst winding_number_reversepath) by (simp_all add: valid_path_imp_path) also have "... = winding_number g p'" using that by auto finally show ?thesis . qed ultimately show ?case unfolding p_circ_def apply (subst (asm) sum.cong[OF refl, of pts _ "\p. winding_number g p * contour_integral (circlepath p (h p)) f"]) by (auto simp add:sum.insert[OF \finite pts\ \p\pts\] algebra_simps) qed lemma Cauchy_theorem_singularities: assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and homo:"\z. (z \ s) \ winding_number g z = 0" and avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" shows "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" (is "?L=?R") proof - define circ where "circ \ \p. winding_number g p * contour_integral (circlepath p (h p)) f" define pts1 where "pts1 \ pts \ s" define pts2 where "pts2 \ pts - pts1" have "pts=pts1 \ pts2" "pts1 \ pts2 = {}" "pts2 \ s={}" "pts1\s" unfolding pts1_def pts2_def by auto have "contour_integral g f = (\p\pts1. circ p)" unfolding circ_def proof (rule Cauchy_theorem_aux[OF \open s\ _ _ \pts1\s\ _ \valid_path g\ loop _ homo]) have "finite pts1" unfolding pts1_def using \finite pts\ by auto then show "connected (s - pts1)" using \open s\ \connected s\ connected_open_delete_finite[of s] by auto next show "finite pts1" using \pts = pts1 \ pts2\ assms(3) by auto show "f holomorphic_on s - pts1" by (metis Diff_Int2 Int_absorb holo pts1_def) show "path_image g \ s - pts1" using assms(7) pts1_def by auto show "\p\s. 0 < h p \ (\w\cball p (h p). w \ s \ (w \ p \ w \ pts1))" by (simp add: avoid pts1_def) qed moreover have "sum circ pts2=0" proof - have "winding_number g p=0" when "p\pts2" for p using \pts2 \ s={}\ that homo[rule_format,of p] by auto thus ?thesis unfolding circ_def apply (intro sum.neutral) by auto qed moreover have "?R=sum circ pts1 + sum circ pts2" unfolding circ_def using sum.union_disjoint[OF _ _ \pts1 \ pts2 = {}\] \finite pts\ \pts=pts1 \ pts2\ by blast ultimately show ?thesis apply (fold circ_def) by auto qed theorem Residue_theorem: fixes s pts::"complex set" and f::"complex \ complex" and g::"real \ complex" assumes "open s" "connected s" "finite pts" and holo:"f holomorphic_on s-pts" and "valid_path g" and loop:"pathfinish g = pathstart g" and "path_image g \ s-pts" and homo:"\z. (z \ s) \ winding_number g z = 0" shows "contour_integral g f = 2 * pi * \ *(\p\pts. winding_number g p * residue f p)" proof - define c where "c \ 2 * pi * \" obtain h where avoid:"\p\s. h p>0 \ (\w\cball p (h p). w\s \ (w\p \ w \ pts))" using finite_cball_avoid[OF \open s\ \finite pts\] by metis have "contour_integral g f = (\p\pts. winding_number g p * contour_integral (circlepath p (h p)) f)" using Cauchy_theorem_singularities[OF assms avoid] . also have "... = (\p\pts. c * winding_number g p * residue f p)" proof (intro sum.cong) show "pts = pts" by simp next fix x assume "x \ pts" show "winding_number g x * contour_integral (circlepath x (h x)) f = c * winding_number g x * residue f x" proof (cases "x\s") case False then have "winding_number g x=0" using homo by auto thus ?thesis by auto next case True have "contour_integral (circlepath x (h x)) f = c* residue f x" using \x\pts\ \finite pts\ avoid[rule_format,OF True] apply (intro base_residue[of "s-(pts-{x})",THEN contour_integral_unique,folded c_def]) by (auto intro:holomorphic_on_subset[OF holo] open_Diff[OF \open s\ finite_imp_closed]) then show ?thesis by auto qed qed also have "... = c * (\p\pts. winding_number g p * residue f p)" by (simp add: sum_distrib_left algebra_simps) finally show ?thesis unfolding c_def . qed subsection \Non-essential singular points\ -definition\<^marker>\tag important\ is_pole :: +definition\<^marker>\tag important\ is_pole :: "('a::topological_space \ 'b::real_normed_vector) \ 'a \ bool" where "is_pole f a = (LIM x (at a). f x :> at_infinity)" lemma is_pole_cong: assumes "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole f a \ is_pole g b" unfolding is_pole_def using assms by (intro filterlim_cong,auto) lemma is_pole_transform: assumes "is_pole f a" "eventually (\x. f x = g x) (at a)" "a=b" shows "is_pole g b" using is_pole_cong assms by auto lemma is_pole_tendsto: fixes f::"('a::topological_space \ 'b::real_normed_div_algebra)" shows "is_pole f x \ ((inverse o f) \ 0) (at x)" unfolding is_pole_def by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at) lemma is_pole_inverse_holomorphic: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and pole:"is_pole f z" and non_z:"\x\s-{z}. f x\0" shows "(\x. if x=z then 0 else inverse (f x)) holomorphic_on s" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" have "isCont g z" unfolding isCont_def using is_pole_tendsto[OF pole] apply (subst Lim_cong_at[where b=z and y=0 and g="inverse \ f"]) by (simp_all add:g_def) moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def by (auto elim!:continuous_on_inverse simp add:non_z) hence "continuous_on (s-{z}) g" unfolding g_def apply (subst continuous_on_cong[where t="s-{z}" and g="inverse o f"]) by auto ultimately have "continuous_on s g" using open_delete[OF \open s\] \open s\ by (auto simp add:continuous_on_eq_continuous_at) moreover have "(inverse o f) holomorphic_on (s-{z})" unfolding comp_def using f_holo by (auto elim!:holomorphic_on_inverse simp add:non_z) hence "g holomorphic_on (s-{z})" apply (subst holomorphic_cong[where t="s-{z}" and g="inverse o f"]) by (auto simp add:g_def) ultimately show ?thesis unfolding g_def using \open s\ by (auto elim!: no_isolated_singularity) qed lemma not_is_pole_holomorphic: assumes "open A" "x \ A" "f holomorphic_on A" shows "\is_pole f x" proof - have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at) hence "f \x\ f x" by (simp add: isCont_def) thus "\is_pole f x" unfolding is_pole_def using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto qed lemma is_pole_inverse_power: "n > 0 \ is_pole (\z::complex. 1 / (z - a) ^ n) a" unfolding is_pole_def inverse_eq_divide [symmetric] by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros) (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros) lemma is_pole_inverse: "is_pole (\z::complex. 1 / (z - a)) a" using is_pole_inverse_power[of 1 a] by simp lemma is_pole_divide: fixes f :: "'a :: t2_space \ 'b :: real_normed_field" assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \ 0" shows "is_pole (\z. f z / g z) z" proof - have "filterlim (\z. f z * inverse (g z)) at_infinity (at z)" by (intro tendsto_mult_filterlim_at_infinity[of _ "f z"] filterlim_compose[OF filterlim_inverse_at_infinity])+ (insert assms, auto simp: isCont_def) thus ?thesis by (simp add: field_split_simps is_pole_def) qed lemma is_pole_basic: assumes "f holomorphic_on A" "open A" "z \ A" "f z \ 0" "n > 0" shows "is_pole (\w. f w / (w - z) ^ n) z" proof (rule is_pole_divide) have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at) have "filterlim (\w. (w - z) ^ n) (nhds 0) (at z)" using assms by (auto intro!: tendsto_eq_intros) thus "filterlim (\w. (w - z) ^ n) (at 0) (at z)" by (intro filterlim_atI tendsto_eq_intros) (insert assms, auto simp: eventually_at_filter) qed fact+ lemma is_pole_basic': assumes "f holomorphic_on A" "open A" "0 \ A" "f 0 \ 0" "n > 0" shows "is_pole (\w. f w / w ^ n) 0" using is_pole_basic[of f A 0] assms by simp -text \The proposition - \<^term>\\x. ((f::complex\complex) \ x) (at z) \ is_pole f z\ -can be interpreted as the complex function \<^term>\f\ has a non-essential singularity at \<^term>\z\ -(i.e. the singularity is either removable or a pole).\ +text \The proposition + \<^term>\\x. ((f::complex\complex) \ x) (at z) \ is_pole f z\ +can be interpreted as the complex function \<^term>\f\ has a non-essential singularity at \<^term>\z\ +(i.e. the singularity is either removable or a pole).\ definition not_essential::"[complex \ complex, complex] \ bool" where "not_essential f z = (\x. f\z\x \ is_pole f z)" definition isolated_singularity_at::"[complex \ complex, complex] \ bool" where "isolated_singularity_at f z = (\r>0. f analytic_on ball z r-{z})" named_theorems singularity_intros "introduction rules for singularities" lemma holomorphic_factor_unique: fixes f::"complex \ complex" and z::complex and r::real and m n::int assumes "r>0" "g z\0" "h z\0" and asm:"\w\ball z r-{z}. f w = g w * (w-z) powr n \ g w\0 \ f w = h w * (w - z) powr m \ h w\0" and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r" shows "n=m" proof - have [simp]:"at z within ball z r \ bot" using \r>0\ by (auto simp add:at_within_ball_bot_iff) have False when "n>m" proof - have "(h \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (n - m) * g w"]) have "\w\ball z r-{z}. h w = (w-z)powr(n-m) * g w" using \n>m\ asm \r>0\ apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (n - m) * g x' = h x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \ \x. (x - z) powr (n-m)" have "f' z=0" using \n>m\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) - using \n>m\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) + using \n>m\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def by (simp add: continuous_within) moreover have "(g \ g z) F" using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * g w) \ 0) F" using tendsto_mult by fastforce qed moreover have "(h \ h z) (at z within ball z r)" using holomorphic_on_imp_continuous_on[OF h_holo] by (auto simp add:continuous_on_def \r>0\) ultimately have "h z=0" by (auto intro!: tendsto_unique) thus False using \h z\0\ by auto qed moreover have False when "m>n" proof - have "(g \ 0) (at z within ball z r)" proof (rule Lim_transform_within[OF _ \r>0\, where f="\w. (w - z) powr (m - n) * h w"]) have "\w\ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \m>n\ asm apply (auto simp add:field_simps powr_diff) by force then show "\x' \ ball z r; 0 < dist x' z;dist x' z < r\ \ (x' - z) powr (m - n) * h x' = g x'" for x' by auto next define F where "F \ at z within ball z r" define f' where "f' \\x. (x - z) powr (m-n)" have "f' z=0" using \m>n\ unfolding f'_def by auto moreover have "continuous F f'" unfolding f'_def F_def continuous_def apply (subst Lim_ident_at) using \m>n\ by (auto intro!:tendsto_powr_complex_0 tendsto_eq_intros) ultimately have "(f' \ 0) F" unfolding F_def by (simp add: continuous_within) moreover have "(h \ h z) F" using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \r>0\ unfolding F_def by auto ultimately show " ((\w. f' w * h w) \ 0) F" using tendsto_mult by fastforce qed moreover have "(g \ g z) (at z within ball z r)" using holomorphic_on_imp_continuous_on[OF g_holo] by (auto simp add:continuous_on_def \r>0\) ultimately have "g z=0" by (auto intro!: tendsto_unique) thus False using \g z\0\ by auto qed ultimately show "n=m" by fastforce qed lemma holomorphic_factor_puncture: - assumes f_iso:"isolated_singularity_at f z" + assumes f_iso:"isolated_singularity_at f z" and "not_essential f z" \ \\<^term>\f\ has either a removable singularity or a pole at \<^term>\z\\ and non_zero:"\\<^sub>Fw in (at z). f w\0" \ \\<^term>\f\ will not be constantly zero in a neighbour of \<^term>\z\\ shows "\!n::int. \g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r-{z}. f w = g w * (w-z) powr n \ g w\0)" proof - define P where "P = (\f n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" - have imp_unique:"\!n::int. \g r. P f n g r" when "\n g r. P f n g r" + have imp_unique:"\!n::int. \g r. P f n g r" when "\n g r. P f n g r" proof (rule ex_ex1I[OF that]) fix n1 n2 :: int assume g1_asm:"\g1 r1. P f n1 g1 r1" and g2_asm:"\g2 r2. P f n2 g2 r2" define fac where "fac \ \n g r. \w\cball z r-{z}. f w = g w * (w - z) powr (of_int n) \ g w \ 0" obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\0" and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\0" and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto define r where "r \ min r1 r2" have "r>0" using \r1>0\ \r2>0\ unfolding r_def by auto - moreover have "\w\ball z r-{z}. f w = g1 w * (w-z) powr n1 \ g1 w\0 + moreover have "\w\ball z r-{z}. f w = g1 w * (w-z) powr n1 \ g1 w\0 \ f w = g2 w * (w - z) powr n2 \ g2 w\0" using \fac n1 g1 r1\ \fac n2 g2 r2\ unfolding fac_def r_def by fastforce ultimately show "n1=n2" using g1_holo g2_holo \g1 z\0\ \g2 z\0\ apply (elim holomorphic_factor_unique) - by (auto simp add:r_def) + by (auto simp add:r_def) qed - have P_exist:"\ n g r. P h n g r" when - "\z'. (h \ z') (at z)" "isolated_singularity_at h z" "\\<^sub>Fw in (at z). h w\0" + have P_exist:"\ n g r. P h n g r" when + "\z'. (h \ z') (at z)" "isolated_singularity_at h z" "\\<^sub>Fw in (at z). h w\0" for h proof - from that(2) obtain r where "r>0" "h analytic_on ball z r - {z}" unfolding isolated_singularity_at_def by auto obtain z' where "(h \ z') (at z)" using \\z'. (h \ z') (at z)\ by auto define h' where "h'=(\x. if x=z then z' else h x)" have "h' holomorphic_on ball z r" - apply (rule no_isolated_singularity'[of "{z}"]) + apply (rule no_isolated_singularity'[of "{z}"]) subgoal by (metis LIM_equal Lim_at_imp_Lim_at_within \h \z\ z'\ empty_iff h'_def insert_iff) - subgoal using \h analytic_on ball z r - {z}\ analytic_imp_holomorphic h'_def holomorphic_transform + subgoal using \h analytic_on ball z r - {z}\ analytic_imp_holomorphic h'_def holomorphic_transform by fastforce by auto have ?thesis when "z'=0" - proof - + proof - have "h' z=0" using that unfolding h'_def by auto - moreover have "\ h' constant_on ball z r" + moreover have "\ h' constant_on ball z r" using \\\<^sub>Fw in (at z). h w\0\ unfolding constant_on_def frequently_def eventually_at h'_def apply simp by (metis \0 < r\ centre_in_ball dist_commute mem_ball that) moreover note \h' holomorphic_on ball z r\ ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \ ball z r" and g:"g holomorphic_on ball z r1" "\w. w \ ball z r1 \ h' w = (w - z) ^ n * g w" - "\w. w \ ball z r1 \ g w \ 0" + "\w. w \ ball z r1 \ g w \ 0" using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified, - OF \h' holomorphic_on ball z r\ \r>0\ \h' z=0\ \\ h' constant_on ball z r\] + OF \h' holomorphic_on ball z r\ \r>0\ \h' z=0\ \\ h' constant_on ball z r\] by (auto simp add:dist_commute) define rr where "rr=r1/2" have "P h' n g rr" unfolding P_def rr_def using \n>0\ \r1>0\ g by (auto simp add:powr_nat) then have "P h n g rr" unfolding h'_def P_def by auto then show ?thesis unfolding P_def by blast qed moreover have ?thesis when "z'\0" proof - have "h' z\0" using that unfolding h'_def by auto obtain r1 where "r1>0" "cball z r1 \ ball z r" "\x\cball z r1. h' x\0" proof - have "isCont h' z" "h' z\0" by (auto simp add: Lim_cong_within \h \z\ z'\ \z'\0\ continuous_at h'_def) then obtain r2 where r2:"r2>0" "\x\ball z r2. h' x\0" using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto define r1 where "r1=min r2 r / 2" - have "0 < r1" "cball z r1 \ ball z r" + have "0 < r1" "cball z r1 \ ball z r" using \r2>0\ \r>0\ unfolding r1_def by auto - moreover have "\x\cball z r1. h' x \ 0" + moreover have "\x\cball z r1. h' x \ 0" using r2 unfolding r1_def by simp ultimately show ?thesis using that by auto qed then have "P h' 0 h' r1" using \h' holomorphic_on ball z r\ unfolding P_def by auto then have "P h 0 h' r1" unfolding P_def h'_def by auto then show ?thesis unfolding P_def by blast qed ultimately show ?thesis by auto qed have ?thesis when "\x. (f \ x) (at z)" apply (rule_tac imp_unique[unfolded P_def]) using P_exist[OF that(1) f_iso non_zero] unfolding P_def . moreover have ?thesis when "is_pole f z" proof (rule imp_unique[unfolded P_def]) obtain e where [simp]:"e>0" and e_holo:"f holomorphic_on ball z e - {z}" and e_nz: "\x\ball z e-{z}. f x\0" proof - have "\\<^sub>F z in at z. f z \ 0" using \is_pole f z\ filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def by auto then obtain e1 where e1:"e1>0" "\x\ball z e1-{z}. f x\0" using that eventually_at[of "\x. f x\0" z UNIV,simplified] by (auto simp add:dist_commute) obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto define e where "e=min e1 e2" show ?thesis apply (rule that[of e]) using e1 e2 unfolding e_def by auto qed - + define h where "h \ \x. inverse (f x)" have "\n g r. P h n g r" proof - - have "h \z\ 0" + have "h \z\ 0" using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce moreover have "\\<^sub>Fw in (at z). h w\0" - using non_zero + using non_zero apply (elim frequently_rev_mp) unfolding h_def eventually_at by (auto intro:exI[where x=1]) moreover have "isolated_singularity_at h z" unfolding isolated_singularity_at_def h_def apply (rule exI[where x=e]) - using e_holo e_nz \e>0\ by (metis open_ball analytic_on_open + using e_holo e_nz \e>0\ by (metis open_ball analytic_on_open holomorphic_on_inverse open_delete) ultimately show ?thesis using P_exist[of h] by auto qed then obtain n g r where "0 < r" and g_holo:"g holomorphic_on cball z r" and "g z\0" and g_fac:"(\w\cball z r-{z}. h w = g w * (w - z) powr of_int n \ g w \ 0)" unfolding P_def by auto have "P f (-n) (inverse o g) r" proof - have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\cball z r - {z}" for w - using g_fac[rule_format,of w] that unfolding h_def + using g_fac[rule_format,of w] that unfolding h_def apply (auto simp add:powr_minus ) by (metis inverse_inverse_eq inverse_mult_distrib) - then show ?thesis + then show ?thesis unfolding P_def comp_def using \r>0\ g_holo g_fac \g z\0\ by (auto intro:holomorphic_intros) qed - then show "\x g r. 0 < r \ g holomorphic_on cball z r \ g z \ 0 + then show "\x g r. 0 < r \ g holomorphic_on cball z r \ g z \ 0 \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int x \ g w \ 0)" unfolding P_def by blast qed ultimately show ?thesis using \not_essential f z\ unfolding not_essential_def by presburger qed lemma not_essential_transform: assumes "not_essential g z" assumes "\\<^sub>F w in (at z). g w = f w" - shows "not_essential f z" + shows "not_essential f z" using assms unfolding not_essential_def by (simp add: filterlim_cong is_pole_cong) lemma isolated_singularity_at_transform: assumes "isolated_singularity_at g z" assumes "\\<^sub>F w in (at z). g w = f w" - shows "isolated_singularity_at f z" + shows "isolated_singularity_at f z" proof - obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto obtain r2 where "r2>0" and r2:" \x. x \ z \ dist x z < r2 \ g x = f x" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "r3>0" unfolding r3_def using \r1>0\ \r2>0\ by auto moreover have "f analytic_on ball z r3 - {z}" proof - have "g holomorphic_on ball z r3 - {z}" using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto) then have "f holomorphic_on ball z r3 - {z}" - using r2 unfolding r3_def + using r2 unfolding r3_def by (auto simp add:dist_commute elim!:holomorphic_transform) - then show ?thesis by (subst analytic_on_open,auto) + then show ?thesis by (subst analytic_on_open,auto) qed ultimately show ?thesis unfolding isolated_singularity_at_def by auto qed lemma not_essential_powr[singularity_intros]: assumes "LIM w (at z). f w :> (at x)" shows "not_essential (\w. (f w) powr (of_int n)) z" proof - define fp where "fp=(\w. (f w) powr (of_int n))" have ?thesis when "n>0" proof - - have "(\w. (f w) ^ (nat n)) \z\ x ^ nat n" + have "(\w. (f w) ^ (nat n)) \z\ x ^ nat n" using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros) - then have "fp \z\ x ^ nat n" unfolding fp_def + then have "fp \z\ x ^ nat n" unfolding fp_def apply (elim Lim_transform_within[where d=1],simp) by (metis less_le powr_0 powr_of_int that zero_less_nat_eq zero_power) then show ?thesis unfolding not_essential_def fp_def by auto qed moreover have ?thesis when "n=0" proof - - have "fp \z\ 1 " + have "fp \z\ 1 " apply (subst tendsto_cong[where g="\_.1"]) using that filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def by auto then show ?thesis unfolding fp_def not_essential_def by auto qed moreover have ?thesis when "n<0" proof (cases "x=0") case True have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity" apply (subst filterlim_inverse_at_iff[symmetric],simp) apply (rule filterlim_atI) subgoal using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros) - subgoal using filterlim_at_within_not_equal[OF assms,of 0] + subgoal using filterlim_at_within_not_equal[OF assms,of 0] by (eventually_elim,insert that,auto) done then have "LIM w (at z). fp w :> at_infinity" proof (elim filterlim_mono_eventually) show "\\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x" using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def apply eventually_elim using powr_of_int that by auto qed auto then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto next case False let ?xx= "inverse (x ^ (nat (-n)))" have "(\w. inverse ((f w) ^ (nat (-n)))) \z\?xx" using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros) then have "fp \z\?xx" apply (elim Lim_transform_within[where d=1],simp) - unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less + unfolding fp_def by (metis inverse_zero nat_mono_iff nat_zero_as_int neg_0_less_iff_less not_le power_eq_0_iff powr_0 powr_of_int that) then show ?thesis unfolding fp_def not_essential_def by auto qed ultimately show ?thesis by linarith qed lemma isolated_singularity_at_powr[singularity_intros]: assumes "isolated_singularity_at f z" "\\<^sub>F w in (at z). f w\0" shows "isolated_singularity_at (\w. (f w) powr (of_int n)) z" proof - obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}" using assms(1) unfolding isolated_singularity_at_def by auto then have r1:"f holomorphic_on ball z r1 - {z}" using analytic_on_open[of "ball z r1-{z}" f] by blast obtain r2 where "r2>0" and r2:"\w. w \ z \ dist w z < r2 \ f w \ 0" using assms(2) unfolding eventually_at by auto define r3 where "r3=min r1 r2" have "(\w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}" apply (rule holomorphic_on_powr_of_int) subgoal unfolding r3_def using r1 by auto subgoal unfolding r3_def using r2 by (auto simp add:dist_commute) done moreover have "r3>0" unfolding r3_def using \0 < r1\ \0 < r2\ by linarith ultimately show ?thesis unfolding isolated_singularity_at_def apply (subst (asm) analytic_on_open[symmetric]) by auto qed lemma non_zero_neighbour: - assumes f_iso:"isolated_singularity_at f z" - and f_ness:"not_essential f z" + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "\\<^sub>F w in (at z). f w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" + and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto have "f w \ 0" when " w \ z" "dist w z < fr" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w \ 0" using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute) moreover have "(w - z) powr of_int fn \0" unfolding powr_eq_0_iff using \w\z\ by auto ultimately show ?thesis by auto qed then show ?thesis using \fr>0\ unfolding eventually_at by auto qed lemma non_zero_neighbour_pole: assumes "is_pole f z" shows "\\<^sub>F w in (at z). f w\0" - using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0] + using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0] unfolding is_pole_def by auto lemma non_zero_neighbour_alt: assumes holo: "f holomorphic_on S" and "open S" "connected S" "z \ S" "\ \ S" "f \ \ 0" shows "\\<^sub>F w in (at z). f w\0 \ w\S" proof (cases "f z = 0") case True - from isolated_zeros[OF holo \open S\ \connected S\ \z \ S\ True \\ \ S\ \f \ \ 0\] - obtain r where "0 < r" "ball z r \ S" "\w \ ball z r - {z}.f w \ 0" by metis - then show ?thesis unfolding eventually_at + from isolated_zeros[OF holo \open S\ \connected S\ \z \ S\ True \\ \ S\ \f \ \ 0\] + obtain r where "0 < r" "ball z r \ S" "\w \ ball z r - {z}.f w \ 0" by metis + then show ?thesis unfolding eventually_at apply (rule_tac x=r in exI) by (auto simp add:dist_commute) next case False obtain r1 where r1:"r1>0" "\y. dist z y < r1 \ f y \ 0" - using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at + using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at holo holomorphic_on_imp_continuous_on by blast - obtain r2 where r2:"r2>0" "ball z r2 \ S" + obtain r2 where r2:"r2>0" "ball z r2 \ S" using assms(2) assms(4) openE by blast - show ?thesis unfolding eventually_at + show ?thesis unfolding eventually_at apply (rule_tac x="min r1 r2" in exI) using r1 r2 by (auto simp add:dist_commute) qed lemma not_essential_times[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w * g w) z" proof - define fg where "fg = (\w. f w * g w)" have ?thesis when "\ ((\\<^sub>Fw in (at z). f w\0) \ (\\<^sub>Fw in (at z). g w\0))" proof - - have "\\<^sub>Fw in (at z). fg w=0" + have "\\<^sub>Fw in (at z). fg w=0" using that[unfolded frequently_def, simplified] unfolding fg_def by (auto elim: eventually_rev_mp) from tendsto_cong[OF this] have "fg \z\0" by auto then show ?thesis unfolding not_essential_def fg_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" and g_nconst:"\\<^sub>Fw in (at z). g w\0" proof - obtain fn fp fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" + and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto obtain gn gp gr where [simp]:"gp z \ 0" and "gr > 0" - and gr: "gp holomorphic_on cball z gr" + and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto - + define r1 where "r1=(min fr gr)" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed have [intro]: "fp \z\fp z" "gp \z\gp z" using fr(1) \fr>0\ gr(1) \gr>0\ - by (meson open_ball ball_subset_cball centre_in_ball - continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on + by (meson open_ball ball_subset_cball centre_in_ball + continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on holomorphic_on_subset)+ - have ?thesis when "fn+gn>0" + have ?thesis when "fn+gn>0" proof - - have "(\w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \z\0" + have "(\w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \z\0" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ 0" apply (elim Lim_transform_within[OF _ \r1>0\]) - by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self - eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int + by (metis (no_types, hide_lams) Diff_iff cball_trivial dist_commute dist_self + eq_iff_diff_eq_0 fg_times less_le linorder_not_le mem_ball mem_cball powr_of_int that) then show ?thesis unfolding not_essential_def fg_def by auto qed - moreover have ?thesis when "fn+gn=0" + moreover have ?thesis when "fn+gn=0" proof - - have "(\w. fp w * gp w) \z\fp z*gp z" + have "(\w. fp w * gp w) \z\fp z*gp z" using that by (auto intro!:tendsto_eq_intros) then have "fg \z\ fp z*gp z" apply (elim Lim_transform_within[OF _ \r1>0\]) apply (subst fg_times) by (auto simp add:dist_commute that) then show ?thesis unfolding not_essential_def fg_def by auto qed - moreover have ?thesis when "fn+gn<0" + moreover have ?thesis when "fn+gn<0" proof - have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity" apply (rule filterlim_divide_at_infinity) apply (insert that, auto intro!:tendsto_eq_intros filterlim_atI) using eventually_at_topological by blast then have "is_pole fg z" unfolding is_pole_def apply (elim filterlim_transform_within[OF _ _ \r1>0\],simp) apply (subst fg_times,simp add:dist_commute) apply (subst powr_of_int) using that by (auto simp add:field_split_simps) then show ?thesis unfolding not_essential_def fg_def by auto qed ultimately show ?thesis unfolding not_essential_def fg_def by fastforce qed ultimately show ?thesis by auto qed lemma not_essential_inverse[singularity_intros]: assumes f_ness:"not_essential f z" assumes f_iso:"isolated_singularity_at f z" shows "not_essential (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - - have "\\<^sub>Fw in (at z). f w=0" + have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto from tendsto_cong[OF this] have "vf \z\0" unfolding vf_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "is_pole f z" proof - have "vf \z\0" using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "\x. f\z\x " and f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - from that obtain fz where fz:"f\z\fz" by auto have ?thesis when "fz=0" proof - have "(\w. inverse (vf w)) \z\0" using fz that unfolding vf_def by auto moreover have "\\<^sub>F w in at z. inverse (vf w) \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] unfolding vf_def by auto ultimately have "is_pole vf z" using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto then show ?thesis unfolding not_essential_def vf_def by auto qed moreover have ?thesis when "fz\0" proof - have "vf \z\inverse fz" using fz that unfolding vf_def by (auto intro:tendsto_eq_intros) then show ?thesis unfolding not_essential_def vf_def by auto qed ultimately show ?thesis by auto qed ultimately show ?thesis using f_ness unfolding not_essential_def by auto qed lemma isolated_singularity_at_inverse[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and f_ness:"not_essential f z" shows "isolated_singularity_at (\w. inverse (f w)) z" proof - define vf where "vf = (\w. inverse (f w))" have ?thesis when "\(\\<^sub>Fw in (at z). f w\0)" proof - - have "\\<^sub>Fw in (at z). f w=0" + have "\\<^sub>Fw in (at z). f w=0" using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp) then have "\\<^sub>Fw in (at z). vf w=0" unfolding vf_def by auto then obtain d1 where "d1>0" and d1:"\x. x \ z \ dist x z < d1 \ vf x = 0" unfolding eventually_at by auto then have "vf holomorphic_on ball z d1-{z}" apply (rule_tac holomorphic_transform[of "\_. 0"]) by (auto simp add:dist_commute) then have "vf analytic_on ball z d1 - {z}" by (simp add: analytic_on_open open_delete) then show ?thesis using \d1>0\ unfolding isolated_singularity_at_def vf_def by auto qed moreover have ?thesis when f_nconst:"\\<^sub>Fw in (at z). f w\0" proof - have "\\<^sub>F w in at z. f w \ 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] . then obtain d1 where d1:"d1>0" "\x. x \ z \ dist x z < d1 \ f x \ 0" unfolding eventually_at by auto obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}" using f_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto moreover have "vf analytic_on ball z d3 - {z}" unfolding vf_def apply (rule analytic_on_inverse) subgoal using d2 unfolding d3_def by (elim analytic_on_subset) auto subgoal for w using d1 unfolding d3_def by (auto simp add:dist_commute) done ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto qed ultimately show ?thesis by auto qed lemma not_essential_divide[singularity_intros]: assumes f_ness:"not_essential f z" and g_ness:"not_essential g z" assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "not_essential (\w. f w / g w) z" proof - have "not_essential (\w. f w * inverse (g w)) z" apply (rule not_essential_times[where g="\w. inverse (g w)"]) using assms by (auto intro: isolated_singularity_at_inverse not_essential_inverse) then show ?thesis by (simp add:field_simps) qed -lemma +lemma assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows isolated_singularity_at_times[singularity_intros]: "isolated_singularity_at (\w. f w * g w) z" and isolated_singularity_at_add[singularity_intros]: "isolated_singularity_at (\w. f w + g w) z" proof - - obtain d1 d2 where "d1>0" "d2>0" + obtain d1 d2 where "d1>0" "d2>0" and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}" using f_iso g_iso unfolding isolated_singularity_at_def by auto define d3 where "d3=min d1 d2" have "d3>0" unfolding d3_def using \d1>0\ \d2>0\ by auto - + have "(\w. f w * g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_mult) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) - then show "isolated_singularity_at (\w. f w * g w) z" + then show "isolated_singularity_at (\w. f w * g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto have "(\w. f w + g w) analytic_on ball z d3 - {z}" apply (rule analytic_on_add) using d1 d2 unfolding d3_def by (auto elim:analytic_on_subset) - then show "isolated_singularity_at (\w. f w + g w) z" + then show "isolated_singularity_at (\w. f w + g w) z" using \d3>0\ unfolding isolated_singularity_at_def by auto qed lemma isolated_singularity_at_uminus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" shows "isolated_singularity_at (\w. - f w) z" using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast lemma isolated_singularity_at_id[singularity_intros]: "isolated_singularity_at (\w. w) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_minus[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" shows "isolated_singularity_at (\w. f w - g w) z" using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\w. - g w"] ,OF g_iso] by simp lemma isolated_singularity_at_divide[singularity_intros]: assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" and g_ness:"not_essential g z" shows "isolated_singularity_at (\w. f w / g w) z" using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso, of "\w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps) lemma isolated_singularity_at_const[singularity_intros]: "isolated_singularity_at (\w. c) z" unfolding isolated_singularity_at_def by (simp add: gt_ex) lemma isolated_singularity_at_holomorphic: assumes "f holomorphic_on s-{z}" "open s" "z\s" shows "isolated_singularity_at f z" - using assms unfolding isolated_singularity_at_def + using assms unfolding isolated_singularity_at_def by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ definition\<^marker>\tag important\ zorder :: "(complex \ complex) \ complex \ int" where "zorder f z = (THE n. (\h r. r>0 \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. f w = h w * (w-z) powr (of_int n) \ h w \0)))" definition\<^marker>\tag important\ zor_poly ::"[complex \ complex, complex] \ complex \ complex" where "zor_poly f z = (SOME h. \r. r > 0 \ h holomorphic_on cball z r \ h z \ 0 \ (\w\cball z r - {z}. f w = h w * (w - z) powr (zorder f z) \ h w \0))" lemma zorder_exist: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" - assumes f_iso:"isolated_singularity_at f z" - and f_ness:"not_essential f z" + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows "g z\0 \ (\r. r>0 \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w-z) powr n \ g w \0))" proof - define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" - have "\!n. \g r. P n g r" + have "\!n. \g r. P n g r" using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto then have "\g r. P n g r" unfolding n_def P_def zorder_def by (drule_tac theI',argo) then have "\r. P n g r" unfolding P_def zor_poly_def g_def n_def by (drule_tac someI_ex,argo) then obtain r1 where "P n g r1" by auto then show ?thesis unfolding P_def by auto qed -lemma +lemma fixes f::"complex \ complex" and z::complex - assumes f_iso:"isolated_singularity_at f z" - and f_ness:"not_essential f z" + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" and f_nconst:"\\<^sub>Fw in (at z). f w\0" shows zorder_inverse: "zorder (\w. inverse (f w)) z = - zorder f z" - and zor_poly_inverse: "\\<^sub>Fw in (at z). zor_poly (\w. inverse (f w)) z w + and zor_poly_inverse: "\\<^sub>Fw in (at z). zor_poly (\w. inverse (f w)) z w = inverse (zor_poly f z w)" proof - define vf where "vf = (\w. inverse (f w))" - define fn vfn where + define fn vfn where "fn = zorder f z" and "vfn = zorder vf z" - define fp vfp where + define fp vfp where "fp = zor_poly f z" and "vfp = zor_poly vf z" obtain fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" + and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def] by auto - have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" + have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" and fr_nz: "inverse (fp w)\0" when "w\ball z fr - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that by auto then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\0" unfolding vf_def by (auto simp add:powr_minus) qed - obtain vfr where [simp]:"vfp z \ 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" + obtain vfr where [simp]:"vfp z \ 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr" "(\w\cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0)" proof - - have "isolated_singularity_at vf z" - using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . - moreover have "not_essential vf z" + have "isolated_singularity_at vf z" + using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def . + moreover have "not_essential vf z" using not_essential_inverse[OF f_ness f_iso] unfolding vf_def . - moreover have "\\<^sub>F w in at z. vf w \ 0" + moreover have "\\<^sub>F w in at z. vf w \ 0" using f_nconst unfolding vf_def by (auto elim:frequently_elim1) ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto qed define r1 where "r1 = min fr vfr" have "r1>0" using \fr>0\ \vfr>0\ unfolding r1_def by simp show "vfn = - fn" apply (rule holomorphic_factor_unique[of r1 vfp z "\w. inverse (fp w)" vf]) subgoal using \r1>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r1 - {z}" then have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" unfolding r1_def by auto - from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] - show "vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0 + from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] + show "vf w = vfp w * (w - z) powr of_int vfn \ vfp w \ 0 \ vf w = inverse (fp w) * (w - z) powr of_int (- fn) \ inverse (fp w) \ 0" by auto qed - subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) + subgoal using vfr(1) unfolding r1_def by (auto intro!:holomorphic_intros) subgoal using fr unfolding r1_def by (auto intro!:holomorphic_intros) done have "vfp w = inverse (fp w)" when "w\ball z r1-{z}" for w proof - have "w \ ball z fr - {z}" "w \ cball z vfr - {z}" "w\z" using that unfolding r1_def by auto from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \vfn = - fn\ \w\z\ show ?thesis by auto qed then show "\\<^sub>Fw in (at z). vfp w = inverse (fp w)" unfolding eventually_at using \r1>0\ apply (rule_tac x=r1 in exI) by (auto simp add:dist_commute) qed -lemma +lemma fixes f g::"complex \ complex" and z::complex - assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" - and f_ness:"not_essential f z" and g_ness:"not_essential g z" + assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" + and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_times:"zorder (\w. f w * g w) z = zorder f z + zorder g z" and - zor_poly_times:"\\<^sub>Fw in (at z). zor_poly (\w. f w * g w) z w + zor_poly_times:"\\<^sub>Fw in (at z). zor_poly (\w. f w * g w) z w = zor_poly f z w *zor_poly g z w" proof - define fg where "fg = (\w. f w * g w)" - define fn gn fgn where + define fn gn fgn where "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z" - define fp gp fgp where + define fp gp fgp where "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z" have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) - obtain fr where [simp]:"fp z \ 0" and "fr > 0" - and fr: "fp holomorphic_on cball z fr" + obtain fr where [simp]:"fp z \ 0" and "fr > 0" + and fr: "fp holomorphic_on cball z fr" "\w\cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \ fp w \ 0" using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto - obtain gr where [simp]:"gp z \ 0" and "gr > 0" - and gr: "gp holomorphic_on cball z gr" + obtain gr where [simp]:"gp z \ 0" and "gr > 0" + and gr: "gp holomorphic_on cball z gr" "\w\cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \ gp w \ 0" using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto define r1 where "r1=min fr gr" have "r1>0" unfolding r1_def using \fr>0\ \gr>0\ by auto have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\0" when "w\ball z r1 - {z}" for w proof - have "f w = fp w * (w - z) powr of_int fn" "fp w\0" using fr(2)[rule_format,of w] that unfolding r1_def by auto moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \ 0" using gr(2)[rule_format, of w] that unfolding r1_def by auto ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\0" unfolding fg_def by (auto simp add:powr_add) qed obtain fgr where [simp]:"fgp z \ 0" and "fgr > 0" - and fgr: "fgp holomorphic_on cball z fgr" + and fgr: "fgp holomorphic_on cball z fgr" "\w\cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0" proof - - have "fgp z \ 0 \ (\r>0. fgp holomorphic_on cball z r + have "fgp z \ 0 \ (\r>0. fgp holomorphic_on cball z r \ (\w\cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0))" apply (rule zorder_exist[of fg z, folded fgn_def fgp_def]) subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] . subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] . subgoal unfolding fg_def using fg_nconst . done then show ?thesis using that by blast qed define r2 where "r2 = min fgr r1" have "r2>0" using \r1>0\ \fgr>0\ unfolding r2_def by simp show "fgn = fn + gn " apply (rule holomorphic_factor_unique[of r2 fgp z "\w. fp w * gp w" fg]) subgoal using \r2>0\ by simp subgoal by simp subgoal by simp subgoal proof (rule ballI) fix w assume "w \ ball z r2 - {z}" then have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" unfolding r2_def by auto - from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] - show "fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0 + from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] + show "fg w = fgp w * (w - z) powr of_int fgn \ fgp w \ 0 \ fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \ fp w * gp w \ 0" by auto qed - subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) + subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros) done have "fgp w = fp w *gp w" when "w\ball z r2-{z}" for w proof - have "w \ ball z r1 - {z}" "w \ cball z fgr - {z}" "w\z" using that unfolding r2_def by auto from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \fgn = fn + gn\ \w\z\ show ?thesis by auto qed - then show "\\<^sub>Fw in (at z). fgp w = fp w * gp w" + then show "\\<^sub>Fw in (at z). fgp w = fp w * gp w" using \r2>0\ unfolding eventually_at by (auto simp add:dist_commute) qed -lemma +lemma fixes f g::"complex \ complex" and z::complex - assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" - and f_ness:"not_essential f z" and g_ness:"not_essential g z" + assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z" + and f_ness:"not_essential f z" and g_ness:"not_essential g z" and fg_nconst: "\\<^sub>Fw in (at z). f w * g w\ 0" shows zorder_divide:"zorder (\w. f w / g w) z = zorder f z - zorder g z" and - zor_poly_divide:"\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w + zor_poly_divide:"\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" proof - have f_nconst:"\\<^sub>Fw in (at z). f w \ 0" and g_nconst:"\\<^sub>Fw in (at z).g w\ 0" using fg_nconst by (auto elim!:frequently_elim1) define vg where "vg=(\w. inverse (g w))" have "zorder (\w. f w * vg w) z = zorder f z + zorder vg z" apply (rule zorder_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "zorder (\w. f w / g w) z = zorder f z - zorder g z" - using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def + using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def by (auto simp add:field_simps) have "\\<^sub>F w in at z. zor_poly (\w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w" apply (rule zor_poly_times[OF f_iso _ f_ness,of vg]) subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] . subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] . subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1) done then show "\\<^sub>Fw in (at z). zor_poly (\w. f w / g w) z w = zor_poly f z w / zor_poly g z w" using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def apply eventually_elim by (auto simp add:field_simps) qed lemma zorder_exist_zero: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" - assumes holo: "f holomorphic_on s" and + assumes holo: "f holomorphic_on s" and "open s" "connected s" "z\s" and non_const: "\w\s. f w \ 0" shows "(if f z=0 then n > 0 else n=0) \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r. f w = g w * (w-z) ^ nat n \ g w \0))" proof - - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" + obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r + have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,6) by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE) - show "not_essential f z" unfolding not_essential_def - using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on + show "not_essential f z" unfolding not_essential_def + using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce have "\\<^sub>F w in at z. f w \ 0 \ w\s" proof - obtain w where "w\s" "f w\0" using non_const by auto - then show ?thesis + then show ?thesis by (rule non_zero_neighbour_alt[OF holo \open s\ \connected s\ \z\s\]) qed then show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto - obtain r2 where r2: "r2>0" "cball z r2 \ s" + obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,6) open_contains_cball_eq by blast define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto - moreover have "g holomorphic_on cball z r3" + moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto - moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto - ultimately show ?thesis using that[of r3] \g z\0\ by auto + ultimately show ?thesis using that[of r3] \g z\0\ by auto qed - have if_0:"if f z=0 then n > 0 else n=0" + have if_0:"if f z=0 then n > 0 else n=0" proof - have "f\ z \ f z" by (metis assms(4,6,7) at_within_open continuous_on holo holomorphic_on_imp_continuous_on) then have "(\w. g w * (w - z) powr of_int n) \z\ f z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto moreover have "g \z\g z" - by (metis (mono_tags, lifting) open_ball at_within_open_subset + by (metis (mono_tags, lifting) open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1,3) subsetCE) ultimately have "(\w. (g w * (w - z) powr of_int n) / g w) \z\ f z/g z" apply (rule_tac tendsto_divide) using \g z\0\ by auto then have powr_tendsto:"(\w. (w - z) powr of_int n) \z\ f z/g z" apply (elim Lim_transform_within_open[where s="ball z r"]) using r by auto - have ?thesis when "n\0" "f z=0" + have ?thesis when "n\0" "f z=0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" - using powr_tendsto + using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) then have *:"(\w. (w - z) ^ nat n) \z\ 0" using \f z=0\ by simp moreover have False when "n=0" proof - have "(\w. (w - z) ^ nat n) \z\ 1" using \n=0\ by auto then show False using * using LIM_unique zero_neq_one by blast qed ultimately show ?thesis using that by fastforce qed - moreover have ?thesis when "n\0" "f z\0" + moreover have ?thesis when "n\0" "f z\0" proof - have False when "n>0" proof - have "(\w. (w - z) ^ nat n) \z\ f z/g z" - using powr_tendsto + using powr_tendsto apply (elim Lim_transform_within[where d=r]) by (auto simp add: powr_of_int \n\0\ \r>0\) moreover have "(\w. (w - z) ^ nat n) \z\ 0" using \n>0\ by (auto intro!:tendsto_eq_intros) ultimately show False using \f z\0\ \g z\0\ using LIM_unique divide_eq_0_iff by blast qed then show ?thesis using that by force qed moreover have False when "n<0" proof - have "(\w. inverse ((w - z) ^ nat (- n))) \z\ f z/g z" "(\w.((w - z) ^ nat (- n))) \z\ 0" subgoal using powr_tendsto powr_of_int that by (elim Lim_transform_within_open[where s=UNIV],auto) subgoal using that by (auto intro!:tendsto_eq_intros) done - from tendsto_mult[OF this,simplified] + from tendsto_mult[OF this,simplified] have "(\x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \z\ 0" . - then have "(\x. 1::complex) \z\ 0" + then have "(\x. 1::complex) \z\ 0" by (elim Lim_transform_within_open[where s=UNIV],auto) then show False using LIM_const_eq by fastforce qed ultimately show ?thesis by fastforce qed moreover have "f w = g w * (w-z) ^ nat n \ g w \0" when "w\cball z r" for w proof (cases "w=z") case True - then have "f \z\f w" + then have "f \z\f w" using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce then have "(\w. g w * (w-z) ^ nat n) \z\f w" proof (elim Lim_transform_within[OF _ \r>0\]) fix x assume "0 < dist x z" "dist x z < r" then have "x \ cball z r - {z}" "x\z" unfolding cball_def by (auto simp add: dist_commute) then have "f x = g x * (x - z) powr of_int n" using r(4)[rule_format,of x] by simp also have "... = g x * (x - z) ^ nat n" apply (subst powr_of_int) using if_0 \x\z\ by (auto split:if_splits) finally show "f x = g x * (x - z) ^ nat n" . qed moreover have "(\w. g w * (w-z) ^ nat n) \z\ g w * (w-z) ^ nat n" using True apply (auto intro!:tendsto_eq_intros) - by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball + by (metis open_ball at_within_open_subset ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on r(1) r(3) that) ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast then show ?thesis using \g z\0\ True by auto next case False then have "f w = g w * (w - z) powr of_int n \ g w \ 0" using r(4) that by auto then show ?thesis using False if_0 powr_of_int by (auto split:if_splits) qed ultimately show ?thesis using r by auto qed lemma zorder_exist_pole: fixes f::"complex \ complex" and z::complex defines "n\zorder f z" and "g\zor_poly f z" - assumes holo: "f holomorphic_on s-{z}" and + assumes holo: "f holomorphic_on s-{z}" and "open s" "z\s" and "is_pole f z" shows "n < 0 \ g z\0 \ (\r. r>0 \ cball z r \ s \ g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0))" proof - - obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" + obtain r where "g z \ 0" and r: "r>0" "cball z r \ s" "g holomorphic_on cball z r" "(\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" proof - - have "g z \ 0 \ (\r>0. g holomorphic_on cball z r + have "g z \ 0 \ (\r>0. g holomorphic_on cball z r \ (\w\cball z r - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0))" proof (rule zorder_exist[of f z,folded g_def n_def]) show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using holo assms(4,5) by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) - show "not_essential f z" unfolding not_essential_def - using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on + show "not_essential f z" unfolding not_essential_def + using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on by fastforce from non_zero_neighbour_pole[OF \is_pole f z\] show "\\<^sub>F w in at z. f w \ 0" apply (elim eventually_frequentlyE) by auto qed then obtain r1 where "g z \ 0" "r1>0" and r1:"g holomorphic_on cball z r1" "(\w\cball z r1 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" by auto - obtain r2 where r2: "r2>0" "cball z r2 \ s" + obtain r2 where r2: "r2>0" "cball z r2 \ s" using assms(4,5) open_contains_cball_eq by metis define r3 where "r3=min r1 r2" have "r3>0" "cball z r3 \ s" using \r1>0\ r2 unfolding r3_def by auto - moreover have "g holomorphic_on cball z r3" + moreover have "g holomorphic_on cball z r3" using r1(1) unfolding r3_def by auto - moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" + moreover have "(\w\cball z r3 - {z}. f w = g w * (w - z) powr of_int n \ g w \ 0)" using r1(2) unfolding r3_def by auto - ultimately show ?thesis using that[of r3] \g z\0\ by auto + ultimately show ?thesis using that[of r3] \g z\0\ by auto qed have "n<0" proof (rule ccontr) assume " \ n < 0" define c where "c=(if n=0 then g z else 0)" - have [simp]:"g \z\ g z" - by (metis open_ball at_within_open ball_subset_cball centre_in_ball + have [simp]:"g \z\ g z" + by (metis open_ball at_within_open ball_subset_cball centre_in_ball continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) ) have "\\<^sub>F x in at z. f x = g x * (x - z) ^ nat n" unfolding eventually_at_topological apply (rule_tac exI[where x="ball z r"]) using r powr_of_int \\ n < 0\ by auto moreover have "(\x. g x * (x - z) ^ nat n) \z\c" proof (cases "n=0") case True then show ?thesis unfolding c_def by simp next case False then have "(\x. (x - z) ^ nat n) \z\ 0" using \\ n < 0\ by (auto intro!:tendsto_eq_intros) - from tendsto_mult[OF _ this,of g "g z",simplified] + from tendsto_mult[OF _ this,of g "g z",simplified] show ?thesis unfolding c_def using False by simp qed ultimately have "f \z\c" using tendsto_cong by fast - then show False using \is_pole f z\ at_neq_bot not_tendsto_and_filterlim_at_infinity + then show False using \is_pole f z\ at_neq_bot not_tendsto_and_filterlim_at_infinity unfolding is_pole_def by blast qed moreover have "\w\cball z r - {z}. f w = g w / (w-z) ^ nat (- n) \ g w \0" - using r(4) \n<0\ powr_of_int + using r(4) \n<0\ powr_of_int by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le) ultimately show ?thesis using r(1-3) \g z\0\ by auto qed lemma zorder_eqI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes fg_eq:"\w. \w \ s;w\z\ \ f w = g w * (w - z) powr n" shows "zorder f z = n" proof - have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact moreover have "open (-{0::complex})" by auto ultimately have "open ((g -` (-{0})) \ s)" unfolding continuous_on_open_vimage[OF \open s\] by blast moreover from assms have "z \ (g -` (-{0})) \ s" by auto ultimately obtain r where r: "r > 0" "cball z r \ s \ (g -` (-{0}))" unfolding open_contains_cball by blast let ?gg= "(\w. g w * (w - z) powr n)" define P where "P = (\n g r. 0 < r \ g holomorphic_on cball z r \ g z\0 \ (\w\cball z r - {z}. f w = g w * (w-z) powr (of_int n) \ g w\0))" have "P n g r" unfolding P_def using r assms(3,4,5) by auto then have "\g r. P n g r" by auto moreover have unique: "\!n. \g r. P n g r" unfolding P_def proof (rule holomorphic_factor_puncture) have "ball z r-{z} \ s" using r using ball_subset_cball by blast then have "?gg holomorphic_on ball z r-{z}" using \g holomorphic_on s\ r by (auto intro!: holomorphic_intros) then have "f holomorphic_on ball z r - {z}" apply (elim holomorphic_transform) using fg_eq \ball z r-{z} \ s\ by auto then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def using analytic_on_open open_delete r(1) by blast next have "not_essential ?gg z" proof (intro singularity_intros) - show "not_essential g z" - by (meson \continuous_on s g\ assms(1) assms(2) continuous_on_eq_continuous_at + show "not_essential g z" + by (meson \continuous_on s g\ assms(1) assms(2) continuous_on_eq_continuous_at isCont_def not_essential_def) show " \\<^sub>F w in at z. w - z \ 0" by (simp add: eventually_at_filter) - then show "LIM w at z. w - z :> at 0" + then show "LIM w at z. w - z :> at 0" unfolding filterlim_at by (auto intro:tendsto_eq_intros) - show "isolated_singularity_at g z" - by (meson Diff_subset open_ball analytic_on_holomorphic + show "isolated_singularity_at g z" + by (meson Diff_subset open_ball analytic_on_holomorphic assms(1,2,3) holomorphic_on_subset isolated_singularity_at_def openE) qed then show "not_essential f z" apply (elim not_essential_transform) - unfolding eventually_at using assms(1,2) assms(5)[symmetric] + unfolding eventually_at using assms(1,2) assms(5)[symmetric] by (metis dist_commute mem_ball openE subsetCE) - show "\\<^sub>F w in at z. f w \ 0" unfolding frequently_at + show "\\<^sub>F w in at z. f w \ 0" unfolding frequently_at proof (rule,rule) fix d::real assume "0 < d" define z' where "z'=z+min d r / 2" have "z' \ z" " dist z' z < d " - unfolding z'_def using \d>0\ \r>0\ + unfolding z'_def using \d>0\ \r>0\ by (auto simp add:dist_norm) - moreover have "f z' \ 0" + moreover have "f z' \ 0" proof (subst fg_eq[OF _ \z'\z\]) have "z' \ cball z r" unfolding z'_def using \r>0\ \d>0\ by (auto simp add:dist_norm) then show " z' \ s" using r(2) by blast - show "g z' * (z' - z) powr of_int n \ 0" + show "g z' * (z' - z) powr of_int n \ 0" using P_def \P n g r\ \z' \ cball z r\ calculation(1) by auto qed ultimately show "\x\UNIV. x \ z \ dist x z < d \ f x \ 0" by auto qed qed ultimately have "(THE n. \g r. P n g r) = n" by (rule_tac the1_equality) then show ?thesis unfolding zorder_def P_def by blast qed lemma residue_pole_order: fixes f::"complex \ complex" and z::complex defines "n \ nat (- zorder f z)" and "h \ zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and pole:"is_pole f z" shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ ball z e" and h_holo: "h holomorphic_on cball z r" and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" proof - - obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" + obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" "(\w\cball z r - {z}. f w = h w / (w - z) ^ n \ h w \ 0)" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\,folded n_def h_def] by auto have "n>0" using \zorder f z < 0\ unfolding n_def by simp moreover have "(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" using \h z\0\ r(6) by blast ultimately show ?thesis using r(3,4,5) that by blast qed have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" using h_divide by simp define c where "c \ 2 * pi * \" define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" define h' where "h' \ \u. h u / (u - z) ^ n" have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" unfolding h'_def proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", folded c_def Suc_pred'[OF \n>0\]]) show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp show "h holomorphic_on ball z r" using h_holo by auto show " z \ ball z r" using \r>0\ by auto qed then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto then have "(f has_contour_integral c * der_f) (circlepath z r)" proof (elim has_contour_integral_eq) fix x assume "x \ path_image (circlepath z r)" hence "x\cball z r - {z}" using \r>0\ by auto then show "h' x = f x" using h_divide unfolding h'_def by auto qed moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" - using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] + using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] unfolding c_def by simp ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast hence "der_f = residue f z" unfolding c_def by auto thus ?thesis unfolding der_f_def by auto qed lemma simple_zeroI: assumes "open s" "z \ s" "g holomorphic_on s" "g z \ 0" assumes "\w. w \ s \ f w = g w * (w - z)" shows "zorder f z = 1" using assms(1-4) by (rule zorder_eqI) (use assms(5) in auto) lemma higher_deriv_power: - shows "(deriv ^^ j) (\w. (w - z) ^ n) w = + shows "(deriv ^^ j) (\w. (w - z) ^ n) w = pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)" proof (induction j arbitrary: w) case 0 thus ?case by auto next case (Suc j w) have "(deriv ^^ Suc j) (\w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\w. (w - z) ^ n)) w" by simp - also have "(deriv ^^ j) (\w. (w - z) ^ n) = + also have "(deriv ^^ j) (\w. (w - z) ^ n) = (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))" using Suc by (intro Suc.IH ext) also { have "(\ has_field_derivative of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)" using Suc.prems by (auto intro!: derivative_eq_intros) - also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = + also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j = pochhammer (of_nat (Suc n - Suc j)) (Suc j)" - by (cases "Suc j \ n", subst pochhammer_rec) + by (cases "Suc j \ n", subst pochhammer_rec) (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left) finally have "deriv (\w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w = \ * (w - z) ^ (n - Suc j)" by (rule DERIV_imp_deriv) } finally show ?case . qed lemma zorder_zero_eqI: assumes f_holo:"f holomorphic_on s" and "open s" "z \ s" assumes zero: "\i. i < nat n \ (deriv ^^ i) f z = 0" assumes nz: "(deriv ^^ nat n) f z \ 0" and "n\0" shows "zorder f z = n" proof - obtain r where [simp]:"r>0" and "ball z r \ s" using \open s\ \z\s\ openE by blast have nz':"\w\ball z r. f w \ 0" proof (rule ccontr) assume "\ (\w\ball z r. f w \ 0)" then have "eventually (\u. f u = 0) (nhds z)" - using \r>0\ unfolding eventually_nhds + using \r>0\ unfolding eventually_nhds apply (rule_tac x="ball z r" in exI) by auto then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\_. 0) z" by (intro higher_deriv_cong_ev) auto also have "(deriv ^^ nat n) (\_. 0) z = 0" by (induction n) simp_all finally show False using nz by contradiction qed define zn g where "zn = zorder f z" and "g = zor_poly f z" obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and [simp]:"e>0" and "cball z e \ ball z r" and g_holo:"g holomorphic_on cball z e" and e_fac:"(\w\cball z e. f w = g w * (w - z) ^ nat zn \ g w \ 0)" proof - have "f holomorphic_on ball z r" using f_holo \ball z r \ s\ by auto from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def] show ?thesis by blast qed from this(1,2,5) have "zn\0" "g z\0" - subgoal by (auto split:if_splits) + subgoal by (auto split:if_splits) subgoal using \0 < e\ ball_subset_cball centre_in_ball e_fac by blast done define A where "A = (\i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)" have deriv_A:"(deriv ^^ i) f z = (if zn \ int i then A i else 0)" for i proof - have "eventually (\w. w \ ball z e) (nhds z)" using \cball z e \ ball z r\ \e>0\ by (intro eventually_nhds_in_open) auto hence "eventually (\w. f w = (w - z) ^ (nat zn) * g w) (nhds z)" - apply eventually_elim + apply eventually_elim by (use e_fac in auto) hence "(deriv ^^ i) f z = (deriv ^^ i) (\w. (w - z) ^ nat zn * g w) z" by (intro higher_deriv_cong_ev) auto also have "\ = (\j=0..i. of_nat (i choose j) * (deriv ^^ j) (\w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)" - using g_holo \e>0\ + using g_holo \e>0\ by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros) - also have "\ = (\j=0..i. if j = nat zn then + also have "\ = (\j=0..i. if j = nat zn then of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" proof (intro sum.cong refl, goal_cases) case (1 j) - have "(deriv ^^ j) (\w. (w - z) ^ nat zn) z = + have "(deriv ^^ j) (\w. (w - z) ^ nat zn) z = pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)" by (subst higher_deriv_power) auto also have "\ = (if j = nat zn then fact j else 0)" by (auto simp: not_less pochhammer_0_left pochhammer_fact) - also have "of_nat (i choose j) * \ * (deriv ^^ (i - j)) g z = - (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) + also have "of_nat (i choose j) * \ * (deriv ^^ (i - j)) g z = + (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)" by simp finally show ?case . qed also have "\ = (if i \ zn then A i else 0)" by (auto simp: A_def) finally show "(deriv ^^ i) f z = \" . qed have False when "nn\0\ by auto + using deriv_A[of "nat n"] that \n\0\ by auto with nz show False by auto qed moreover have "n\zn" proof - - have "g z \ 0" using e_fac[rule_format,of z] \e>0\ by simp + have "g z \ 0" using e_fac[rule_format,of z] \e>0\ by simp then have "(deriv ^^ nat zn) f z \ 0" using deriv_A[of "nat zn"] by(auto simp add:A_def) then have "nat zn \ nat n" using zero[of "nat zn"] by linarith moreover have "zn\0" using e_if by (auto split:if_splits) ultimately show ?thesis using nat_le_eq_zle by blast qed ultimately show ?thesis unfolding zn_def by fastforce qed -lemma +lemma assumes "eventually (\z. f z = g z) (at z)" "z = z'" shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'" proof - define P where "P = (\ff n h r. 0 < r \ h holomorphic_on cball z r \ h z\0 \ (\w\cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \ h w\0))" - have "(\r. P f n h r) = (\r. P g n h r)" for n h + have "(\r. P f n h r) = (\r. P g n h r)" for n h proof - - have *: "\r. P g n h r" if "\r. P f n h r" and "eventually (\x. f x = g x) (at z)" for f g + have *: "\r. P g n h r" if "\r. P f n h r" and "eventually (\x. f x = g x) (at z)" for f g proof - from that(1) obtain r1 where r1_P:"P f n h r1" by auto from that(2) obtain r2 where "r2>0" and r2_dist:"\x. x \ z \ dist x z \ r2 \ f x = g x" unfolding eventually_at_le by auto define r where "r=min r1 r2" have "r>0" "h z\0" using r1_P \r2>0\ unfolding r_def P_def by auto moreover have "h holomorphic_on cball z r" using r1_P unfolding P_def r_def by auto moreover have "g w = h w * (w - z) powr of_int n \ h w \ 0" when "w\cball z r - {z}" for w proof - have "f w = h w * (w - z) powr of_int n \ h w \ 0" using r1_P that unfolding P_def r_def by auto - moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def - by (simp add: dist_commute) + moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def + by (simp add: dist_commute) ultimately show ?thesis by simp qed ultimately show ?thesis unfolding P_def by auto qed from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) show ?thesis by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']]) qed - then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'" + then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'" using \z=z'\ unfolding P_def zorder_def zor_poly_def by auto qed lemma zorder_nonzero_div_power: assumes "open s" "z \ s" "f holomorphic_on s" "f z \ 0" "n > 0" shows "zorder (\w. f w / (w - z) ^ n) z = - n" apply (rule zorder_eqI[OF \open s\ \z\s\ \f holomorphic_on s\ \f z\0\]) apply (subst powr_of_int) using \n>0\ by (auto simp add:field_simps) lemma zor_poly_eq: assumes "isolated_singularity_at f z" "not_essential f z" "\\<^sub>F w in at z. f w \ 0" shows "eventually (\w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)" proof - - obtain r where r:"r>0" + obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))" using zorder_exist[OF assms] by blast - then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" + then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_zero_eq: assumes "f holomorphic_on s" "open s" "connected s" "z \ s" "\w\s. f w \ 0" shows "eventually (\w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)" proof - - obtain r where r:"r>0" + obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))" using zorder_exist_zero[OF assms] by auto - then have *: "\w\ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" + then have *: "\w\ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)" by (auto simp: field_simps powr_minus) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_pole_eq: assumes f_iso:"isolated_singularity_at f z" "is_pole f z" shows "eventually (\w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)" proof - obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast - obtain r where r:"r>0" + obtain r where r:"r>0" "(\w\cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\] by auto - then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" + then have *: "\w\ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)" by (auto simp: field_simps) have "eventually (\w. w \ ball z r - {z}) (at z)" using r eventually_at_ball'[of r z UNIV] by auto thus ?thesis by eventually_elim (insert *, auto) qed lemma zor_poly_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "isolated_singularity_at f z0" "not_essential f z0" "\\<^sub>F w in at z0. f w \ 0" assumes lim: "((\x. f (g x) * (g x - z0) powr - n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist[OF assms(2-4)] obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r - {z0} \ f w = zor_poly f z0 w * (w - z0) powr n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)" by eventually_elim (insert r, auto simp: field_simps powr_minus) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) powr - n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) powr - n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_zero_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes "f holomorphic_on A" "open A" "connected A" "z0 \ A" "\z\A. f z \ 0" assumes lim: "((\x. f (g x) / (g x - z0) ^ nat n) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - from zorder_exist_zero[OF assms(2-6)] obtain r where r: "r > 0" "cball z0 r \ A" "zor_poly f z0 holomorphic_on cball z0 r" "\w. w \ cball z0 r \ f w = zor_poly f z0 w * (w - z0) ^ nat n" unfolding n_def by blast from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto hence "eventually (\w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)" by eventually_elim (insert r, auto simp: field_simps) moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w / (w - z0) ^ nat n) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) / (g x - z0) ^ nat n) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma zor_poly_pole_eqI: fixes f :: "complex \ complex" and z0 :: complex defines "n \ zorder f z0" assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0" assumes lim: "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ c) F" assumes g: "filterlim g (at z0) F" and "F \ bot" shows "zor_poly f z0 z0 = c" proof - obtain r where r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r" - proof - - have "\\<^sub>F w in at z0. f w \ 0" + proof - + have "\\<^sub>F w in at z0. f w \ 0" using non_zero_neighbour_pole[OF \is_pole f z0\] by (auto elim:eventually_frequentlyE) moreover have "not_essential f z0" unfolding not_essential_def using \is_pole f z0\ by simp ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto qed from r(1) have "eventually (\w. w \ ball z0 r \ w \ z0) (at z0)" using eventually_at_ball'[of r z0 UNIV] by auto have "eventually (\w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)" using zor_poly_pole_eq[OF f_iso \is_pole f z0\] unfolding n_def . moreover have "continuous_on (ball z0 r) (zor_poly f z0)" using r by (intro holomorphic_on_imp_continuous_on) auto with r(1,2) have "isCont (zor_poly f z0) z0" by (auto simp: continuous_on_eq_continuous_at) hence "(zor_poly f z0 \ zor_poly f z0 z0) (at z0)" unfolding isCont_def . ultimately have "((\w. f w * (w - z0) ^ nat (-n)) \ zor_poly f z0 z0) (at z0)" by (blast intro: Lim_transform_eventually) hence "((\x. f (g x) * (g x - z0) ^ nat (-n)) \ zor_poly f z0 z0) F" by (rule filterlim_compose[OF _ g]) from tendsto_unique[OF \F \ bot\ this lim] show ?thesis . qed lemma residue_simple_pole: - assumes "isolated_singularity_at f z0" + assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" shows "residue f z0 = zor_poly f z0 z0" using assms by (subst residue_pole_order) simp_all lemma residue_simple_pole_limit: - assumes "isolated_singularity_at f z0" + assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" assumes "((\x. f (g x) * (g x - z0)) \ c) F" assumes "filterlim g (at z0) F" "F \ bot" shows "residue f z0 = c" proof - have "residue f z0 = zor_poly f z0 z0" by (rule residue_simple_pole assms)+ also have "\ = c" apply (rule zor_poly_pole_eqI) using assms by auto finally show ?thesis . qed lemma lhopital_complex_simple: - assumes "(f has_field_derivative f') (at z)" + assumes "(f has_field_derivative f') (at z)" assumes "(g has_field_derivative g') (at z)" assumes "f z = 0" "g z = 0" "g' \ 0" "f' / g' = c" shows "((\w. f w / g w) \ c) (at z)" proof - have "eventually (\w. w \ z) (at z)" by (auto simp: eventually_at_filter) hence "eventually (\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z)) = f w / g w) (at z)" by eventually_elim (simp add: assms field_split_simps) moreover have "((\w. ((f w - f z) / (w - z)) / ((g w - g z) / (w - z))) \ f' / g') (at z)" by (intro tendsto_divide has_field_derivativeD assms) ultimately have "((\w. f w / g w) \ f' / g') (at z)" by (blast intro: Lim_transform_eventually) with assms show ?thesis by simp qed lemma - assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" - and "open s" "connected s" "z \ s" + assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" + and "open s" "connected s" "z \ s" assumes g_deriv:"(g has_field_derivative g') (at z)" assumes "f z \ 0" "g z = 0" "g' \ 0" shows porder_simple_pole_deriv: "zorder (\w. f w / g w) z = - 1" and residue_simple_pole_deriv: "residue (\w. f w / g w) z = f z / g'" proof - have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z" using isolated_singularity_at_holomorphic[OF _ \open s\ \z\s\] f_holo g_holo by (meson Diff_subset holomorphic_on_subset)+ have [simp]:"not_essential f z" "not_essential g z" unfolding not_essential_def using f_holo g_holo assms(3,5) by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+ - have g_nconst:"\\<^sub>F w in at z. g w \0 " + have g_nconst:"\\<^sub>F w in at z. g w \0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at z. g w \ 0)" then have "\\<^sub>F w in nhds z. g w = 0" - unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ + unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball) then have "deriv g z = deriv (\_. 0) z" by (intro deriv_cong_ev) auto then have "deriv g z = 0" by auto then have "g' = 0" using g_deriv DERIV_imp_deriv by blast then show False using \g'\0\ by auto qed - + have "zorder (\w. f w / g w) z = zorder f z - zorder g z" proof - - have "\\<^sub>F w in at z. f w \0 \ w\s" + have "\\<^sub>F w in at z. f w \0 \ w\s" apply (rule non_zero_neighbour_alt) using assms by auto - with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" + with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" by (elim frequently_rev_mp eventually_rev_mp,auto) then show ?thesis using zorder_divide[of f z g] by auto qed moreover have "zorder f z=0" apply (rule zorder_zero_eqI[OF f_holo \open s\ \z\s\]) using \f z\0\ by auto moreover have "zorder g z=1" apply (rule zorder_zero_eqI[OF g_holo \open s\ \z\s\]) subgoal using assms(8) by auto subgoal using DERIV_imp_deriv assms(9) g_deriv by auto subgoal by simp done ultimately show "zorder (\w. f w / g w) z = - 1" by auto - + show "residue (\w. f w / g w) z = f z / g'" proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified]) show "zorder (\w. f w / g w) z = - 1" by fact - show "isolated_singularity_at (\w. f w / g w) z" + show "isolated_singularity_at (\w. f w / g w) z" by (auto intro: singularity_intros) - show "is_pole (\w. f w / g w) z" + show "is_pole (\w. f w / g w) z" proof (rule is_pole_divide) - have "\\<^sub>F x in at z. g x \ 0" + have "\\<^sub>F x in at z. g x \ 0" apply (rule non_zero_neighbour) using g_nconst by auto - moreover have "g \z\ 0" + moreover have "g \z\ 0" using DERIV_isCont assms(8) continuous_at g_deriv by force ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp - show "isCont f z" - using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on + show "isCont f z" + using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on by auto show "f z \ 0" by fact qed show "filterlim id (at z) (at z)" by (simp add: filterlim_iff) have "((\w. (f w * (w - z)) / g w) \ f z / g') (at z)" proof (rule lhopital_complex_simple) show "((\w. f w * (w - z)) has_field_derivative f z) (at z)" using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo]) show "(g has_field_derivative g') (at z)" by fact qed (insert assms, auto) then show "((\w. (f w / g w) * (w - z)) \ f z / g') (at z)" by (simp add: field_split_simps) qed qed subsection \The argument principle\ theorem argument_principle: fixes f::"complex \ complex" and poles s:: "complex set" defines "pz \ {w. f w = 0 \ w \ poles}" \ \\<^term>\pz\ is the set of poles and zeros\ assumes "open s" and "connected s" and f_holo:"f holomorphic_on s-poles" and h_holo:"h holomorphic_on s" and "valid_path g" and loop:"pathfinish g = pathstart g" and path_img:"path_image g \ s - pz" and homo:"\z. (z \ s) \ winding_number g z = 0" and finite:"finite pz" and poles:"\p\poles. is_pole f p" shows "contour_integral g (\x. deriv f x * h x / f x) = 2 * pi * \ * (\p\pz. winding_number g p * h p * zorder f p)" (is "?L=?R") proof - define c where "c \ 2 * complex_of_real pi * \ " define ff where "ff \ (\x. deriv f x * h x / f x)" define cont where "cont \ \ff p e. (ff has_contour_integral c * zorder f p * h p ) (circlepath p e)" define avoid where "avoid \ \p e. \w\cball p e. w \ s \ (w \ p \ w \ pz)" have "\e>0. avoid p e \ (p\pz \ cont ff p e)" when "p\s" for p proof - obtain e1 where "e1>0" and e1_avoid:"avoid p e1" using finite_cball_avoid[OF \open s\ finite] \p\s\ unfolding avoid_def by auto have "\e2>0. cball p e2 \ ball p e1 \ cont ff p e2" when "p\pz" proof - define po where "po \ zorder f p" define pp where "pp \ zor_poly f p" define f' where "f' \ \w. pp w * (w - p) powr po" define ff' where "ff' \ (\x. deriv f' x * h x / f' x)" obtain r where "pp p\0" "r>0" and "rw\cball p r-{p}. f w = pp w * (w - p) powr po \ pp w \ 0)" proof - have "isolated_singularity_at f p" proof - have "f holomorphic_on ball p e1 - {p}" apply (intro holomorphic_on_subset[OF f_holo]) using e1_avoid \p\pz\ unfolding avoid_def pz_def by force - then show ?thesis unfolding isolated_singularity_at_def + then show ?thesis unfolding isolated_singularity_at_def using \e1>0\ analytic_on_open open_delete by blast qed - moreover have "not_essential f p" + moreover have "not_essential f p" proof (cases "is_pole f p") case True then show ?thesis unfolding not_essential_def by auto next case False then have "p\s-poles" using \p\s\ poles unfolding pz_def by auto moreover have "open (s-poles)" - using \open s\ + using \open s\ apply (elim open_Diff) apply (rule finite_imp_closed) using finite unfolding pz_def by simp ultimately have "isCont f p" using holomorphic_on_imp_continuous_on[OF f_holo] continuous_on_eq_continuous_at by auto then show ?thesis unfolding isCont_def not_essential_def by auto - qed + qed moreover have "\\<^sub>F w in at p. f w \ 0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at p. f w \ 0)" then have "\\<^sub>F w in at p. f w= 0" unfolding frequently_def by auto then obtain rr where "rr>0" "\w\ball p rr - {p}. f w =0" unfolding eventually_at by (auto simp add:dist_commute) then have "ball p rr - {p} \ {w\ball p rr-{p}. f w=0}" by blast moreover have "infinite (ball p rr - {p})" using \rr>0\ using finite_imp_not_open by fastforce ultimately have "infinite {w\ball p rr-{p}. f w=0}" using infinite_super by blast then have "infinite pz" unfolding pz_def infinite_super by auto then show False using \finite pz\ by auto qed - ultimately obtain r where "pp p \ 0" and r:"r>0" "pp holomorphic_on cball p r" + ultimately obtain r where "pp p \ 0" and r:"r>0" "pp holomorphic_on cball p r" "(\w\cball p r - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" using zorder_exist[of f p,folded po_def pp_def] by auto define r1 where "r1=min r e1 / 2" have "r1e1>0\ \r>0\ by auto - moreover have "r1>0" "pp holomorphic_on cball p r1" + moreover have "r1>0" "pp holomorphic_on cball p r1" "(\w\cball p r1 - {p}. f w = pp w * (w - p) powr of_int po \ pp w \ 0)" unfolding r1_def using \e1>0\ r by auto ultimately show ?thesis using that \pp p\0\ by auto qed - + define e2 where "e2 \ r/2" have "e2>0" using \r>0\ unfolding e2_def by auto define anal where "anal \ \w. deriv pp w * h w / pp w" define prin where "prin \ \w. po * h w / (w - p)" have "((\w. prin w + anal w) has_contour_integral c * po * h p) (circlepath p e2)" proof (rule has_contour_integral_add[of _ _ _ _ 0,simplified]) have "ball p r \ s" using \r avoid_def ball_subset_cball e1_avoid by (simp add: subset_eq) then have "cball p e2 \ s" using \r>0\ unfolding e2_def by auto then have "(\w. po * h w) holomorphic_on cball p e2" using h_holo by (auto intro!: holomorphic_intros) then show "(prin has_contour_integral c * po * h p ) (circlepath p e2)" using Cauchy_integral_circlepath_simple[folded c_def, of "\w. po * h w"] \e2>0\ unfolding prin_def by (auto simp add: mult.assoc) have "anal holomorphic_on ball p r" unfolding anal_def using pp_holo h_holo pp_po \ball p r \ s\ \pp p\0\ by (auto intro!: holomorphic_intros) then show "(anal has_contour_integral 0) (circlepath p e2)" using e2_def \r>0\ by (auto elim!: Cauchy_theorem_disc_simple) qed then have "cont ff' p e2" unfolding cont_def po_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto define wp where "wp \ w-p" have "wp\0" and "pp w \0" unfolding wp_def using \w\p\ \w\ball p r\ pp_po by auto moreover have der_f':"deriv f' w = po * pp w * (w-p) powr (po - 1) + deriv pp w * (w-p) powr po" proof (rule DERIV_imp_deriv) have "(pp has_field_derivative (deriv pp w)) (at w)" using DERIV_deriv_iff_has_field_derivative pp_holo \w\p\ by (meson open_ball \w \ ball p r\ ball_subset_cball holomorphic_derivI holomorphic_on_subset) - then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1) + then show " (f' has_field_derivative of_int po * pp w * (w - p) powr of_int (po - 1) + deriv pp w * (w - p) powr of_int po) (at w)" unfolding f'_def using \w\p\ by (auto intro!: derivative_eq_intros DERIV_cong[OF has_field_derivative_powr_of_int]) qed ultimately show "prin w + anal w = ff' w" unfolding ff'_def prin_def anal_def apply simp apply (unfold f'_def) apply (fold wp_def) apply (auto simp add:field_simps) by (metis (no_types, lifting) diff_add_cancel mult.commute powr_add powr_to_1) qed then have "cont ff p e2" unfolding cont_def proof (elim has_contour_integral_eq) fix w assume "w \ path_image (circlepath p e2)" then have "w\ball p r" and "w\p" unfolding e2_def using \r>0\ by auto have "deriv f' w = deriv f w" proof (rule complex_derivative_transform_within_open[where s="ball p r - {p}"]) show "f' holomorphic_on ball p r - {p}" unfolding f'_def using pp_holo by (auto intro!: holomorphic_intros) next have "ball p e1 - {p} \ s - poles" using ball_subset_cball e1_avoid[unfolded avoid_def] unfolding pz_def by auto - then have "ball p r - {p} \ s - poles" + then have "ball p r - {p} \ s - poles" apply (elim dual_order.trans) using \r by auto then show "f holomorphic_on ball p r - {p}" using f_holo by auto next show "open (ball p r - {p})" by auto show "w \ ball p r - {p}" using \w\ball p r\ \w\p\ by auto next fix x assume "x \ ball p r - {p}" then show "f' x = f x" using pp_po unfolding f'_def by auto qed moreover have " f' w = f w " using \w \ ball p r\ ball_subset_cball subset_iff pp_po \w\p\ unfolding f'_def by auto ultimately show "ff' w = ff w" unfolding ff'_def ff_def by simp qed moreover have "cball p e2 \ ball p e1" using \0 < r\ \r e2_def by auto ultimately show ?thesis using \e2>0\ by auto qed then obtain e2 where e2:"p\pz \ e2>0 \ cball p e2 \ ball p e1 \ cont ff p e2" by auto define e4 where "e4 \ if p\pz then e2 else e1" have "e4>0" using e2 \e1>0\ unfolding e4_def by auto moreover have "avoid p e4" using e2 \e1>0\ e1_avoid unfolding e4_def avoid_def by auto moreover have "p\pz \ cont ff p e4" by (auto simp add: e2 e4_def) ultimately show ?thesis by auto qed then obtain get_e where get_e:"\p\s. get_e p>0 \ avoid p (get_e p) \ (p\pz \ cont ff p (get_e p))" by metis define ci where "ci \ \p. contour_integral (circlepath p (get_e p)) ff" define w where "w \ \p. winding_number g p" have "contour_integral g ff = (\p\pz. w p * ci p)" unfolding ci_def w_def proof (rule Cauchy_theorem_singularities[OF \open s\ \connected s\ finite _ \valid_path g\ loop path_img homo]) have "open (s - pz)" using open_Diff[OF _ finite_imp_closed[OF finite]] \open s\ by auto then show "ff holomorphic_on s - pz" unfolding ff_def using f_holo h_holo by (auto intro!: holomorphic_intros simp add:pz_def) next show "\p\s. 0 < get_e p \ (\w\cball p (get_e p). w \ s \ (w \ p \ w \ pz))" using get_e using avoid_def by blast qed also have "... = (\p\pz. c * w p * h p * zorder f p)" proof (rule sum.cong[of pz pz,simplified]) fix p assume "p \ pz" show "w p * ci p = c * w p * h p * (zorder f p)" proof (cases "p\s") assume "p \ s" have "ci p = c * h p * (zorder f p)" unfolding ci_def apply (rule contour_integral_unique) using get_e \p\s\ \p\pz\ unfolding cont_def by (metis mult.assoc mult.commute) thus ?thesis by auto next assume "p\s" then have "w p=0" using homo unfolding w_def by auto then show ?thesis by auto qed qed also have "... = c*(\p\pz. w p * h p * zorder f p)" unfolding sum_distrib_left by (simp add:algebra_simps) finally have "contour_integral g ff = c * (\p\pz. w p * h p * of_int (zorder f p))" . then show ?thesis unfolding ff_def c_def w_def by simp qed subsection \Rouche's theorem \ theorem Rouche_theorem: fixes f g::"complex \ complex" and s:: "complex set" defines "fg\(\p. f p + g p)" defines "zeros_fg\{p. fg p = 0}" and "zeros_f\{p. f p = 0}" assumes "open s" and "connected s" and "finite zeros_fg" and "finite zeros_f" and f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and "valid_path \" and loop:"pathfinish \ = pathstart \" and path_img:"path_image \ \ s " and path_less:"\z\path_image \. cmod(f z) > cmod(g z)" and homo:"\z. (z \ s) \ winding_number \ z = 0" shows "(\p\zeros_fg. winding_number \ p * zorder fg p) = (\p\zeros_f. winding_number \ p * zorder f p)" proof - have path_fg:"path_image \ \ s - zeros_fg" proof - have False when "z\path_image \" and "f z + g z=0" for z proof - have "cmod (f z) > cmod (g z)" using \z\path_image \\ path_less by auto moreover have "f z = - g z" using \f z + g z =0\ by (simp add: eq_neg_iff_add_eq_0) then have "cmod (f z) = cmod (g z)" by auto ultimately show False by auto qed then show ?thesis unfolding zeros_fg_def fg_def using path_img by auto qed have path_f:"path_image \ \ s - zeros_f" proof - have False when "z\path_image \" and "f z =0" for z proof - have "cmod (g z) < cmod (f z) " using \z\path_image \\ path_less by auto then have "cmod (g z) < 0" using \f z=0\ by auto then show False by auto qed then show ?thesis unfolding zeros_f_def using path_img by auto qed define w where "w \ \p. winding_number \ p" define c where "c \ 2 * complex_of_real pi * \" define h where "h \ \p. g p / f p + 1" obtain spikes where "finite spikes" and spikes: "\x\{0..1} - spikes. \ differentiable at x" using \valid_path \\ by (auto simp: valid_path_def piecewise_C1_differentiable_on_def C1_differentiable_on_eq) have h_contour:"((\x. deriv h x / h x) has_contour_integral 0) \" proof - have outside_img:"0 \ outside (path_image (h o \))" proof - have "h p \ ball 1 1" when "p\path_image \" for p proof - have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) then show ?thesis unfolding h_def by (auto simp add:dist_complex_def) qed then have "path_image (h o \) \ ball 1 1" by (simp add: image_subset_iff path_image_compose) moreover have " (0::complex) \ ball 1 1" by (simp add: dist_norm) ultimately show "?thesis" using convex_in_outside[of "ball 1 1" 0] outside_mono by blast qed have valid_h:"valid_path (h \ \)" proof (rule valid_path_compose_holomorphic[OF \valid_path \\ _ _ path_f]) show "h holomorphic_on s - zeros_f" unfolding h_def using f_holo g_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) next show "open (s - zeros_f)" using \finite zeros_f\ \open s\ finite_imp_closed by auto qed have "((\z. 1/z) has_contour_integral 0) (h \ \)" proof - have "0 \ path_image (h \ \)" using outside_img by (simp add: outside_def) then have "((\z. 1/z) has_contour_integral c * winding_number (h \ \) 0) (h \ \)" using has_contour_integral_winding_number[of "h o \" 0,simplified] valid_h unfolding c_def by auto moreover have "winding_number (h o \) 0 = 0" proof - have "0 \ outside (path_image (h \ \))" using outside_img . moreover have "path (h o \)" using valid_h by (simp add: valid_path_imp_path) moreover have "pathfinish (h o \) = pathstart (h o \)" by (simp add: loop pathfinish_compose pathstart_compose) ultimately show ?thesis using winding_number_zero_in_outside by auto qed ultimately show ?thesis by auto qed moreover have "vector_derivative (h \ \) (at x) = vector_derivative \ (at x) * deriv h (\ x)" when "x\{0..1} - spikes" for x proof (rule vector_derivative_chain_at_general) show "\ differentiable at x" using that \valid_path \\ spikes by auto next define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" define t where "t \ \ x" have "f t\0" unfolding zeros_f_def t_def by (metis DiffD1 image_eqI norm_not_less_zero norm_zero path_defs(4) path_less that) moreover have "t\s" using contra_subsetD path_image_def path_fg t_def that by fastforce ultimately have "(h has_field_derivative der t) (at t)" unfolding h_def der_def using g_holo f_holo \open s\ by (auto intro!: holomorphic_derivI derivative_eq_intros) - then show "h field_differentiable at (\ x)" + then show "h field_differentiable at (\ x)" unfolding t_def field_differentiable_def by blast qed then have " ((/) 1 has_contour_integral 0) (h \ \) = ((\x. deriv h x / h x) has_contour_integral 0) \" unfolding has_contour_integral apply (intro has_integral_spike_eq[OF negligible_finite, OF \finite spikes\]) by auto ultimately show ?thesis by auto qed then have "contour_integral \ (\x. deriv h x / h x) = 0" using contour_integral_unique by simp moreover have "contour_integral \ (\x. deriv fg x / fg x) = contour_integral \ (\x. deriv f x / f x) + contour_integral \ (\p. deriv h p / h p)" proof - have "(\p. deriv f p / f p) contour_integrable_on \" proof (rule contour_integrable_holomorphic_simple[OF _ _ \valid_path \\ path_f]) show "open (s - zeros_f)" using finite_imp_closed[OF \finite zeros_f\] \open s\ by auto then show "(\p. deriv f p / f p) holomorphic_on s - zeros_f" using f_holo by (auto intro!: holomorphic_intros simp add:zeros_f_def) qed moreover have "(\p. deriv h p / h p) contour_integrable_on \" using h_contour by (simp add: has_contour_integral_integrable) ultimately have "contour_integral \ (\x. deriv f x / f x + deriv h x / h x) = contour_integral \ (\p. deriv f p / f p) + contour_integral \ (\p. deriv h p / h p)" using contour_integral_add[of "(\p. deriv f p / f p)" \ "(\p. deriv h p / h p)" ] by auto moreover have "deriv fg p / fg p = deriv f p / f p + deriv h p / h p" when "p\ path_image \" for p proof - have "fg p\0" and "f p\0" using path_f path_fg that unfolding zeros_f_def zeros_fg_def by auto have "h p\0" proof (rule ccontr) assume "\ h p \ 0" then have "g p / f p= -1" unfolding h_def by (simp add: add_eq_0_iff2) then have "cmod (g p/f p) = 1" by auto moreover have "cmod (g p/f p) <1" using path_less[rule_format,OF that] apply (cases "cmod (f p) = 0") by (auto simp add: norm_divide) ultimately show False by auto qed have der_fg:"deriv fg p = deriv f p + deriv g p" unfolding fg_def using f_holo g_holo holomorphic_on_imp_differentiable_at[OF _ \open s\] path_img that by auto have der_h:"deriv h p = (deriv g p * f p - g p * deriv f p)/(f p * f p)" proof - define der where "der \ \p. (deriv g p * f p - g p * deriv f p)/(f p * f p)" have "p\s" using path_img that by auto then have "(h has_field_derivative der p) (at p)" unfolding h_def der_def using g_holo f_holo \open s\ \f p\0\ by (auto intro!: derivative_eq_intros holomorphic_derivI) then show ?thesis unfolding der_def using DERIV_imp_deriv by auto qed show ?thesis apply (simp only:der_fg der_h) apply (auto simp add:field_simps \h p\0\ \f p\0\ \fg p\0\) by (auto simp add:field_simps h_def \f p\0\ fg_def) qed then have "contour_integral \ (\p. deriv fg p / fg p) = contour_integral \ (\p. deriv f p / f p + deriv h p / h p)" by (elim contour_integral_eq) ultimately show ?thesis by auto qed moreover have "contour_integral \ (\x. deriv fg x / fg x) = c * (\p\zeros_fg. w p * zorder fg p)" unfolding c_def zeros_fg_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "fg holomorphic_on s" unfolding fg_def using f_holo g_holo holomorphic_on_add by auto show "path_image \ \ s - {p. fg p = 0}" using path_fg unfolding zeros_fg_def . show " finite {p. fg p = 0}" using \finite zeros_fg\ unfolding zeros_fg_def . qed moreover have "contour_integral \ (\x. deriv f x / f x) = c * (\p\zeros_f. w p * zorder f p)" unfolding c_def zeros_f_def w_def proof (rule argument_principle[OF \open s\ \connected s\ _ _ \valid_path \\ loop _ homo , of _ "{}" "\_. 1",simplified]) show "f holomorphic_on s" using f_holo g_holo holomorphic_on_add by auto show "path_image \ \ s - {p. f p = 0}" using path_f unfolding zeros_f_def . show " finite {p. f p = 0}" using \finite zeros_f\ unfolding zeros_f_def . qed ultimately have " c* (\p\zeros_fg. w p * (zorder fg p)) = c* (\p\zeros_f. w p * (zorder f p))" by auto then show ?thesis unfolding c_def using w_def by auto qed end diff --git a/src/HOL/Analysis/Further_Topology.thy b/src/HOL/Analysis/Further_Topology.thy --- a/src/HOL/Analysis/Further_Topology.thy +++ b/src/HOL/Analysis/Further_Topology.thy @@ -1,5464 +1,5714 @@ section \Extending Continous Maps, Invariance of Domain, etc\ (*FIX rename? *) text\Ported from HOL Light (moretop.ml) by L C Paulson\ theory Further_Topology imports Weierstrass_Theorems Polytope Complex_Transcendental Equivalence_Lebesgue_Henstock_Integration Retracts begin subsection\A map from a sphere to a higher dimensional sphere is nullhomotopic\ lemma spheremap_lemma1: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes "subspace S" "subspace T" and dimST: "dim S < dim T" and "S \ T" and diff_f: "f differentiable_on sphere 0 1 \ S" shows "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" proof assume fim: "f ` (sphere 0 1 \ S) = sphere 0 1 \ T" have inS: "\x. \x \ S; x \ 0\ \ (x /\<^sub>R norm x) \ S" using subspace_mul \subspace S\ by blast have subS01: "(\x. x /\<^sub>R norm x) ` (S - {0}) \ sphere 0 1 \ S" using \subspace S\ subspace_mul by fastforce then have diff_f': "f differentiable_on (\x. x /\<^sub>R norm x) ` (S - {0})" by (rule differentiable_on_subset [OF diff_f]) define g where "g \ \x. norm x *\<^sub>R f(inverse(norm x) *\<^sub>R x)" have gdiff: "g differentiable_on S - {0}" unfolding g_def by (rule diff_f' derivative_intros differentiable_on_compose [where f=f] | force)+ have geq: "g ` (S - {0}) = T - {0}" proof have "g ` (S - {0}) \ T" apply (auto simp: g_def subspace_mul [OF \subspace T\]) apply (metis (mono_tags, lifting) DiffI subS01 subspace_mul [OF \subspace T\] fim image_subset_iff inf_le2 singletonD) done moreover have "g ` (S - {0}) \ UNIV - {0}" proof (clarsimp simp: g_def) fix y assume "y \ S" and f0: "f (y /\<^sub>R norm y) = 0" then have "y \ 0 \ y /\<^sub>R norm y \ sphere 0 1 \ S" by (auto simp: subspace_mul [OF \subspace S\]) then show "y = 0" by (metis fim f0 Int_iff image_iff mem_sphere_0 norm_eq_zero zero_neq_one) qed ultimately show "g ` (S - {0}) \ T - {0}" by auto next have *: "sphere 0 1 \ T \ f ` (sphere 0 1 \ S)" using fim by (simp add: image_subset_iff) have "x \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" if "x \ T" "x \ 0" for x proof - have "x /\<^sub>R norm x \ T" using \subspace T\ subspace_mul that by blast then show ?thesis using * [THEN subsetD, of "x /\<^sub>R norm x"] that apply clarsimp apply (rule_tac x="norm x *\<^sub>R xa" in image_eqI, simp) apply (metis norm_eq_zero right_inverse scaleR_one scaleR_scaleR) using \subspace S\ subspace_mul apply force done qed then have "T - {0} \ (\x. norm x *\<^sub>R f (x /\<^sub>R norm x)) ` (S - {0})" by force then show "T - {0} \ g ` (S - {0})" by (simp add: g_def) qed define T' where "T' \ {y. \x \ T. orthogonal x y}" have "subspace T'" by (simp add: subspace_orthogonal_to_vectors T'_def) have dim_eq: "dim T' + dim T = DIM('a)" using dim_subspace_orthogonal_to_vectors [of T UNIV] \subspace T\ by (simp add: T'_def) have "\v1 v2. v1 \ span T \ (\w \ span T. orthogonal v2 w) \ x = v1 + v2" for x by (force intro: orthogonal_subspace_decomp_exists [of T x]) then obtain p1 p2 where p1span: "p1 x \ span T" and "\w. w \ span T \ orthogonal (p2 x) w" and eq: "p1 x + p2 x = x" for x by metis then have p1: "\z. p1 z \ T" and ortho: "\w. w \ T \ orthogonal (p2 x) w" for x using span_eq_iff \subspace T\ by blast+ then have p2: "\z. p2 z \ T'" by (simp add: T'_def orthogonal_commute) have p12_eq: "\x y. \x \ T; y \ T'\ \ p1(x + y) = x \ p2(x + y) = y" proof (rule orthogonal_subspace_decomp_unique [OF eq p1span, where T=T']) show "\x y. \x \ T; y \ T'\ \ p2 (x + y) \ span T'" using span_eq_iff p2 \subspace T'\ by blast show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: span_base) then have "\c x. p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" proof - fix c :: real and x :: 'a have f1: "c *\<^sub>R x = c *\<^sub>R p1 x + c *\<^sub>R p2 x" by (metis eq pth_6) have f2: "c *\<^sub>R p2 x \ T'" by (simp add: \subspace T'\ p2 subspace_scale) have "c *\<^sub>R p1 x \ T" by (metis (full_types) assms(2) p1span span_eq_iff subspace_scale) then show "p1 (c *\<^sub>R x) = c *\<^sub>R p1 x \ p2 (c *\<^sub>R x) = c *\<^sub>R p2 x" using f2 f1 p12_eq by presburger qed moreover have lin_add: "\x y. p1 (x + y) = p1 x + p1 y \ p2 (x + y) = p2 x + p2 y" proof (rule orthogonal_subspace_decomp_unique [OF _ p1span, where T=T']) show "\x y. p1 (x + y) + p2 (x + y) = p1 x + p1 y + (p2 x + p2 y)" by (simp add: add.assoc add.left_commute eq) show "\a b. \a \ T; b \ T'\ \ orthogonal a b" using T'_def by blast qed (auto simp: p1span p2 span_base span_add) ultimately have "linear p1" "linear p2" by unfold_locales auto have "(\z. g (p1 z)) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" apply (rule differentiable_on_compose [where f=g]) apply (rule linear_imp_differentiable_on [OF \linear p1\]) apply (rule differentiable_on_subset [OF gdiff]) using p12_eq \S \ T\ apply auto done then have diff: "(\x. g (p1 x) + p2 x) differentiable_on {x + y |x y. x \ S - {0} \ y \ T'}" by (intro derivative_intros linear_imp_differentiable_on [OF \linear p2\]) have "dim {x + y |x y. x \ S - {0} \ y \ T'} \ dim {x + y |x y. x \ S \ y \ T'}" by (blast intro: dim_subset) also have "... = dim S + dim T' - dim (S \ T')" using dim_sums_Int [OF \subspace S\ \subspace T'\] by (simp add: algebra_simps) also have "... < DIM('a)" using dimST dim_eq by auto finally have neg: "negligible {x + y |x y. x \ S - {0} \ y \ T'}" by (rule negligible_lowdim) have "negligible ((\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'})" by (rule negligible_differentiable_image_negligible [OF order_refl neg diff]) then have "negligible {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof (rule negligible_subset) have "\t' \ T'; s \ S; s \ 0\ \ g s + t' \ (\x. g (p1 x) + p2 x) ` {x + t' |x t'. x \ S \ x \ 0 \ t' \ T'}" for t' s apply (rule_tac x="s + t'" in image_eqI) using \S \ T\ p12_eq by auto then show "{x + y |x y. x \ g ` (S - {0}) \ y \ T'} \ (\x. g (p1 x) + p2 x) ` {x + y |x y. x \ S - {0} \ y \ T'}" by auto qed moreover have "- T' \ {x + y |x y. x \ g ` (S - {0}) \ y \ T'}" proof clarsimp fix z assume "z \ T'" show "\x y. z = x + y \ x \ g ` (S - {0}) \ y \ T'" apply (rule_tac x="p1 z" in exI) apply (rule_tac x="p2 z" in exI) apply (simp add: p1 eq p2 geq) by (metis \z \ T'\ add.left_neutral eq p2) qed ultimately have "negligible (-T')" using negligible_subset by blast moreover have "negligible T'" using negligible_lowdim by (metis add.commute assms(3) diff_add_inverse2 diff_self_eq_0 dim_eq le_add1 le_antisym linordered_semidom_class.add_diff_inverse not_less0) ultimately have "negligible (-T' \ T')" by (metis negligible_Un_eq) then show False using negligible_Un_eq non_negligible_UNIV by simp qed lemma spheremap_lemma2: fixes f :: "'a::euclidean_space \ 'a::euclidean_space" assumes ST: "subspace S" "subspace T" "dim S < dim T" and "S \ T" and contf: "continuous_on (sphere 0 1 \ S) f" and fim: "f ` (sphere 0 1 \ S) \ sphere 0 1 \ T" shows "\c. homotopic_with_canon (\x. True) (sphere 0 1 \ S) (sphere 0 1 \ T) f (\x. c)" proof - have [simp]: "\x. \norm x = 1; x \ S\ \ norm (f x) = 1" using fim by (simp add: image_subset_iff) have "compact (sphere 0 1 \ S)" by (simp add: \subspace S\ closed_subspace compact_Int_closed) then obtain g where pfg: "polynomial_function g" and gim: "g ` (sphere 0 1 \ S) \ T" and g12: "\x. x \ sphere 0 1 \ S \ norm(f x - g x) < 1/2" apply (rule Stone_Weierstrass_polynomial_function_subspace [OF _ contf _ \subspace T\, of "1/2"]) using fim apply auto done have gnz: "g x \ 0" if "x \ sphere 0 1 \ S" for x proof - have "norm (f x) = 1" using fim that by (simp add: image_subset_iff) then show ?thesis using g12 [OF that] by auto qed have diffg: "g differentiable_on sphere 0 1 \ S" by (metis pfg differentiable_on_polynomial_function) define h where "h \ \x. inverse(norm(g x)) *\<^sub>R g x" have h: "x \ sphere 0 1 \ S \ h x \ sphere 0 1 \ T" for x unfolding h_def using gnz [of x] by (auto simp: subspace_mul [OF \subspace T\] subsetD [OF gim]) have diffh: "h differentiable_on sphere 0 1 \ S" unfolding h_def apply (intro derivative_intros diffg differentiable_on_compose [OF diffg]) using gnz apply auto done have homfg: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) f g" proof (rule homotopic_with_linear [OF contf]) show "continuous_on (sphere 0 1 \ S) g" using pfg by (simp add: differentiable_imp_continuous_on diffg) next have non0fg: "0 \ closed_segment (f x) (g x)" if "norm x = 1" "x \ S" for x proof - have "f x \ sphere 0 1" using fim that by (simp add: image_subset_iff) moreover have "norm(f x - g x) < 1/2" apply (rule g12) using that by force ultimately show ?thesis by (auto simp: norm_minus_commute dest: segment_bound) qed show "\x. x \ sphere 0 1 \ S \ closed_segment (f x) (g x) \ T - {0}" apply (simp add: subset_Diff_insert non0fg) apply (simp add: segment_convex_hull) apply (rule hull_minimal) using fim image_eqI gim apply force apply (rule subspace_imp_convex [OF \subspace T\]) done qed obtain d where d: "d \ (sphere 0 1 \ T) - h ` (sphere 0 1 \ S)" using h spheremap_lemma1 [OF ST \S \ T\ diffh] by force then have non0hd: "0 \ closed_segment (h x) (- d)" if "norm x = 1" "x \ S" for x using midpoint_between [of 0 "h x" "-d"] that h [of x] by (auto simp: between_mem_segment midpoint_def) have conth: "continuous_on (sphere 0 1 \ S) h" using differentiable_imp_continuous_on diffh by blast have hom_hd: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (T - {0}) h (\x. -d)" apply (rule homotopic_with_linear [OF conth continuous_on_const]) apply (simp add: subset_Diff_insert non0hd) apply (simp add: segment_convex_hull) apply (rule hull_minimal) using h d apply (force simp: subspace_neg [OF \subspace T\]) apply (rule subspace_imp_convex [OF \subspace T\]) done have conT0: "continuous_on (T - {0}) (\y. inverse(norm y) *\<^sub>R y)" by (intro continuous_intros) auto have sub0T: "(\y. y /\<^sub>R norm y) ` (T - {0}) \ sphere 0 1 \ T" by (fastforce simp: assms(2) subspace_mul) obtain c where homhc: "homotopic_with_canon (\z. True) (sphere 0 1 \ S) (sphere 0 1 \ T) h (\x. c)" apply (rule_tac c="-d" in that) apply (rule homotopic_with_eq) apply (rule homotopic_compose_continuous_left [OF hom_hd conT0 sub0T]) using d apply (auto simp: h_def) done show ?thesis apply (rule_tac x=c in exI) apply (rule homotopic_with_trans [OF _ homhc]) apply (rule homotopic_with_eq) apply (rule homotopic_compose_continuous_left [OF homfg conT0 sub0T]) apply (auto simp: h_def) done qed lemma spheremap_lemma3: assumes "bounded S" "convex S" "subspace U" and affSU: "aff_dim S \ dim U" obtains T where "subspace T" "T \ U" "S \ {} \ aff_dim T = aff_dim S" "(rel_frontier S) homeomorphic (sphere 0 1 \ T)" proof (cases "S = {}") case True with \subspace U\ subspace_0 show ?thesis by (rule_tac T = "{0}" in that) auto next case False then obtain a where "a \ S" by auto then have affS: "aff_dim S = int (dim ((\x. -a+x) ` S))" by (metis hull_inc aff_dim_eq_dim) with affSU have "dim ((\x. -a+x) ` S) \ dim U" by linarith with choose_subspace_of_subspace obtain T where "subspace T" "T \ span U" and dimT: "dim T = dim ((\x. -a+x) ` S)" . show ?thesis proof (rule that [OF \subspace T\]) show "T \ U" using span_eq_iff \subspace U\ \T \ span U\ by blast show "aff_dim T = aff_dim S" using dimT \subspace T\ affS aff_dim_subspace by fastforce show "rel_frontier S homeomorphic sphere 0 1 \ T" proof - have "aff_dim (ball 0 1 \ T) = aff_dim (T)" by (metis IntI interior_ball \subspace T\ aff_dim_convex_Int_nonempty_interior centre_in_ball empty_iff inf_commute subspace_0 subspace_imp_convex zero_less_one) then have affS_eq: "aff_dim S = aff_dim (ball 0 1 \ T)" using \aff_dim T = aff_dim S\ by simp have "rel_frontier S homeomorphic rel_frontier(ball 0 1 \ T)" apply (rule homeomorphic_rel_frontiers_convex_bounded_sets [OF \convex S\ \bounded S\]) apply (simp add: \subspace T\ convex_Int subspace_imp_convex) apply (simp add: bounded_Int) apply (rule affS_eq) done also have "... = frontier (ball 0 1) \ T" apply (rule convex_affine_rel_frontier_Int [OF convex_ball]) apply (simp add: \subspace T\ subspace_imp_affine) using \subspace T\ subspace_0 by force also have "... = sphere 0 1 \ T" by auto finally show ?thesis . qed qed qed proposition inessential_spheremap_lowdim_gen: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "convex S" "bounded S" "convex T" "bounded T" and affST: "aff_dim S < aff_dim T" and contf: "continuous_on (rel_frontier S) f" and fim: "f ` (rel_frontier S) \ rel_frontier T" obtains c where "homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" proof (cases "S = {}") case True then show ?thesis by (simp add: that) next case False then show ?thesis proof (cases "T = {}") case True then show ?thesis using fim that by auto next case False obtain T':: "'a set" where "subspace T'" and affT': "aff_dim T' = aff_dim T" and homT: "rel_frontier T homeomorphic sphere 0 1 \ T'" apply (rule spheremap_lemma3 [OF \bounded T\ \convex T\ subspace_UNIV, where 'b='a]) apply (simp add: aff_dim_le_DIM) using \T \ {}\ by blast with homeomorphic_imp_homotopy_eqv have relT: "sphere 0 1 \ T' homotopy_eqv rel_frontier T" using homotopy_equivalent_space_sym by blast have "aff_dim S \ int (dim T')" using affT' \subspace T'\ affST aff_dim_subspace by force with spheremap_lemma3 [OF \bounded S\ \convex S\ \subspace T'\] \S \ {}\ obtain S':: "'a set" where "subspace S'" "S' \ T'" and affS': "aff_dim S' = aff_dim S" and homT: "rel_frontier S homeomorphic sphere 0 1 \ S'" by metis with homeomorphic_imp_homotopy_eqv have relS: "sphere 0 1 \ S' homotopy_eqv rel_frontier S" using homotopy_equivalent_space_sym by blast have dimST': "dim S' < dim T'" by (metis \S' \ T'\ \subspace S'\ \subspace T'\ affS' affST affT' less_irrefl not_le subspace_dim_equal) have "\c. homotopic_with_canon (\z. True) (rel_frontier S) (rel_frontier T) f (\x. c)" apply (rule homotopy_eqv_homotopic_triviality_null_imp [OF relT contf fim]) apply (rule homotopy_eqv_cohomotopic_triviality_null[OF relS, THEN iffD1, rule_format]) apply (metis dimST' \subspace S'\ \subspace T'\ \S' \ T'\ spheremap_lemma2, blast) done with that show ?thesis by blast qed qed lemma inessential_spheremap_lowdim: fixes f :: "'M::euclidean_space \ 'a::euclidean_space" assumes "DIM('M) < DIM('a)" and f: "continuous_on (sphere a r) f" "f ` (sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis by (meson nullhomotopic_into_contractible f contractible_sphere that) next case False show ?thesis proof (cases "r \ 0") case True then show ?thesis by (meson f nullhomotopic_from_contractible contractible_sphere that) next case False with \\ s \ 0\ have "r > 0" "s > 0" by auto show ?thesis apply (rule inessential_spheremap_lowdim_gen [of "cball a r" "cball b s" f]) using \0 < r\ \0 < s\ assms(1) apply (simp_all add: f aff_dim_cball) using that by blast qed qed subsection\ Some technical lemmas about extending maps from cell complexes\ lemma extending_maps_Union_aux: assumes fin: "finite \" and "\S. S \ \ \ closed S" and "\S T. \S \ \; T \ \; S \ T\ \ S \ T \ K" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" using assms proof (induction \) case empty show ?case by simp next case (insert S \) then obtain f where contf: "continuous_on (S) f" and fim: "f ` S \ T" and feq: "\x \ S \ K. f x = h x" by (meson insertI1) obtain g where contg: "continuous_on (\\) g" and gim: "g ` \\ \ T" and geq: "\x \ \\ \ K. g x = h x" using insert by auto have fg: "f x = g x" if "x \ T" "T \ \" "x \ S" for x T proof - have "T \ S \ K \ S = T" using that by (metis (no_types) insert.prems(2) insertCI) then show ?thesis using UnionI feq geq \S \ \\ subsetD that by fastforce qed show ?case apply (rule_tac x="\x. if x \ S then f x else g x" in exI, simp) apply (intro conjI continuous_on_cases) apply (simp_all add: insert closed_Union contf contg) using fim gim feq geq apply (force simp: insert closed_Union contf contg inf_commute intro: fg)+ done qed lemma extending_maps_Union: assumes fin: "finite \" and "\S. S \ \ \ \g. continuous_on S g \ g ` S \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" shows "\g. continuous_on (\\) g \ g ` (\\) \ T \ (\x \ \\ \ K. g x = h x)" apply (simp add: Union_maximal_sets [OF fin, symmetric]) apply (rule extending_maps_Union_aux) apply (simp_all add: Union_maximal_sets [OF fin] assms) by (metis K psubsetI) lemma extend_map_lemma: assumes "finite \" "\ \ \" "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ - \ \ aff_dim X < aff_dim T" and face: "\S T. \S \ \; T \ \\ \ (S \ T) face_of S \ (S \ T) face_of T" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof (cases "\ - \ = {}") case True then have "\\ \ \\" by (simp add: Union_mono) then show ?thesis apply (rule_tac g=f in that) using contf continuous_on_subset apply blast using fim apply blast by simp next case False then have "0 \ aff_dim T" by (metis aff aff_dim_empty aff_dim_geq aff_dim_negative_iff all_not_in_conv not_less) then obtain i::nat where i: "int i = aff_dim T" by (metis nonneg_eq_int) have Union_empty_eq: "\{D. D = {} \ P D} = {}" for P :: "'a set \ bool" by auto have extendf: "\g. continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) g \ g ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < i})) \ rel_frontier T \ (\x \ \\. g x = f x)" if "i \ aff_dim T" for i::nat using that proof (induction i) case 0 then show ?case apply (simp add: Union_empty_eq) apply (rule_tac x=f in exI) apply (intro conjI) using contf continuous_on_subset apply blast using fim apply blast by simp next case (Suc p) with \bounded T\ have "rel_frontier T \ {}" by (auto simp: rel_frontier_eq_empty affine_bounded_eq_lowdim [of T]) then obtain t where t: "t \ rel_frontier T" by auto have ple: "int p \ aff_dim T" using Suc.prems by force obtain h where conth: "continuous_on (\(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) h" and him: "h ` (\ (\ \ {D. \C \ \. D face_of C \ aff_dim D < p})) \ rel_frontier T" and heq: "\x. x \ \\ \ h x = f x" using Suc.IH [OF ple] by auto let ?Faces = "{D. \C \ \. D face_of C \ aff_dim D \ p}" have extendh: "\g. continuous_on D g \ g ` D \ rel_frontier T \ (\x \ D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p}). g x = h x)" if D: "D \ \ \ ?Faces" for D proof (cases "D \ \(\ \ {D. \C \ \. D face_of C \ aff_dim D < p})") case True then show ?thesis apply (rule_tac x=h in exI) apply (intro conjI) apply (blast intro: continuous_on_subset [OF conth]) using him apply blast by simp next case False note notDsub = False show ?thesis proof (cases "\a. D = {a}") case True then obtain a where "D = {a}" by auto with notDsub t show ?thesis by (rule_tac x="\x. t" in exI) simp next case False have "D \ {}" using notDsub by auto have Dnotin: "D \ \ \ {D. \C \ \. D face_of C \ aff_dim D < p}" using notDsub by auto then have "D \ \" by simp have "D \ ?Faces - {D. \C \ \. D face_of C \ aff_dim D < p}" using Dnotin that by auto then obtain C where "C \ \" "D face_of C" and affD: "aff_dim D = int p" by auto then have "bounded D" using face_of_polytope_polytope poly polytope_imp_bounded by blast then have [simp]: "\ affine D" using affine_bounded_eq_trivial False \D \ {}\ \bounded D\ by blast have "{F. F facet_of D} \ {E. E face_of C \ aff_dim E < int p}" apply clarify apply (metis \D face_of C\ affD eq_iff face_of_trans facet_of_def zle_diff1_eq) done moreover have "polyhedron D" using \C \ \\ \D face_of C\ face_of_polytope_polytope poly polytope_imp_polyhedron by auto ultimately have relf_sub: "rel_frontier D \ \ {E. E face_of C \ aff_dim E < p}" by (simp add: rel_frontier_of_polyhedron Union_mono) then have him_relf: "h ` rel_frontier D \ rel_frontier T" using \C \ \\ him by blast have "convex D" by (simp add: \polyhedron D\ polyhedron_imp_convex) have affD_lessT: "aff_dim D < aff_dim T" using Suc.prems affD by linarith have contDh: "continuous_on (rel_frontier D) h" using \C \ \\ relf_sub by (blast intro: continuous_on_subset [OF conth]) then have *: "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c)) = (\g. continuous_on UNIV g \ range g \ rel_frontier T \ (\x\rel_frontier D. g x = h x))" apply (rule nullhomotopic_into_rel_frontier_extension [OF closed_rel_frontier]) apply (simp_all add: assms rel_frontier_eq_empty him_relf) done have "(\c. homotopic_with_canon (\x. True) (rel_frontier D) (rel_frontier T) h (\x. c))" by (metis inessential_spheremap_lowdim_gen [OF \convex D\ \bounded D\ \convex T\ \bounded T\ affD_lessT contDh him_relf]) then obtain g where contg: "continuous_on UNIV g" and gim: "range g \ rel_frontier T" and gh: "\x. x \ rel_frontier D \ g x = h x" by (metis *) have "D \ E \ rel_frontier D" if "E \ \ \ {D. Bex \ ((face_of) D) \ aff_dim D < int p}" for E proof (rule face_of_subset_rel_frontier) show "D \ E face_of D" using that \C \ \\ \D face_of C\ face apply auto apply (meson face_of_Int_subface \\ \ \\ face_of_refl_eq poly polytope_imp_convex subsetD) using face_of_Int_subface apply blast done show "D \ E \ D" using that notDsub by auto qed then show ?thesis apply (rule_tac x=g in exI) apply (intro conjI ballI) using continuous_on_subset contg apply blast using gim apply blast using gh by fastforce qed qed have intle: "i < 1 + int j \ i \ int j" for i j by auto have "finite \" using \finite \\ \\ \ \\ rev_finite_subset by blast then have fin: "finite (\ \ ?Faces)" apply simp apply (rule_tac B = "\{{D. D face_of C}| C. C \ \}" in finite_subset) by (auto simp: \finite \\ finite_polytope_faces poly) have clo: "closed S" if "S \ \ \ ?Faces" for S using that \\ \ \\ face_of_polytope_polytope poly polytope_imp_closed by blast have K: "X \ Y \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < int p})" if "X \ \ \ ?Faces" "Y \ \ \ ?Faces" "\ Y \ X" for X Y proof - have ff: "X \ Y face_of X \ X \ Y face_of Y" if XY: "X face_of D" "Y face_of E" and DE: "D \ \" "E \ \" for D E apply (rule face_of_Int_subface [OF _ _ XY]) apply (auto simp: face DE) done show ?thesis using that apply auto apply (drule_tac x="X \ Y" in spec, safe) using ff face_of_imp_convex [of X] face_of_imp_convex [of Y] apply (fastforce dest: face_of_aff_dim_lt) by (meson face_of_trans ff) qed obtain g where "continuous_on (\(\ \ ?Faces)) g" "g ` \(\ \ ?Faces) \ rel_frontier T" "(\x \ \(\ \ ?Faces) \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < p}). g x = h x)" apply (rule exE [OF extending_maps_Union [OF fin extendh clo K]], blast+) done then show ?case apply (simp add: intle local.heq [symmetric], blast) done qed have eq: "\(\ \ {D. \C \ \. D face_of C \ aff_dim D < i}) = \\" proof show "\(\ \ {D. \C\\. D face_of C \ aff_dim D < int i}) \ \\" apply (rule Union_subsetI) using \\ \ \\ face_of_imp_subset apply force done show "\\ \ \(\ \ {D. \C\\. D face_of C \ aff_dim D < i})" apply (rule Union_mono) using face apply (fastforce simp: aff i) done qed have "int i \ aff_dim T" by (simp add: i) then show ?thesis using extendf [of i] unfolding eq by (metis that) qed lemma extend_map_lemma_cofinite0: assumes "finite \" and "pairwise (\S T. S \ T \ K) \" and "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and "\S. S \ \ \ closed S" shows "\C g. finite C \ disjnt C U \ card C \ card \ \ continuous_on (\\ - C) g \ g ` (\\ - C) \ T \ (\x \ (\\ - C) \ K. g x = h x)" using assms proof induction case empty then show ?case by force next case (insert X \) then have "closed X" and clo: "\X. X \ \ \ closed X" and \: "\S. S \ \ \ \a g. a \ U \ continuous_on (S - {a}) g \ g ` (S - {a}) \ T \ (\x \ S \ K. g x = h x)" and pwX: "\Y. Y \ \ \ Y \ X \ X \ Y \ K \ Y \ X \ K" and pwF: "pairwise (\ S T. S \ T \ K) \" by (simp_all add: pairwise_insert) obtain C g where C: "finite C" "disjnt C U" "card C \ card \" and contg: "continuous_on (\\ - C) g" and gim: "g ` (\\ - C) \ T" and gh: "\x. x \ (\\ - C) \ K \ g x = h x" using insert.IH [OF pwF \ clo] by auto obtain a f where "a \ U" and contf: "continuous_on (X - {a}) f" and fim: "f ` (X - {a}) \ T" and fh: "(\x \ X \ K. f x = h x)" using insert.prems by (meson insertI1) show ?case proof (intro exI conjI) show "finite (insert a C)" by (simp add: C) show "disjnt (insert a C) U" using C \a \ U\ by simp show "card (insert a C) \ card (insert X \)" by (simp add: C card_insert_if insert.hyps le_SucI) have "closed (\\)" using clo insert.hyps by blast have "continuous_on (X - insert a C \ (\\ - insert a C)) (\x. if x \ X then f x else g x)" apply (rule continuous_on_cases_local) apply (simp_all add: closedin_closed) using \closed X\ apply blast using \closed (\\)\ apply blast using contf apply (force simp: elim: continuous_on_subset) using contg apply (force simp: elim: continuous_on_subset) using fh gh insert.hyps pwX by fastforce then show "continuous_on (\(insert X \) - insert a C) (\a. if a \ X then f a else g a)" by (blast intro: continuous_on_subset) show "\x\(\(insert X \) - insert a C) \ K. (if x \ X then f x else g x) = h x" using gh by (auto simp: fh) show "(\a. if a \ X then f a else g a) ` (\(insert X \) - insert a C) \ T" using fim gim by auto force qed qed lemma extend_map_lemma_cofinite1: assumes "finite \" and \: "\X. X \ \ \ \a g. a \ U \ continuous_on (X - {a}) g \ g ` (X - {a}) \ T \ (\x \ X \ K. g x = h x)" and clo: "\X. X \ \ \ closed X" and K: "\X Y. \X \ \; Y \ \; \ X \ Y; \ Y \ X\ \ X \ Y \ K" obtains C g where "finite C" "disjnt C U" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ T" "\x. x \ (\\ - C) \ K \ g x = h x" proof - let ?\ = "{X \ \. \Y\\. \ X \ Y}" have [simp]: "\?\ = \\" by (simp add: Union_maximal_sets assms) have fin: "finite ?\" by (force intro: finite_subset [OF _ \finite \\]) have pw: "pairwise (\ S T. S \ T \ K) ?\" by (simp add: pairwise_def) (metis K psubsetI) have "card {X \ \. \Y\\. \ X \ Y} \ card \" by (simp add: \finite \\ card_mono) moreover obtain C g where "finite C \ disjnt C U \ card C \ card ?\ \ continuous_on (\?\ - C) g \ g ` (\?\ - C) \ T \ (\x \ (\?\ - C) \ K. g x = h x)" apply (rule exE [OF extend_map_lemma_cofinite0 [OF fin pw, of U T h]]) apply (fastforce intro!: clo \)+ done ultimately show ?thesis by (rule_tac C=C and g=g in that) auto qed lemma extend_map_lemma_cofinite: assumes "finite \" "\ \ \" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and contf: "continuous_on (\\) f" and fim: "f ` (\\) \ rel_frontier T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" and aff: "\X. X \ \ - \ \ aff_dim X \ aff_dim T" obtains C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\ \ - C) \ rel_frontier T" "\x. x \ \\ \ g x = f x" proof - define \ where "\ \ \ \ {D. \C \ \ - \. D face_of C \ aff_dim D < aff_dim T}" have "finite \" using assms finite_subset by blast moreover have "finite (\{{D. D face_of C} |C. C \ \})" apply (rule finite_Union) apply (simp add: \finite \\) using finite_polytope_faces poly by auto ultimately have "finite \" apply (simp add: \_def) apply (rule finite_subset [of _ "\ {{D. D face_of C} | C. C \ \}"], auto) done have *: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" unfolding \_def apply (elim UnE bexE CollectE DiffE) using subsetD [OF \\ \ \\] apply (simp_all add: face) apply (meson subsetD [OF \\ \ \\] face face_of_Int_subface face_of_imp_subset face_of_refl poly polytope_imp_convex)+ done obtain h where conth: "continuous_on (\\) h" and him: "h ` (\\) \ rel_frontier T" and hf: "\x. x \ \\ \ h x = f x" using \finite \\ unfolding \_def apply (rule extend_map_lemma [OF _ Un_upper1 T _ _ _ contf fim]) using \\ \ \\ face_of_polytope_polytope poly apply fastforce using * apply (auto simp: \_def) done have "bounded (\\)" using \finite \\ \\ \ \\ poly polytope_imp_bounded by blast then have "\\ \ UNIV" by auto then obtain a where a: "a \ \\" by blast have \: "\a g. a \ \\ \ continuous_on (D - {a}) g \ g ` (D - {a}) \ rel_frontier T \ (\x \ D \ \\. g x = h x)" if "D \ \" for D proof (cases "D \ \\") case True then show ?thesis apply (rule_tac x=a in exI) apply (rule_tac x=h in exI) using him apply (blast intro!: \a \ \\\ continuous_on_subset [OF conth]) + done next case False note D_not_subset = False show ?thesis proof (cases "D \ \") case True with D_not_subset show ?thesis by (auto simp: \_def) next case False then have affD: "aff_dim D \ aff_dim T" by (simp add: \D \ \\ aff) show ?thesis proof (cases "rel_interior D = {}") case True with \D \ \\ poly a show ?thesis by (force simp: rel_interior_eq_empty polytope_imp_convex) next case False then obtain b where brelD: "b \ rel_interior D" by blast have "polyhedron D" by (simp add: poly polytope_imp_polyhedron that) have "rel_frontier D retract_of affine hull D - {b}" by (simp add: rel_frontier_retract_of_punctured_affine_hull poly polytope_imp_bounded polytope_imp_convex that brelD) then obtain r where relfD: "rel_frontier D \ affine hull D - {b}" and contr: "continuous_on (affine hull D - {b}) r" and rim: "r ` (affine hull D - {b}) \ rel_frontier D" and rid: "\x. x \ rel_frontier D \ r x = x" by (auto simp: retract_of_def retraction_def) show ?thesis proof (intro exI conjI ballI) show "b \ \\" proof clarify fix E assume "b \ E" "E \ \" then have "E \ D face_of E \ E \ D face_of D" using \\ \ \\ face that by auto with face_of_subset_rel_frontier \E \ \\ \b \ E\ brelD rel_interior_subset [of D] D_not_subset rel_frontier_def \_def show False by blast qed have "r ` (D - {b}) \ r ` (affine hull D - {b})" by (simp add: Diff_mono hull_subset image_mono) also have "... \ rel_frontier D" by (rule rim) also have "... \ \{E. E face_of D \ aff_dim E < aff_dim T}" using affD by (force simp: rel_frontier_of_polyhedron [OF \polyhedron D\] facet_of_def) also have "... \ \(\)" using D_not_subset \_def that by fastforce finally have rsub: "r ` (D - {b}) \ \(\)" . show "continuous_on (D - {b}) (h \ r)" apply (intro conjI \b \ \\\ continuous_on_compose) apply (rule continuous_on_subset [OF contr]) apply (simp add: Diff_mono hull_subset) apply (rule continuous_on_subset [OF conth rsub]) done show "(h \ r) ` (D - {b}) \ rel_frontier T" using brelD him rsub by fastforce show "(h \ r) x = h x" if x: "x \ D \ \\" for x proof - consider A where "x \ D" "A \ \" "x \ A" | A B where "x \ D" "A face_of B" "B \ \" "B \ \" "aff_dim A < aff_dim T" "x \ A" using x by (auto simp: \_def) then have xrel: "x \ rel_frontier D" proof cases case 1 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) show "D \ A face_of D" using \A \ \\ \\ \ \\ face \D \ \\ by blast show "D \ A \ D" using \A \ \\ D_not_subset \_def by blast qed (auto simp: 1) next case 2 show ?thesis proof (rule face_of_subset_rel_frontier [THEN subsetD]) show "D \ A face_of D" apply (rule face_of_Int_subface [of D B _ A, THEN conjunct1]) apply (simp_all add: 2 \D \ \\ face) apply (simp add: \polyhedron D\ polyhedron_imp_convex face_of_refl) done show "D \ A \ D" using "2" D_not_subset \_def by blast qed (auto simp: 2) qed show ?thesis by (simp add: rid xrel) qed qed qed qed qed have clo: "\S. S \ \ \ closed S" by (simp add: poly polytope_imp_closed) obtain C g where "finite C" "disjnt C (\\)" "card C \ card \" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" and gh: "\x. x \ (\\ - C) \ \\ \ g x = h x" proof (rule extend_map_lemma_cofinite1 [OF \finite \\ \ clo]) show "X \ Y \ \\" if XY: "X \ \" "Y \ \" and "\ X \ Y" "\ Y \ X" for X Y proof (cases "X \ \") case True then show ?thesis by (auto simp: \_def) next case False have "X \ Y \ X" using \\ X \ Y\ by blast with XY show ?thesis by (clarsimp simp: \_def) (metis Diff_iff Int_iff aff antisym_conv face face_of_aff_dim_lt face_of_refl not_le poly polytope_imp_convex) qed qed (blast)+ with \\ \ \\ show ?thesis apply (rule_tac C=C and g=g in that) apply (auto simp: disjnt_def hf [symmetric] \_def intro!: gh) done qed text\The next two proofs are similar\ theorem extend_map_cell_complex_to_sphere: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X < aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains g where "continuous_on (\\) g" "g ` (\\) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T - 1" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" proof (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly _ face]) show "\X. X \ \ \ aff_dim X \ aff_dim T - 1" by (simp add: aff) qed auto obtain h where conth: "continuous_on (\\) h" and him: "h ` \\ \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) qed (use \finite \\ T polyG affG faceG gim in fastforce)+ show ?thesis proof show "continuous_on (\\) h" using \\\ = \\\ conth by auto show "h ` \\ \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "x \ \\" using \\\ = \\\ \S \ \\\ that by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce have "h x = g x" apply (rule hg) using \X \ \\ \X \ V\ \x \ X\ by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed qed qed theorem extend_map_cell_complex_to_sphere_cofinite: assumes "finite \" and S: "S \ \\" "closed S" and T: "convex T" "bounded T" and poly: "\X. X \ \ \ polytope X" and aff: "\X. X \ \ \ aff_dim X \ aff_dim T" and face: "\X Y. \X \ \; Y \ \\ \ (X \ Y) face_of X \ (X \ Y) face_of Y" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier T" obtains C g where "finite C" "disjnt C S" "continuous_on (\\ - C) g" "g ` (\\ - C) \ rel_frontier T" "\x. x \ S \ g x = f x" proof - obtain V g where "S \ V" "open V" "continuous_on V g" and gim: "g ` V \ rel_frontier T" and gf: "\x. x \ S \ g x = f x" using neighbourhood_extension_into_ANR [OF contf fim _ \closed S\] ANR_rel_frontier_convex T by blast have "compact S" by (meson assms compact_Union poly polytope_imp_compact seq_compact_closed_subset seq_compact_eq_compact) then obtain d where "d > 0" and d: "\x y. \x \ S; y \ - V\ \ d \ dist x y" using separate_compact_closed [of S "-V"] \open V\ \S \ V\ by force obtain \ where "finite \" "\\ = \\" and diaG: "\X. X \ \ \ diameter X < d" and polyG: "\X. X \ \ \ polytope X" and affG: "\X. X \ \ \ aff_dim X \ aff_dim T" and faceG: "\X Y. \X \ \; Y \ \\ \ X \ Y face_of X \ X \ Y face_of Y" by (rule cell_complex_subdivision_exists [OF \d>0\ \finite \\ poly aff face]) auto obtain C h where "finite C" and dis: "disjnt C (\(\ \ Pow V))" and card: "card C \ card \" and conth: "continuous_on (\\ - C) h" and him: "h ` (\\ - C) \ rel_frontier T" and hg: "\x. x \ \(\ \ Pow V) \ h x = g x" proof (rule extend_map_lemma_cofinite [of \ "\ \ Pow V" T g]) show "continuous_on (\(\ \ Pow V)) g" by (metis Union_Int_subset Union_Pow_eq \continuous_on V g\ continuous_on_subset le_inf_iff) show "g ` \(\ \ Pow V) \ rel_frontier T" using gim by force qed (auto intro: \finite \\ T polyG affG dest: faceG) have Ssub: "S \ \(\ \ Pow V)" proof fix x assume "x \ S" then have "x \ \\" using \\\ = \\\ \S \ \\\ by auto then obtain X where "x \ X" "X \ \" by blast then have "diameter X < d" "bounded X" by (auto simp: diaG \X \ \\ polyG polytope_imp_bounded) then have "X \ V" using d [OF \x \ S\] diameter_bounded_bound [OF \bounded X\ \x \ X\] by fastforce then show "x \ \(\ \ Pow V)" using \X \ \\ \x \ X\ by blast qed show ?thesis proof show "continuous_on (\\-C) h" using \\\ = \\\ conth by auto show "h ` (\\ - C) \ rel_frontier T" using \\\ = \\\ him by auto show "h x = f x" if "x \ S" for x proof - have "h x = g x" apply (rule hg) using Ssub that by blast also have "... = f x" by (simp add: gf that) finally show "h x = f x" . qed show "disjnt C S" using dis Ssub by (meson disjnt_iff subset_eq) qed (intro \finite C\) qed subsection\ Special cases and corollaries involving spheres\ lemma disjnt_Diff1: "X \ Y' \ disjnt (X - Y) (X' - Y')" by (auto simp: disjnt_def) proposition extend_map_affine_to_sphere_cofinite_simple: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" and aff: "aff_dim T \ aff_dim U" and "S \ T" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" obtains K g where "finite K" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - have "\K g. finite K \ disjnt K S \ continuous_on (T - K) g \ g ` (T - K) \ rel_frontier U \ (\x \ S. g x = f x)" if "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" for T proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with \bounded U\ have "aff_dim U \ 0" using affine_bounded_eq_lowdim rel_frontier_eq_empty by auto with aff have "aff_dim T \ 0" by auto then obtain a where "T \ {a}" using \affine T\ affine_bounded_eq_lowdim affine_bounded_eq_trivial by auto then show ?thesis using \S = {}\ fim by (metis Diff_cancel contf disjnt_empty2 finite.emptyI finite_insert finite_subset) next case False then obtain a where "a \ rel_frontier U" by auto then show ?thesis using continuous_on_const [of _ a] \S = {}\ by force qed next case False have "bounded S" by (simp add: \compact S\ compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define bbox where "bbox \ cbox (-(b+One)) (b+One)" have "cbox (-b) b \ bbox" by (auto simp: bbox_def algebra_simps intro!: subset_box_imp) with b \S \ T\ have "S \ bbox \ T" by auto then have Ssub: "S \ \{bbox \ T}" by auto then have "aff_dim (bbox \ T) \ aff_dim U" by (metis aff aff_dim_subset inf_commute inf_le1 order_trans) obtain K g where K: "finite K" "disjnt K S" and contg: "continuous_on (\{bbox \ T} - K) g" and gim: "g ` (\{bbox \ T} - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_cell_complex_to_sphere_cofinite [OF _ Ssub _ \convex U\ \bounded U\ _ _ _ contf fim]) show "closed S" using \compact S\ compact_eq_bounded_closed by auto show poly: "\X. X \ {bbox \ T} \ polytope X" by (simp add: polytope_Int_polyhedron bbox_def polytope_interval affine_imp_polyhedron \affine T\) show "\X Y. \X \ {bbox \ T}; Y \ {bbox \ T}\ \ X \ Y face_of X \ X \ Y face_of Y" by (simp add:poly face_of_refl polytope_imp_convex) show "\X. X \ {bbox \ T} \ aff_dim X \ aff_dim U" by (simp add: \aff_dim (bbox \ T) \ aff_dim U\) qed auto define fro where "fro \ \d. frontier(cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))" obtain d where d12: "1/2 \ d" "d \ 1" and dd: "disjnt K (fro d)" proof (rule disjoint_family_elem_disjnt [OF _ \finite K\]) show "infinite {1/2..1::real}" by (simp add: infinite_Icc) have dis1: "disjnt (fro x) (fro y)" if "x b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ bbox" using d12 by (auto simp: algebra_simps subset_box_imp c_def bbox_def) have clo_cbT: "closed (cbox (- c) c \ T)" by (simp add: affine_closed closed_Int closed_cbox \affine T\) have cpT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ b cbsub(2) \S \ T\ by fastforce have "closest_point (cbox (- c) c \ T) x \ K" if "x \ T" "x \ K" for x proof (cases "x \ cbox (-c) c") case True with that show ?thesis by (simp add: closest_point_self) next case False have int_ne: "interior (cbox (-c) c) \ T \ {}" using \S \ {}\ \S \ T\ b \cbox (- b) b \ box (- c) c\ by force have "convex T" by (meson \affine T\ affine_imp_convex) then have "x \ affine hull (cbox (- c) c \ T)" by (metis Int_commute Int_iff \S \ {}\ \S \ T\ cbsub(1) \x \ T\ affine_hull_convex_Int_nonempty_interior all_not_in_conv b hull_inc inf.orderE interior_cbox) then have "x \ affine hull (cbox (- c) c \ T) - rel_interior (cbox (- c) c \ T)" by (meson DiffI False Int_iff rel_interior_subset subsetCE) then have "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" by (rule closest_point_in_rel_frontier [OF clo_cbT cpT_ne]) moreover have "(rel_frontier (cbox (- c) c \ T)) \ fro d" apply (subst convex_affine_rel_frontier_Int [OF _ \affine T\ int_ne]) apply (auto simp: fro_def c_def) done ultimately show ?thesis using dd by (force simp: disjnt_def) qed then have cpt_subset: "closest_point (cbox (- c) c \ T) ` (T - K) \ \{bbox \ T} - K" using closest_point_in_set [OF clo_cbT cpT_ne] cbsub(3) by force show ?thesis proof (intro conjI ballI exI) have "continuous_on (T - K) (closest_point (cbox (- c) c \ T))" apply (rule continuous_on_closest_point) using \S \ {}\ cbsub(2) b that by (auto simp: affine_imp_convex convex_Int affine_closed closed_Int closed_cbox \affine T\) then show "continuous_on (T - K) (g \ closest_point (cbox (- c) c \ T))" by (metis continuous_on_compose continuous_on_subset [OF contg cpt_subset]) have "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ g ` (\{bbox \ T} - K)" by (metis image_comp image_mono cpt_subset) also have "... \ rel_frontier U" by (rule gim) finally show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K) \ rel_frontier U" . show "(g \ closest_point (cbox (- c) c \ T)) x = f x" if "x \ S" for x proof - have "(g \ closest_point (cbox (- c) c \ T)) x = g x" unfolding o_def by (metis IntI \S \ T\ b cbsub(2) closest_point_self subset_eq that) also have "... = f x" by (simp add: that gf) finally show ?thesis . qed qed (auto simp: K) qed then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (affine hull T - K) g" and gim: "g ` (affine hull T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (metis aff affine_affine_hull aff_dim_affine_hull order_trans [OF \S \ T\ hull_subset [of T affine]]) then obtain K g where "finite K" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" by (rule_tac K=K and g=g in that) (auto simp: hull_inc elim: continuous_on_subset) then show ?thesis by (rule_tac K="K \ T" and g=g in that) (auto simp: disjnt_iff Diff_Int contg) qed subsection\Extending maps to spheres\ (*Up to extend_map_affine_to_sphere_cofinite_gen*) lemma extend_map_affine_to_sphere1: fixes f :: "'a::euclidean_space \ 'b::topological_space" assumes "finite K" "affine U" and contf: "continuous_on (U - K) f" and fim: "f ` (U - K) \ T" and comps: "\C. \C \ components(U - S); C \ K \ {}\ \ C \ L \ {}" and clo: "closedin (top_of_set U) S" and K: "disjnt K S" "K \ U" obtains g where "continuous_on (U - L) g" "g ` (U - L) \ T" "\x. x \ S \ g x = f x" proof (cases "K = {}") case True then show ?thesis by (metis Diff_empty Diff_subset contf fim continuous_on_subset image_subsetI rev_image_eqI subset_iff that) next case False have "S \ U" using clo closedin_limpt by blast then have "(U - S) \ K \ {}" by (metis Diff_triv False Int_Diff K disjnt_def inf.absorb_iff2 inf_commute) then have "\(components (U - S)) \ K \ {}" using Union_components by simp then obtain C0 where C0: "C0 \ components (U - S)" "C0 \ K \ {}" by blast have "convex U" by (simp add: affine_imp_convex \affine U\) then have "locally connected U" by (rule convex_imp_locally_connected) have "\a g. a \ C \ a \ L \ continuous_on (S \ (C - {a})) g \ g ` (S \ (C - {a})) \ T \ (\x \ S. g x = f x)" if C: "C \ components (U - S)" and CK: "C \ K \ {}" for C proof - have "C \ U-S" "C \ L \ {}" by (simp_all add: in_components_subset comps that) then obtain a where a: "a \ C" "a \ L" by auto have opeUC: "openin (top_of_set U) C" proof (rule openin_trans) show "openin (top_of_set (U-S)) C" by (simp add: \locally connected U\ clo locally_diff_closed openin_components_locally_connected [OF _ C]) show "openin (top_of_set U) (U - S)" by (simp add: clo openin_diff) qed then obtain d where "C \ U" "0 < d" and d: "cball a d \ U \ C" using openin_contains_cball by (metis \a \ C\) then have "ball a d \ U \ C" by auto obtain h k where homhk: "homeomorphism (S \ C) (S \ C) h k" and subC: "{x. (\ (h x = x \ k x = x))} \ C" and bou: "bounded {x. (\ (h x = x \ k x = x))}" and hin: "\x. x \ C \ K \ h x \ ball a d \ U" proof (rule homeomorphism_grouping_points_exists_gen [of C "ball a d \ U" "C \ K" "S \ C"]) show "openin (top_of_set C) (ball a d \ U)" by (metis open_ball \C \ U\ \ball a d \ U \ C\ inf.absorb_iff2 inf.orderE inf_assoc open_openin openin_subtopology) show "openin (top_of_set (affine hull C)) C" by (metis \a \ C\ \openin (top_of_set U) C\ affine_hull_eq affine_hull_openin all_not_in_conv \affine U\) show "ball a d \ U \ {}" using \0 < d\ \C \ U\ \a \ C\ by force show "finite (C \ K)" by (simp add: \finite K\) show "S \ C \ affine hull C" by (metis \C \ U\ \S \ U\ \a \ C\ opeUC affine_hull_eq affine_hull_openin all_not_in_conv assms(2) sup.bounded_iff) show "connected C" by (metis C in_components_connected) qed auto have a_BU: "a \ ball a d \ U" using \0 < d\ \C \ U\ \a \ C\ by auto have "rel_frontier (cball a d \ U) retract_of (affine hull (cball a d \ U) - {a})" apply (rule rel_frontier_retract_of_punctured_affine_hull) apply (auto simp: \convex U\ convex_Int) by (metis \affine U\ convex_cball empty_iff interior_cball a_BU rel_interior_convex_Int_affine) moreover have "rel_frontier (cball a d \ U) = frontier (cball a d) \ U" apply (rule convex_affine_rel_frontier_Int) using a_BU by (force simp: \affine U\)+ moreover have "affine hull (cball a d \ U) = U" by (metis \convex U\ a_BU affine_hull_convex_Int_nonempty_interior affine_hull_eq \affine U\ equals0D inf.commute interior_cball) ultimately have "frontier (cball a d) \ U retract_of (U - {a})" by metis then obtain r where contr: "continuous_on (U - {a}) r" and rim: "r ` (U - {a}) \ sphere a d" "r ` (U - {a}) \ U" and req: "\x. x \ sphere a d \ U \ r x = x" using \affine U\ by (auto simp: retract_of_def retraction_def hull_same) define j where "j \ \x. if x \ ball a d then r x else x" have kj: "\x. x \ S \ k (j x) = x" using \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def subC by auto have Uaeq: "U - {a} = (cball a d - {a}) \ U \ (U - ball a d)" using \0 < d\ by auto have jim: "j ` (S \ (C - {a})) \ (S \ C) - ball a d" proof clarify fix y assume "y \ S \ (C - {a})" then have "y \ U - {a}" using \C \ U - S\ \S \ U\ \a \ C\ by auto then have "r y \ sphere a d" using rim by auto then show "j y \ S \ C - ball a d" apply (simp add: j_def) using \r y \ sphere a d\ \y \ U - {a}\ \y \ S \ (C - {a})\ d rim by fastforce qed have contj: "continuous_on (U - {a}) j" unfolding j_def Uaeq proof (intro continuous_on_cases_local continuous_on_id, simp_all add: req closedin_closed Uaeq [symmetric]) show "\T. closed T \ (cball a d - {a}) \ U = (U - {a}) \ T" apply (rule_tac x="(cball a d) \ U" in exI) using affine_closed \affine U\ by blast show "\T. closed T \ U - ball a d = (U - {a}) \ T" apply (rule_tac x="U - ball a d" in exI) using \0 < d\ by (force simp: affine_closed \affine U\ closed_Diff) show "continuous_on ((cball a d - {a}) \ U) r" by (force intro: continuous_on_subset [OF contr]) qed have fT: "x \ U - K \ f x \ T" for x using fim by blast show ?thesis proof (intro conjI exI) show "continuous_on (S \ (C - {a})) (f \ k \ j)" proof (intro continuous_on_compose) show "continuous_on (S \ (C - {a})) j" apply (rule continuous_on_subset [OF contj]) using \C \ U - S\ \S \ U\ \a \ C\ by force show "continuous_on (j ` (S \ (C - {a}))) k" apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) using jim \C \ U - S\ \S \ U\ \ball a d \ U \ C\ j_def by fastforce show "continuous_on (k ` j ` (S \ (C - {a}))) f" proof (clarify intro!: continuous_on_subset [OF contf]) fix y assume "y \ S \ (C - {a})" have ky: "k y \ S \ C" using homeomorphism_image2 [OF homhk] \y \ S \ (C - {a})\ by blast have jy: "j y \ S \ C - ball a d" using Un_iff \y \ S \ (C - {a})\ jim by auto show "k (j y) \ U - K" apply safe using \C \ U\ \S \ U\ homeomorphism_image2 [OF homhk] jy apply blast by (metis DiffD1 DiffD2 Int_iff Un_iff \disjnt K S\ disjnt_def empty_iff hin homeomorphism_apply2 homeomorphism_image2 homhk imageI jy) qed qed have ST: "\x. x \ S \ (f \ k \ j) x \ T" apply (simp add: kj) apply (metis DiffI \S \ U\ \disjnt K S\ subsetD disjnt_iff fim image_subset_iff) done moreover have "(f \ k \ j) x \ T" if "x \ C" "x \ a" "x \ S" for x proof - have rx: "r x \ sphere a d" using \C \ U\ rim that by fastforce have jj: "j x \ S \ C - ball a d" using jim that by blast have "k (j x) = j x \ k (j x) \ C \ j x \ C" by (metis Diff_iff Int_iff Un_iff \S \ U\ subsetD d j_def jj rx sphere_cball that(1)) then have "k (j x) \ C" using homeomorphism_apply2 [OF homhk, of "j x"] \C \ U\ \S \ U\ a rx by (metis (mono_tags, lifting) Diff_iff subsetD jj mem_Collect_eq subC) with jj \C \ U\ show ?thesis apply safe using ST j_def apply fastforce apply (auto simp: not_less intro!: fT) by (metis DiffD1 DiffD2 Int_iff hin homeomorphism_apply2 [OF homhk] jj) qed ultimately show "(f \ k \ j) ` (S \ (C - {a})) \ T" by force show "\x\S. (f \ k \ j) x = f x" using kj by simp qed (auto simp: a) qed then obtain a h where ah: "\C. \C \ components (U - S); C \ K \ {}\ \ a C \ C \ a C \ L \ continuous_on (S \ (C - {a C})) (h C) \ h C ` (S \ (C - {a C})) \ T \ (\x \ S. h C x = f x)" using that by metis define F where "F \ {C \ components (U - S). C \ K \ {}}" define G where "G \ {C \ components (U - S). C \ K = {}}" define UF where "UF \ (\C\F. C - {a C})" have "C0 \ F" by (auto simp: F_def C0) have "finite F" proof (subst finite_image_iff [of "\C. C \ K" F, symmetric]) show "inj_on (\C. C \ K) F" unfolding F_def inj_on_def using components_nonoverlap by blast show "finite ((\C. C \ K) ` F)" unfolding F_def by (rule finite_subset [of _ "Pow K"]) (auto simp: \finite K\) qed obtain g where contg: "continuous_on (S \ UF) g" and gh: "\x i. \i \ F; x \ (S \ UF) \ (S \ (i - {a i}))\ \ g x = h i x" proof (rule pasting_lemma_exists_closed [OF \finite F\]) let ?X = "top_of_set (S \ UF)" show "topspace ?X \ (\C\F. S \ (C - {a C}))" using \C0 \ F\ by (force simp: UF_def) show "closedin (top_of_set (S \ UF)) (S \ (C - {a C}))" if "C \ F" for C proof (rule closedin_closed_subset [of U "S \ C"]) show "closedin (top_of_set U) (S \ C)" apply (rule closedin_Un_complement_component [OF \locally connected U\ clo]) using F_def that by blast next have "x = a C'" if "C' \ F" "x \ C'" "x \ U" for x C' proof - have "\A. x \ \A \ C' \ A" using \x \ C'\ by blast with that show "x = a C'" by (metis (lifting) DiffD1 F_def Union_components mem_Collect_eq) qed then show "S \ UF \ U" using \S \ U\ by (force simp: UF_def) next show "S \ (C - {a C}) = (S \ C) \ (S \ UF)" using F_def UF_def components_nonoverlap that by auto qed show "continuous_map (subtopology ?X (S \ (C' - {a C'}))) euclidean (h C')" if "C' \ F" for C' proof - have C': "C' \ components (U - S)" "C' \ K \ {}" using F_def that by blast+ show ?thesis using ah [OF C'] by (auto simp: F_def subtopology_subtopology intro: continuous_on_subset) qed show "\i j x. \i \ F; j \ F; x \ topspace ?X \ (S \ (i - {a i})) \ (S \ (j - {a j}))\ \ h i x = h j x" using components_eq by (fastforce simp: components_eq F_def ah) qed auto have SU': "S \ \G \ (S \ UF) \ U" using \S \ U\ in_components_subset by (auto simp: F_def G_def UF_def) have clo1: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ \G)" proof (rule closedin_closed_subset [OF _ SU']) have *: "\C. C \ F \ openin (top_of_set U) C" unfolding F_def by clarify (metis (no_types, lifting) \locally connected U\ clo closedin_def locally_diff_closed openin_components_locally_connected openin_trans topspace_euclidean_subtopology) show "closedin (top_of_set U) (U - UF)" unfolding UF_def by (force intro: openin_delete *) show "S \ \G = (U - UF) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) apply (metis Diff_iff UnionI Union_components) apply (metis DiffD1 UnionI Union_components) by (metis (no_types, lifting) IntI components_nonoverlap empty_iff) qed have clo2: "closedin (top_of_set (S \ \G \ (S \ UF))) (S \ UF)" proof (rule closedin_closed_subset [OF _ SU']) show "closedin (top_of_set U) (\C\F. S \ C)" apply (rule closedin_Union) apply (simp add: \finite F\) using F_def \locally connected U\ clo closedin_Un_complement_component by blast show "S \ UF = (\C\F. S \ C) \ (S \ \G \ (S \ UF))" using \S \ U\ apply (auto simp: F_def G_def UF_def) using C0 apply blast by (metis components_nonoverlap disjnt_def disjnt_iff) qed have SUG: "S \ \G \ U - K" using \S \ U\ K apply (auto simp: G_def disjnt_iff) by (meson Diff_iff subsetD in_components_subset) then have contf': "continuous_on (S \ \G) f" by (rule continuous_on_subset [OF contf]) have contg': "continuous_on (S \ UF) g" apply (rule continuous_on_subset [OF contg]) using \S \ U\ by (auto simp: F_def G_def) have "\x. \S \ U; x \ S\ \ f x = g x" by (subst gh) (auto simp: ah C0 intro: \C0 \ F\) then have f_eq_g: "\x. x \ S \ UF \ x \ S \ \G \ f x = g x" using \S \ U\ apply (auto simp: F_def G_def UF_def dest: in_components_subset) using components_eq by blast have cont: "continuous_on (S \ \G \ (S \ UF)) (\x. if x \ S \ \G then f x else g x)" by (blast intro: continuous_on_cases_local [OF clo1 clo2 contf' contg' f_eq_g, of "\x. x \ S \ \G"]) show ?thesis proof have UF: "\F - L \ UF" unfolding F_def UF_def using ah by blast have "U - S - L = \(components (U - S)) - L" by simp also have "... = \F \ \G - L" unfolding F_def G_def by blast also have "... \ UF \ \G" using UF by blast finally have "U - L \ S \ \G \ (S \ UF)" by blast then show "continuous_on (U - L) (\x. if x \ S \ \G then f x else g x)" by (rule continuous_on_subset [OF cont]) have "((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ ((U - L) \ (-S \ UF))" using \U - L \ S \ \G \ (S \ UF)\ by auto moreover have "g ` ((U - L) \ (-S \ UF)) \ T" proof - have "g x \ T" if "x \ U" "x \ L" "x \ S" "C \ F" "x \ C" "x \ a C" for x C proof (subst gh) show "x \ (S \ UF) \ (S \ (C - {a C}))" using that by (auto simp: UF_def) show "h C x \ T" using ah that by (fastforce simp add: F_def) qed (rule that) then show ?thesis by (force simp: UF_def) qed ultimately have "g ` ((U - L) \ {x. x \ S \ (\xa\G. x \ xa)}) \ T" using image_mono order_trans by blast moreover have "f ` ((U - L) \ (S \ \G)) \ T" using fim SUG by blast ultimately show "(\x. if x \ S \ \G then f x else g x) ` (U - L) \ T" by force show "\x. x \ S \ (if x \ S \ \G then f x else g x) = f x" by (simp add: F_def G_def) qed qed lemma extend_map_affine_to_sphere2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "compact S" "convex U" "bounded U" "affine T" "S \ T" and affTU: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and ovlap: "\C. C \ components(T - S) \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof - obtain K g where K: "finite K" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" using assms extend_map_affine_to_sphere_cofinite_simple by metis have "(\y C. C \ components (T - S) \ x \ C \ y \ C \ y \ L)" if "x \ K" for x proof - have "x \ T-S" using \K \ T\ \disjnt K S\ disjnt_def that by fastforce then obtain C where "C \ components(T - S)" "x \ C" by (metis UnionE Union_components) with ovlap [of C] show ?thesis by blast qed then obtain \ where \: "\x. x \ K \ \C. C \ components (T - S) \ x \ C \ \ x \ C \ \ x \ L" by metis obtain h where conth: "continuous_on (T - \ ` K) h" and him: "h ` (T - \ ` K) \ rel_frontier U" and hg: "\x. x \ S \ h x = g x" proof (rule extend_map_affine_to_sphere1 [OF \finite K\ \affine T\ contg gim, of S "\ ` K"]) show cloTS: "closedin (top_of_set T) S" by (simp add: \compact S\ \S \ T\ closed_subset compact_imp_closed) show "\C. \C \ components (T - S); C \ K \ {}\ \ C \ \ ` K \ {}" using \ components_eq by blast qed (use K in auto) show ?thesis proof show *: "\ ` K \ L" using \ by blast show "finite (\ ` K)" by (simp add: K) show "\ ` K \ T" by clarify (meson \ Diff_iff contra_subsetD in_components_subset) show "continuous_on (T - \ ` K) h" by (rule conth) show "disjnt (\ ` K) S" using K apply (auto simp: disjnt_def) by (metis \ DiffD2 UnionI Union_components) qed (simp_all add: him hg gf) qed proposition extend_map_affine_to_sphere_cofinite_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "convex U" "bounded U" "affine T" "S \ T" and aff: "aff_dim T \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ rel_frontier U" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ rel_frontier U" "\x. x \ S \ g x = f x" proof (cases "S = {}") case True show ?thesis proof (cases "rel_frontier U = {}") case True with aff have "aff_dim T \ 0" apply (simp add: rel_frontier_eq_empty) using affine_bounded_eq_lowdim \bounded U\ order_trans by auto with aff_dim_geq [of T] consider "aff_dim T = -1" | "aff_dim T = 0" by linarith then show ?thesis proof cases assume "aff_dim T = -1" then have "T = {}" by (simp add: aff_dim_empty) then show ?thesis by (rule_tac K="{}" in that) auto next assume "aff_dim T = 0" then obtain a where "T = {a}" using aff_dim_eq_0 by blast then have "a \ L" using dis [of "{a}"] \S = {}\ by (auto simp: in_components_self) with \S = {}\ \T = {a}\ show ?thesis by (rule_tac K="{a}" and g=f in that) auto qed next case False then obtain y where "y \ rel_frontier U" by auto with \S = {}\ show ?thesis by (rule_tac K="{}" and g="\x. y" in that) (auto simp: continuous_on_const) qed next case False have "bounded S" by (simp add: assms compact_imp_bounded) then obtain b where b: "S \ cbox (-b) b" using bounded_subset_cbox_symmetric by blast define LU where "LU \ L \ (\ {C \ components (T - S). \bounded C} - cbox (-(b+One)) (b+One))" obtain K g where "finite K" "K \ LU" "K \ T" "disjnt K S" and contg: "continuous_on (T - K) g" and gim: "g ` (T - K) \ rel_frontier U" and gf: "\x. x \ S \ g x = f x" proof (rule extend_map_affine_to_sphere2 [OF SUT aff contf fim]) show "C \ LU \ {}" if "C \ components (T - S)" for C proof (cases "bounded C") case True with dis that show ?thesis unfolding LU_def by fastforce next case False then have "\ bounded (\{C \ components (T - S). \ bounded C})" by (metis (no_types, lifting) Sup_upper bounded_subset mem_Collect_eq that) then show ?thesis apply (clarsimp simp: LU_def Int_Un_distrib Diff_Int_distrib Int_UN_distrib) by (metis (no_types, lifting) False Sup_upper bounded_cbox bounded_subset inf.orderE mem_Collect_eq that) qed qed blast have *: False if "x \ cbox (- b - m *\<^sub>R One) (b + m *\<^sub>R One)" "x \ box (- b - n *\<^sub>R One) (b + n *\<^sub>R One)" "0 \ m" "m < n" "n \ 1" for m n x using that by (auto simp: mem_box algebra_simps) have "disjoint_family_on (\d. frontier (cbox (- b - d *\<^sub>R One) (b + d *\<^sub>R One))) {1 / 2..1}" by (auto simp: disjoint_family_on_def neq_iff frontier_def dest: *) then obtain d where d12: "1/2 \ d" "d \ 1" and ddis: "disjnt K (frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One)))" using disjoint_family_elem_disjnt [of "{1/2..1::real}" K "\d. frontier (cbox (-(b + d *\<^sub>R One)) (b + d *\<^sub>R One))"] by (auto simp: \finite K\) define c where "c \ b + d *\<^sub>R One" have cbsub: "cbox (-b) b \ box (-c) c" "cbox (-b) b \ cbox (-c) c" "cbox (-c) c \ cbox (-(b+One)) (b+One)" using d12 by (simp_all add: subset_box c_def inner_diff_left inner_left_distrib) have clo_cT: "closed (cbox (- c) c \ T)" using affine_closed \affine T\ by blast have cT_ne: "cbox (- c) c \ T \ {}" using \S \ {}\ \S \ T\ b cbsub by fastforce have S_sub_cc: "S \ cbox (- c) c" using \cbox (- b) b \ cbox (- c) c\ b by auto show ?thesis proof show "finite (K \ cbox (-(b+One)) (b+One))" using \finite K\ by blast show "K \ cbox (- (b + One)) (b + One) \ L" using \K \ LU\ by (auto simp: LU_def) show "K \ cbox (- (b + One)) (b + One) \ T" using \K \ T\ by auto show "disjnt (K \ cbox (- (b + One)) (b + One)) S" using \disjnt K S\ by (simp add: disjnt_def disjoint_eq_subset_Compl inf.coboundedI1) have cloTK: "closest_point (cbox (- c) c \ T) x \ T - K" if "x \ T" and Knot: "x \ K \ x \ cbox (- b - One) (b + One)" for x proof (cases "x \ cbox (- c) c") case True with \x \ T\ show ?thesis using cbsub(3) Knot by (force simp: closest_point_self) next case False have clo_in_rf: "closest_point (cbox (- c) c \ T) x \ rel_frontier (cbox (- c) c \ T)" proof (intro closest_point_in_rel_frontier [OF clo_cT cT_ne] DiffI notI) have "T \ interior (cbox (- c) c) \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then show "x \ affine hull (cbox (- c) c \ T)" by (simp add: Int_commute affine_hull_affine_Int_nonempty_interior \affine T\ hull_inc that(1)) next show "False" if "x \ rel_interior (cbox (- c) c \ T)" proof - have "interior (cbox (- c) c) \ T \ {}" using \S \ {}\ \S \ T\ b cbsub(1) by fastforce then have "affine hull (T \ cbox (- c) c) = T" using affine_hull_convex_Int_nonempty_interior [of T "cbox (- c) c"] by (simp add: affine_imp_convex \affine T\ inf_commute) then show ?thesis by (meson subsetD le_inf_iff rel_interior_subset that False) qed qed have "closest_point (cbox (- c) c \ T) x \ K" proof assume inK: "closest_point (cbox (- c) c \ T) x \ K" have "\x. x \ K \ x \ frontier (cbox (- (b + d *\<^sub>R One)) (b + d *\<^sub>R One))" by (metis ddis disjnt_iff) then show False by (metis DiffI Int_iff \affine T\ cT_ne c_def clo_cT clo_in_rf closest_point_in_set convex_affine_rel_frontier_Int convex_box(1) empty_iff frontier_cbox inK interior_cbox) qed then show ?thesis using cT_ne clo_cT closest_point_in_set by blast qed show "continuous_on (T - K \ cbox (- (b + One)) (b + One)) (g \ closest_point (cbox (-c) c \ T))" apply (intro continuous_on_compose continuous_on_closest_point continuous_on_subset [OF contg]) apply (simp_all add: clo_cT affine_imp_convex \affine T\ convex_Int cT_ne) using cloTK by blast have "g (closest_point (cbox (- c) c \ T) x) \ rel_frontier U" if "x \ T" "x \ K \ x \ cbox (- b - One) (b + One)" for x apply (rule gim [THEN subsetD]) using that cloTK by blast then show "(g \ closest_point (cbox (- c) c \ T)) ` (T - K \ cbox (- (b + One)) (b + One)) \ rel_frontier U" by force show "\x. x \ S \ (g \ closest_point (cbox (- c) c \ T)) x = f x" by simp (metis (mono_tags, lifting) IntI \S \ T\ cT_ne clo_cT closest_point_refl gf subsetD S_sub_cc) qed qed corollary extend_map_affine_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes SUT: "compact S" "affine T" "S \ T" and aff: "aff_dim T \ DIM('b)" and "0 \ r" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(T - S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "K \ T" "disjnt K S" "continuous_on (T - K) g" "g ` (T - K) \ sphere a r" "\x. x \ S \ g x = f x" proof (cases "r = 0") case True with fim show ?thesis by (rule_tac K="{}" and g = "\x. a" in that) (auto simp: continuous_on_const) next case False with assms have "0 < r" by auto then have "aff_dim T \ aff_dim (cball a r)" by (simp add: aff aff_dim_cball) then show ?thesis apply (rule extend_map_affine_to_sphere_cofinite_gen [OF \compact S\ convex_cball bounded_cball \affine T\ \S \ T\ _ contf]) using fim apply (auto simp: assms False that dest: dis) done qed corollary extend_map_UNIV_to_sphere_cofinite: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. \C \ components(- S); bounded C\ \ C \ L \ {}" obtains K g where "finite K" "K \ L" "disjnt K S" "continuous_on (- K) g" "g ` (- K) \ sphere a r" "\x. x \ S \ g x = f x" apply (rule extend_map_affine_to_sphere_cofinite [OF \compact S\ affine_UNIV subset_UNIV _ \0 \ r\ contf fim dis]) apply (auto simp: assms that Compl_eq_Diff_UNIV [symmetric]) done corollary extend_map_UNIV_to_sphere_no_bounded_component: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes aff: "DIM('a) \ DIM('b)" and "0 \ r" and SUT: "compact S" and contf: "continuous_on S f" and fim: "f ` S \ sphere a r" and dis: "\C. C \ components(- S) \ \ bounded C" obtains g where "continuous_on UNIV g" "g ` UNIV \ sphere a r" "\x. x \ S \ g x = f x" apply (rule extend_map_UNIV_to_sphere_cofinite [OF aff \0 \ r\ \compact S\ contf fim, of "{}"]) apply (auto simp: that dest: dis) done theorem Borsuk_separation_theorem_gen: fixes S :: "'a::euclidean_space set" assumes "compact S" shows "(\c \ components(- S). \bounded c) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L [rule_format]: ?lhs show ?rhs proof clarify fix f :: "'a \ 'a" assume contf: "continuous_on S f" and fim: "f ` S \ sphere 0 1" obtain g where contg: "continuous_on UNIV g" and gim: "range g \ sphere 0 1" and gf: "\x. x \ S \ g x = f x" by (rule extend_map_UNIV_to_sphere_no_bounded_component [OF _ _ \compact S\ contf fim L]) auto then obtain c where c: "homotopic_with_canon (\h. True) UNIV (sphere 0 1) g (\x. c)" using contractible_UNIV nullhomotopic_from_contractible by blast then show "\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)" by (metis assms compact_imp_closed contf contg contractible_empty fim gf gim nullhomotopic_from_contractible nullhomotopic_into_sphere_extension) qed next assume R [rule_format]: ?rhs show ?lhs unfolding components_def proof clarify fix a assume "a \ S" and a: "bounded (connected_component_set (- S) a)" have cont: "continuous_on S (\x. inverse(norm(x - a)) *\<^sub>R (x - a))" apply (intro continuous_intros) using \a \ S\ by auto have im: "(\x. inverse(norm(x - a)) *\<^sub>R (x - a)) ` S \ sphere 0 1" by clarsimp (metis \a \ S\ eq_iff_diff_eq_0 left_inverse norm_eq_zero) show False using R cont im Borsuk_map_essential_bounded_component [OF \compact S\ \a \ S\] a by blast qed qed corollary Borsuk_separation_theorem: fixes S :: "'a::euclidean_space set" assumes "compact S" and 2: "2 \ DIM('a)" shows "connected(- S) \ (\f. continuous_on S f \ f ` S \ sphere (0::'a) 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) f (\x. c)))" (is "?lhs = ?rhs") proof assume L: ?lhs show ?rhs proof (cases "S = {}") case True then show ?thesis by auto next case False then have "(\c\components (- S). \ bounded c)" by (metis L assms(1) bounded_empty cobounded_imp_unbounded compact_imp_bounded in_components_maximal order_refl) then show ?thesis by (simp add: Borsuk_separation_theorem_gen [OF \compact S\]) qed next assume R: ?rhs then show ?lhs apply (simp add: Borsuk_separation_theorem_gen [OF \compact S\, symmetric]) apply (auto simp: components_def connected_iff_eq_connected_component_set) using connected_component_in apply fastforce using cobounded_unique_unbounded_component [OF _ 2, of "-S"] \compact S\ compact_eq_bounded_closed by fastforce qed lemma homotopy_eqv_separation: fixes S :: "'a::euclidean_space set" and T :: "'a set" assumes "S homotopy_eqv T" and "compact S" and "compact T" shows "connected(- S) \ connected(- T)" proof - consider "DIM('a) = 1" | "2 \ DIM('a)" by (metis DIM_ge_Suc0 One_nat_def Suc_1 dual_order.antisym not_less_eq_eq) then show ?thesis proof cases case 1 then show ?thesis using bounded_connected_Compl_1 compact_imp_bounded homotopy_eqv_empty1 homotopy_eqv_empty2 assms by metis next case 2 with assms show ?thesis by (simp add: Borsuk_separation_theorem homotopy_eqv_cohomotopic_triviality_null) qed qed proposition Jordan_Brouwer_separation: fixes S :: "'a::euclidean_space set" and a::'a assumes hom: "S homeomorphic sphere a r" and "0 < r" shows "\ connected(- S)" proof - have "- sphere a r \ ball a r \ {}" using \0 < r\ by (simp add: Int_absorb1 subset_eq) moreover have eq: "- sphere a r - ball a r = - cball a r" by auto have "- cball a r \ {}" proof - have "frontier (cball a r) \ {}" using \0 < r\ by auto then show ?thesis by (metis frontier_complement frontier_empty) qed with eq have "- sphere a r - ball a r \ {}" by auto moreover have "connected (- S) = connected (- sphere a r)" proof (rule homotopy_eqv_separation) show "S homotopy_eqv sphere a r" using hom homeomorphic_imp_homotopy_eqv by blast show "compact (sphere a r)" by simp then show " compact S" using hom homeomorphic_compactness by blast qed ultimately show ?thesis using connected_Int_frontier [of "- sphere a r" "ball a r"] by (auto simp: \0 < r\) qed proposition Jordan_Brouwer_frontier: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and T: "T \ components(- S)" and 2: "2 \ DIM('a)" shows "frontier T = S" proof (cases r rule: linorder_cases) assume "r < 0" with S T show ?thesis by auto next assume "r = 0" with S T card_eq_SucD obtain b where "S = {b}" by (auto simp: homeomorphic_finite [of "{a}" S]) have "components (- {b}) = { -{b}}" using T \S = {b}\ by (auto simp: components_eq_sing_iff connected_punctured_universe 2) with T show ?thesis by (metis \S = {b}\ cball_trivial frontier_cball frontier_complement singletonD sphere_trivial) next assume "r > 0" have "compact S" using homeomorphic_compactness compact_sphere S by blast show ?thesis proof (rule frontier_minimal_separating_closed) show "closed S" using \compact S\ compact_eq_bounded_closed by blast show "\ connected (- S)" using Jordan_Brouwer_separation S \0 < r\ by blast obtain f g where hom: "homeomorphism S (sphere a r) f g" using S by (auto simp: homeomorphic_def) show "connected (- T)" if "closed T" "T \ S" for T proof - have "f ` T \ sphere a r" using \T \ S\ hom homeomorphism_image1 by blast moreover have "f ` T \ sphere a r" using \T \ S\ hom by (metis homeomorphism_image2 homeomorphism_of_subsets order_refl psubsetE) ultimately have "f ` T \ sphere a r" by blast then have "connected (- f ` T)" by (rule psubset_sphere_Compl_connected [OF _ \0 < r\ 2]) moreover have "compact T" using \compact S\ bounded_subset compact_eq_bounded_closed that by blast moreover then have "compact (f ` T)" by (meson compact_continuous_image continuous_on_subset hom homeomorphism_def psubsetE \T \ S\) moreover have "T homotopy_eqv f ` T" by (meson \f ` T \ sphere a r\ dual_order.strict_implies_order hom homeomorphic_def homeomorphic_imp_homotopy_eqv homeomorphism_of_subsets \T \ S\) ultimately show ?thesis using homotopy_eqv_separation [of T "f`T"] by blast qed qed (rule T) qed proposition Jordan_Brouwer_nonseparation: fixes S :: "'a::euclidean_space set" and a::'a assumes S: "S homeomorphic sphere a r" and "T \ S" and 2: "2 \ DIM('a)" shows "connected(- T)" proof - have *: "connected(C \ (S - T))" if "C \ components(- S)" for C proof (rule connected_intermediate_closure) show "connected C" using in_components_connected that by auto have "S = frontier C" using "2" Jordan_Brouwer_frontier S that by blast with closure_subset show "C \ (S - T) \ closure C" by (auto simp: frontier_def) qed auto have "components(- S) \ {}" by (metis S bounded_empty cobounded_imp_unbounded compact_eq_bounded_closed compact_sphere components_eq_empty homeomorphic_compactness) then have "- T = (\C \ components(- S). C \ (S - T))" using Union_components [of "-S"] \T \ S\ by auto then show ?thesis apply (rule ssubst) apply (rule connected_Union) using \T \ S\ apply (auto simp: *) done qed subsection\ Invariance of domain and corollaries\ lemma invariance_of_domain_ball: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on (cball a r) f" and "0 < r" and inj: "inj_on f (cball a r)" shows "open(f ` ball a r)" proof (cases "DIM('a) = 1") case True obtain h::"'a\real" and k where "linear h" "linear k" "h ` UNIV = UNIV" "k ` UNIV = UNIV" "\x. norm(h x) = norm x" "\x. norm(k x) = norm x" "\x. k(h x) = x" "\x. h(k x) = x" apply (rule isomorphisms_UNIV_UNIV [where 'M='a and 'N=real]) using True apply force by (metis UNIV_I UNIV_eq_I imageI) have cont: "continuous_on S h" "continuous_on T k" for S T by (simp_all add: \linear h\ \linear k\ linear_continuous_on linear_linear) have "continuous_on (h ` cball a r) (h \ f \ k)" apply (intro continuous_on_compose cont continuous_on_subset [OF contf]) apply (auto simp: \\x. k (h x) = x\) done moreover have "is_interval (h ` cball a r)" by (simp add: is_interval_connected_1 \linear h\ linear_continuous_on linear_linear connected_continuous_image) moreover have "inj_on (h \ f \ k) (h ` cball a r)" using inj by (simp add: inj_on_def) (metis \\x. k (h x) = x\) ultimately have *: "\T. \open T; T \ h ` cball a r\ \ open ((h \ f \ k) ` T)" using injective_eq_1d_open_map_UNIV by blast have "open ((h \ f \ k) ` (h ` ball a r))" by (rule *) (auto simp: \linear h\ \range h = UNIV\ open_surjective_linear_image) then have "open ((h \ f) ` ball a r)" by (simp add: image_comp \\x. k (h x) = x\ cong: image_cong) then show ?thesis apply (simp only: image_comp [symmetric]) apply (metis open_bijective_linear_image_eq \linear h\ \\x. k (h x) = x\ \range h = UNIV\ bijI inj_on_def) done next case False then have 2: "DIM('a) \ 2" by (metis DIM_ge_Suc0 One_nat_def Suc_1 antisym not_less_eq_eq) have fimsub: "f ` ball a r \ - f ` sphere a r" using inj by clarsimp (metis inj_onD less_eq_real_def mem_cball order_less_irrefl) have hom: "f ` sphere a r homeomorphic sphere a r" by (meson compact_sphere contf continuous_on_subset homeomorphic_compact homeomorphic_sym inj inj_on_subset sphere_cball) then have nconn: "\ connected (- f ` sphere a r)" by (rule Jordan_Brouwer_separation) (auto simp: \0 < r\) obtain C where C: "C \ components (- f ` sphere a r)" and "bounded C" apply (rule cobounded_has_bounded_component [OF _ nconn]) apply (simp_all add: 2) by (meson compact_imp_bounded compact_continuous_image_eq compact_sphere contf inj sphere_cball) moreover have "f ` (ball a r) = C" proof have "C \ {}" by (rule in_components_nonempty [OF C]) show "C \ f ` ball a r" proof (rule ccontr) assume nonsub: "\ C \ f ` ball a r" have "- f ` cball a r \ C" proof (rule components_maximal [OF C]) have "f ` cball a r homeomorphic cball a r" using compact_cball contf homeomorphic_compact homeomorphic_sym inj by blast then show "connected (- f ` cball a r)" by (auto intro: connected_complement_homeomorphic_convex_compact 2) show "- f ` cball a r \ - f ` sphere a r" by auto then show "C \ - f ` cball a r \ {}" using \C \ {}\ in_components_subset [OF C] nonsub using image_iff by fastforce qed then have "bounded (- f ` cball a r)" using bounded_subset \bounded C\ by auto then have "\ bounded (f ` cball a r)" using cobounded_imp_unbounded by blast then show "False" using compact_continuous_image [OF contf] compact_cball compact_imp_bounded by blast qed with \C \ {}\ have "C \ f ` ball a r \ {}" by (simp add: inf.absorb_iff1) then show "f ` ball a r \ C" by (metis components_maximal [OF C _ fimsub] connected_continuous_image ball_subset_cball connected_ball contf continuous_on_subset) qed moreover have "open (- f ` sphere a r)" using hom compact_eq_bounded_closed compact_sphere homeomorphic_compactness by blast ultimately show ?thesis using open_components by blast qed text\Proved by L. E. J. Brouwer (1912)\ theorem invariance_of_domain: fixes f :: "'a \ 'a::euclidean_space" assumes "continuous_on S f" "open S" "inj_on f S" shows "open(f ` S)" unfolding open_subopen [of "f`S"] proof clarify fix a assume "a \ S" obtain \ where "\ > 0" and \: "cball a \ \ S" using \open S\ \a \ S\ open_contains_cball_eq by blast show "\T. open T \ f a \ T \ T \ f ` S" proof (intro exI conjI) show "open (f ` (ball a \))" by (meson \ \0 < \\ assms continuous_on_subset inj_on_subset invariance_of_domain_ball) show "f a \ f ` ball a \" by (simp add: \0 < \\) show "f ` ball a \ \ f ` S" using \ ball_subset_cball by blast qed qed lemma inv_of_domain_ss0: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and dimS: "dim S = DIM('b::euclidean_space)" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - have "U \ S" using ope openin_imp_subset by blast have "(UNIV::'b set) homeomorphic S" by (simp add: \subspace S\ dimS homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism (UNIV::'b set) S h k" using homeomorphic_def by blast have homkh: "homeomorphism S (k ` S) k h" using homhk homeomorphism_image2 homeomorphism_sym by fastforce have "open ((k \ f \ h) ` k ` U)" proof (rule invariance_of_domain) show "continuous_on (k ` U) (k \ f \ h)" proof (intro continuous_intros) show "continuous_on (k ` U) h" by (meson continuous_on_subset [OF homeomorphism_cont1 [OF homhk]] top_greatest) show "continuous_on (h ` k ` U) f" apply (rule continuous_on_subset [OF contf], clarify) apply (metis homhk homeomorphism_def ope openin_imp_subset rev_subsetD) done show "continuous_on (f ` h ` k ` U) k" apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF homhk]]) using fim homhk homeomorphism_apply2 ope openin_subset by fastforce qed have ope_iff: "\T. open T \ openin (top_of_set (k ` S)) T" using homhk homeomorphism_image2 open_openin by fastforce show "open (k ` U)" by (simp add: ope_iff homeomorphism_imp_open_map [OF homkh ope]) show "inj_on (k \ f \ h) (k ` U)" apply (clarsimp simp: inj_on_def) by (metis subsetD fim homeomorphism_apply2 [OF homhk] image_subset_iff inj_on_eq_iff injf \U \ S\) qed moreover have eq: "f ` U = h ` (k \ f \ h \ k) ` U" unfolding image_comp [symmetric] using \U \ S\ fim by (metis homeomorphism_image2 homeomorphism_of_subsets homkh subset_image_iff) ultimately show ?thesis by (metis (no_types, hide_lams) homeomorphism_imp_open_map homhk image_comp open_openin subtopology_UNIV) qed lemma inv_of_domain_ss1: fixes f :: "'a \ 'a::euclidean_space" assumes contf: "continuous_on U f" and injf: "inj_on f U" and fim: "f ` U \ S" and "subspace S" and ope: "openin (top_of_set S) U" shows "openin (top_of_set S) (f ` U)" proof - define S' where "S' \ {y. \x \ S. orthogonal x y}" have "subspace S'" by (simp add: S'_def subspace_orthogonal_to_vectors) define g where "g \ \z::'a*'a. ((f \ fst)z, snd z)" have "openin (top_of_set (S \ S')) (g ` (U \ S'))" proof (rule inv_of_domain_ss0) show "continuous_on (U \ S') g" apply (simp add: g_def) apply (intro continuous_intros continuous_on_compose2 [OF contf continuous_on_fst], auto) done show "g ` (U \ S') \ S \ S'" using fim by (auto simp: g_def) show "inj_on g (U \ S')" using injf by (auto simp: g_def inj_on_def) show "subspace (S \ S')" by (simp add: \subspace S'\ \subspace S\ subspace_Times) show "openin (top_of_set (S \ S')) (U \ S')" by (simp add: openin_Times [OF ope]) have "dim (S \ S') = dim S + dim S'" by (simp add: \subspace S'\ \subspace S\ dim_Times) also have "... = DIM('a)" using dim_subspace_orthogonal_to_vectors [OF \subspace S\ subspace_UNIV] by (simp add: add.commute S'_def) finally show "dim (S \ S') = DIM('a)" . qed moreover have "g ` (U \ S') = f ` U \ S'" by (auto simp: g_def image_iff) moreover have "0 \ S'" using \subspace S'\ subspace_affine by blast ultimately show ?thesis by (auto simp: openin_Times_eq) qed corollary invariance_of_domain_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and VU: "dim V \ dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof - obtain V' where "subspace V'" "V' \ U" "dim V' = dim V" using choose_subspace_of_subspace [OF VU] by (metis span_eq_iff \subspace U\) then have "V homeomorphic V'" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V V' h k" using homeomorphic_def by blast have eq: "f ` S = k ` (h \ f) ` S" proof - have "k ` h ` f ` S = f ` S" by (meson fim homeomorphism_def homeomorphism_of_subsets homhk subset_refl) then show ?thesis by (simp add: image_comp) qed show ?thesis unfolding eq proof (rule homeomorphism_imp_open_map) show homkh: "homeomorphism V' V k h" by (simp add: homeomorphism_symD homhk) have hfV': "(h \ f) ` S \ V'" using fim homeomorphism_image1 homhk by fastforce moreover have "openin (top_of_set U) ((h \ f) ` S)" proof (rule inv_of_domain_ss1) show "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) show "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply2 [OF homkh] image_subset_iff inj_onD injf) show "(h \ f) ` S \ U" using \V' \ U\ hfV' by auto qed (auto simp: assms) ultimately show "openin (top_of_set V') ((h \ f) ` S)" using openin_subset_trans \V' \ U\ by force qed qed corollary invariance_of_dimension_subspaces: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and "subspace U" "subspace V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "dim U \ dim V" proof - have "False" if "dim V < dim U" proof - obtain T where "subspace T" "T \ U" "dim T = dim V" using choose_subspace_of_subspace [of "dim V" U] by (metis \dim V < dim U\ assms(2) order.strict_implies_order span_eq_iff) then have "V homeomorphic T" by (simp add: \subspace V\ homeomorphic_subspaces) then obtain h k where homhk: "homeomorphism V T h k" using homeomorphic_def by blast have "continuous_on S (h \ f)" by (meson contf continuous_on_compose continuous_on_subset fim homeomorphism_cont1 homhk) moreover have "(h \ f) ` S \ U" using \T \ U\ fim homeomorphism_image1 homhk by fastforce moreover have "inj_on (h \ f) S" apply (clarsimp simp: inj_on_def) by (metis fim homeomorphism_apply1 homhk image_subset_iff inj_onD injf) ultimately have ope_hf: "openin (top_of_set U) ((h \ f) ` S)" using invariance_of_domain_subspaces [OF ope \subspace U\ \subspace U\] by blast have "(h \ f) ` S \ T" using fim homeomorphism_image1 homhk by fastforce then have "dim ((h \ f) ` S) \ dim T" by (rule dim_subset) also have "dim ((h \ f) ` S) = dim U" using \S \ {}\ \subspace U\ by (blast intro: dim_openin ope_hf) finally show False using \dim V < dim U\ \dim T = dim V\ by simp qed then show ?thesis using not_less by blast qed corollary invariance_of_domain_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" "aff_dim V \ aff_dim U" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" shows "openin (top_of_set V) (f ` S)" proof (cases "S = {}") case True then show ?thesis by auto next case False obtain a b where "a \ S" "a \ U" "b \ V" using False fim ope openin_contains_cball by fastforce have "openin (top_of_set ((+) (- b) ` V)) (((+) (- b) \ f \ (+) a) ` (+) (- a) ` S)" proof (rule invariance_of_domain_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "dim ((+) (- b) ` V) \ dim ((+) (- a) ` U)" by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed then show ?thesis by (metis (no_types, lifting) homeomorphism_imp_open_map homeomorphism_translation image_comp translation_galois) qed corollary invariance_of_dimension_affine_sets: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes ope: "openin (top_of_set U) S" and aff: "affine U" "affine V" and contf: "continuous_on S f" and fim: "f ` S \ V" and injf: "inj_on f S" and "S \ {}" shows "aff_dim U \ aff_dim V" proof - obtain a b where "a \ S" "a \ U" "b \ V" using \S \ {}\ fim ope openin_contains_cball by fastforce have "dim ((+) (- a) ` U) \ dim ((+) (- b) ` V)" proof (rule invariance_of_dimension_subspaces) show "openin (top_of_set ((+) (- a) ` U)) ((+) (- a) ` S)" by (metis ope homeomorphism_imp_open_map homeomorphism_translation translation_galois) show "subspace ((+) (- a) ` U)" by (simp add: \a \ U\ affine_diffs_subspace_subtract \affine U\ cong: image_cong_simp) show "subspace ((+) (- b) ` V)" by (simp add: \b \ V\ affine_diffs_subspace_subtract \affine V\ cong: image_cong_simp) show "continuous_on ((+) (- a) ` S) ((+) (- b) \ f \ (+) a)" by (metis contf continuous_on_compose homeomorphism_cont2 homeomorphism_translation translation_galois) show "((+) (- b) \ f \ (+) a) ` (+) (- a) ` S \ (+) (- b) ` V" using fim by auto show "inj_on ((+) (- b) \ f \ (+) a) ((+) (- a) ` S)" by (auto simp: inj_on_def) (meson inj_onD injf) qed (use \S \ {}\ in auto) then show ?thesis by (metis \a \ U\ \b \ V\ aff_dim_eq_dim affine_hull_eq aff of_nat_le_iff) qed corollary invariance_of_dimension: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and "open S" and injf: "inj_on f S" and "S \ {}" shows "DIM('a) \ DIM('b)" using invariance_of_dimension_subspaces [of UNIV S UNIV f] assms by auto corollary continuous_injective_image_subspace_dim_le: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "subspace S" "subspace T" and contf: "continuous_on S f" and fim: "f ` S \ T" and injf: "inj_on f S" shows "dim S \ dim T" apply (rule invariance_of_dimension_subspaces [of S S _ f]) using assms by (auto simp: subspace_affine) lemma invariance_of_dimension_convex_domain: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "convex S" and contf: "continuous_on S f" and fim: "f ` S \ affine hull T" and injf: "inj_on f S" shows "aff_dim S \ aff_dim T" proof (cases "S = {}") case True then show ?thesis by (simp add: aff_dim_geq) next case False have "aff_dim (affine hull S) \ aff_dim (affine hull T)" proof (rule invariance_of_dimension_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "f ` rel_interior S \ affine hull T" using fim rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast show "rel_interior S \ {}" by (simp add: False \convex S\ rel_interior_eq_empty) qed auto then show ?thesis by simp qed lemma homeomorphic_convex_sets_le: assumes "convex S" "S homeomorphic T" shows "aff_dim S \ aff_dim T" proof - obtain h k where homhk: "homeomorphism S T h k" using homeomorphic_def assms by blast show ?thesis proof (rule invariance_of_dimension_convex_domain [OF \convex S\]) show "continuous_on S h" using homeomorphism_def homhk by blast show "h ` S \ affine hull T" by (metis homeomorphism_def homhk hull_subset) show "inj_on h S" by (meson homeomorphism_apply1 homhk inj_on_inverseI) qed qed lemma homeomorphic_convex_sets: assumes "convex S" "convex T" "S homeomorphic T" shows "aff_dim S = aff_dim T" by (meson assms dual_order.antisym homeomorphic_convex_sets_le homeomorphic_sym) lemma homeomorphic_convex_compact_sets_eq: assumes "convex S" "compact S" "convex T" "compact T" shows "S homeomorphic T \ aff_dim S = aff_dim T" by (meson assms homeomorphic_convex_compact_sets homeomorphic_convex_sets) lemma invariance_of_domain_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "open(f ` S)" using invariance_of_domain_subspaces [of UNIV S UNIV f] assms by auto lemma injective_into_1d_imp_open_map_UNIV: fixes f :: "'a::euclidean_space \ real" assumes "open T" "continuous_on S f" "inj_on f S" "T \ S" shows "open (f ` T)" apply (rule invariance_of_domain_gen [OF \open T\]) using assms apply (auto simp: elim: continuous_on_subset subset_inj_on) done lemma continuous_on_inverse_open: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" and gf: "\x. x \ S \ g(f x) = x" shows "continuous_on (f ` S) g" proof (clarsimp simp add: continuous_openin_preimage_eq) fix T :: "'a set" assume "open T" have eq: "f ` S \ g -` T = f ` (S \ T)" by (auto simp: gf) show "openin (top_of_set (f ` S)) (f ` S \ g -` T)" apply (subst eq) apply (rule open_openin_trans) apply (rule invariance_of_domain_gen) using assms apply auto using inj_on_inverseI apply auto[1] by (metis \open T\ continuous_on_subset inj_onI inj_on_subset invariance_of_domain_gen openin_open openin_open_eq) qed lemma invariance_of_domain_homeomorphism: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" obtains g where "homeomorphism S (f ` S) f g" proof show "homeomorphism S (f ` S) f (inv_into S f)" by (simp add: assms continuous_on_inverse_open homeomorphism_def) qed corollary invariance_of_domain_homeomorphic: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "open S" "continuous_on S f" "DIM('b) \ DIM('a)" "inj_on f S" shows "S homeomorphic (f ` S)" using invariance_of_domain_homeomorphism [OF assms] by (meson homeomorphic_def) lemma continuous_image_subset_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "continuous_on S f" "inj_on f S" "DIM('b) \ DIM('a)" shows "f ` (interior S) \ interior(f ` S)" apply (rule interior_maximal) apply (simp add: image_mono interior_subset) apply (rule invariance_of_domain_gen) using assms apply (auto simp: subset_inj_on interior_subset continuous_on_subset) done lemma homeomorphic_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and dimeq: "DIM('a) = DIM('b)" shows "(interior S) homeomorphic (interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp have gim: "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp show "homeomorphism (interior S) (interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show "\x. x \ interior S \ g (f x) = x" by (meson \\x\S. f x \ T \ g (f x) = x\ subsetD interior_subset) have "interior T \ f ` interior S" proof fix x assume "x \ interior T" then have "g x \ interior S" using gim by blast then show "x \ f ` interior S" by (metis T \x \ interior T\ image_iff interior_subset subsetCE) qed then show "f ` interior S = interior T" using fim by blast show "continuous_on (interior S) f" by (metis interior_subset continuous_on_subset contf) show "\y. y \ interior T \ f (g y) = y" by (meson T subsetD interior_subset) have "interior S \ g ` interior T" proof fix x assume "x \ interior S" then have "f x \ interior T" using fim by blast then show "x \ g ` interior T" by (metis S \x \ interior S\ image_iff interior_subset subsetCE) qed then show "g ` interior T = interior S" using gim by blast show "continuous_on (interior T) g" by (metis interior_subset continuous_on_subset contg) qed qed lemma homeomorphic_open_imp_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "open S" "S \ {}" "open T" "T \ {}" shows "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis inj_onI invariance_of_dimension) done proposition homeomorphic_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "interior S = {} \ interior T = {}" shows "(interior S) homeomorphic (interior T)" proof (cases "interior T = {}") case True with assms show ?thesis by auto next case False then have "DIM('a) = DIM('b)" using assms apply (simp add: homeomorphic_minimal) apply (rule order_antisym; metis continuous_on_subset inj_onI inj_on_subset interior_subset invariance_of_dimension open_interior) done then show ?thesis by (rule homeomorphic_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_frontiers_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" and dimeq: "DIM('a) = DIM('b)" shows "(frontier S) homeomorphic (frontier T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have "g ` interior T \ interior S" using continuous_image_subset_interior [OF contg \inj_on g T\] dimeq gTS by simp then have fim: "f ` frontier S \ frontier T" apply (simp add: frontier_def) using continuous_image_subset_interior assms(2) assms(3) S by auto have "f ` interior S \ interior T" using continuous_image_subset_interior [OF contf \inj_on f S\] dimeq fST by simp then have gim: "g ` frontier T \ frontier S" apply (simp add: frontier_def) using continuous_image_subset_interior T assms(2) assms(3) by auto show "homeomorphism (frontier S) (frontier T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ frontier S \ g (f x) = x" by (simp add: S assms(2) frontier_def) show fg: "\y. y \ frontier T \ f (g y) = y" by (simp add: T assms(3) frontier_def) have "frontier T \ f ` frontier S" proof fix x assume "x \ frontier T" then have "g x \ frontier S" using gim by blast then show "x \ f ` frontier S" by (metis fg \x \ frontier T\ imageI) qed then show "f ` frontier S = frontier T" using fim by blast show "continuous_on (frontier S) f" by (metis Diff_subset assms(2) closure_eq contf continuous_on_subset frontier_def) have "frontier S \ g ` frontier T" proof fix x assume "x \ frontier S" then have "f x \ frontier T" using fim by blast then show "x \ g ` frontier T" by (metis gf \x \ frontier S\ imageI) qed then show "g ` frontier T = frontier S" using gim by blast show "continuous_on (frontier T) g" by (metis Diff_subset assms(3) closure_closed contg continuous_on_subset frontier_def) qed qed lemma homeomorphic_frontiers: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "closed S" "closed T" "interior S = {} \ interior T = {}" shows "(frontier S) homeomorphic (frontier T)" proof (cases "interior T = {}") case True then show ?thesis by (metis Diff_empty assms closure_eq frontier_def) next case False show ?thesis apply (rule homeomorphic_frontiers_same_dimension) apply (simp_all add: assms) using False assms homeomorphic_interiors homeomorphic_open_imp_same_dimension by blast qed lemma continuous_image_subset_rel_interior: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and TS: "aff_dim T \ aff_dim S" shows "f ` (rel_interior S) \ rel_interior(f ` S)" proof (rule rel_interior_maximal) show "f ` rel_interior S \ f ` S" by(simp add: image_mono rel_interior_subset) show "openin (top_of_set (affine hull f ` S)) (f ` rel_interior S)" proof (rule invariance_of_domain_affine_sets) show "openin (top_of_set (affine hull S)) (rel_interior S)" by (simp add: openin_rel_interior) show "aff_dim (affine hull f ` S) \ aff_dim (affine hull S)" by (metis aff_dim_affine_hull aff_dim_subset fim TS order_trans) show "f ` rel_interior S \ affine hull f ` S" by (meson \f ` rel_interior S \ f ` S\ hull_subset order_trans) show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "inj_on f (rel_interior S)" using inj_on_subset injf rel_interior_subset by blast qed auto qed lemma homeomorphic_rel_interiors_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(rel_interior S) homeomorphic (rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (rel_interior S) (rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast have "rel_interior T \ f ` rel_interior S" proof fix x assume "x \ rel_interior T" then have "g x \ rel_interior S" using gim by blast then show "x \ f ` rel_interior S" by (metis fg \x \ rel_interior T\ imageI) qed moreover have "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) ultimately show "f ` rel_interior S = rel_interior T" by blast show "continuous_on (rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast have "rel_interior S \ g ` rel_interior T" proof fix x assume "x \ rel_interior S" then have "f x \ rel_interior T" using fim by blast then show "x \ g ` rel_interior T" by (metis gf \x \ rel_interior S\ imageI) qed then show "g ` rel_interior T = rel_interior S" using gim by blast show "continuous_on (rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_interiors: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(rel_interior S) homeomorphic (rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_interiors_same_dimension [OF \S homeomorphic T\]) qed lemma homeomorphic_rel_boundaries_same_dimension: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" and aff: "aff_dim S = aff_dim T" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" using assms [unfolded homeomorphic_minimal] unfolding homeomorphic_def proof (clarify elim!: ex_forward) fix f g assume S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" then have fST: "f ` S = T" and gTS: "g ` T = S" and "inj_on f S" "inj_on g T" by (auto simp: inj_on_def intro: rev_image_eqI) metis+ have fim: "f ` rel_interior S \ rel_interior T" by (metis \inj_on f S\ aff contf continuous_image_subset_rel_interior fST order_refl) have gim: "g ` rel_interior T \ rel_interior S" by (metis \inj_on g T\ aff contg continuous_image_subset_rel_interior gTS order_refl) show "homeomorphism (S - rel_interior S) (T - rel_interior T) f g" unfolding homeomorphism_def proof (intro conjI ballI) show gf: "\x. x \ S - rel_interior S \ g (f x) = x" using S rel_interior_subset by blast show fg: "\y. y \ T - rel_interior T \ f (g y) = y" using T mem_rel_interior_ball by blast show "f ` (S - rel_interior S) = T - rel_interior T" using S fST fim gim by auto show "continuous_on (S - rel_interior S) f" using contf continuous_on_subset rel_interior_subset by blast show "g ` (T - rel_interior T) = S - rel_interior S" using T gTS gim fim by auto show "continuous_on (T - rel_interior T) g" using contg continuous_on_subset rel_interior_subset by blast qed qed lemma homeomorphic_rel_boundaries: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "S homeomorphic T" "rel_interior S = {} \ rel_interior T = {}" shows "(S - rel_interior S) homeomorphic (T - rel_interior T)" proof (cases "rel_interior T = {}") case True with assms show ?thesis by auto next case False obtain f g where S: "\x\S. f x \ T \ g (f x) = x" and T: "\y\T. g y \ S \ f (g y) = y" and contf: "continuous_on S f" and contg: "continuous_on T g" using assms [unfolded homeomorphic_minimal] by auto have "aff_dim (affine hull S) \ aff_dim (affine hull T)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior S" _ f]) apply (simp_all add: openin_rel_interior False assms) using contf continuous_on_subset rel_interior_subset apply blast apply (meson S hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis S inj_on_inverseI inj_on_subset rel_interior_subset) done moreover have "aff_dim (affine hull T) \ aff_dim (affine hull S)" apply (rule invariance_of_dimension_affine_sets [of _ "rel_interior T" _ g]) apply (simp_all add: openin_rel_interior False assms) using contg continuous_on_subset rel_interior_subset apply blast apply (meson T hull_subset image_subsetI rel_interior_subset rev_subsetD) apply (metis T inj_on_inverseI inj_on_subset rel_interior_subset) done ultimately have "aff_dim S = aff_dim T" by force then show ?thesis by (rule homeomorphic_rel_boundaries_same_dimension [OF \S homeomorphic T\]) qed proposition uniformly_continuous_homeomorphism_UNIV_trivial: fixes f :: "'a::euclidean_space \ 'a" assumes contf: "uniformly_continuous_on S f" and hom: "homeomorphism S UNIV f g" shows "S = UNIV" proof (cases "S = {}") case True then show ?thesis by (metis UNIV_I hom empty_iff homeomorphism_def image_eqI) next case False have "inj g" by (metis UNIV_I hom homeomorphism_apply2 injI) then have "open (g ` UNIV)" by (blast intro: invariance_of_domain hom homeomorphism_cont2) then have "open S" using hom homeomorphism_image2 by blast moreover have "complete S" unfolding complete_def proof clarify fix \ assume \: "\n. \ n \ S" and "Cauchy \" have "Cauchy (f o \)" using uniformly_continuous_imp_Cauchy_continuous \Cauchy \\ \ contf by blast then obtain l where "(f \ \) \ l" by (auto simp: convergent_eq_Cauchy [symmetric]) show "\l\S. \ \ l" proof show "g l \ S" using hom homeomorphism_image2 by blast have "(g \ (f \ \)) \ g l" by (meson UNIV_I \(f \ \) \ l\ continuous_on_sequentially hom homeomorphism_cont2) then show "\ \ g l" proof - have "\n. \ n = (g \ (f \ \)) n" by (metis (no_types) \ comp_eq_dest_lhs hom homeomorphism_apply1) then show ?thesis by (metis (no_types) LIMSEQ_iff \(g \ (f \ \)) \ g l\) qed qed qed then have "closed S" by (simp add: complete_eq_closed) ultimately show ?thesis using clopen [of S] False by simp qed +subsection\Formulation of loop homotopy in terms of maps out of type complex\ + +lemma homotopic_circlemaps_imp_homotopic_loops: + assumes "homotopic_with_canon (\h. True) (sphere 0 1) S f g" + shows "homotopic_loops S (f \ exp \ (\t. 2 * of_real pi * of_real t * \)) + (g \ exp \ (\t. 2 * of_real pi * of_real t * \))" +proof - + have "homotopic_with_canon (\f. True) {z. cmod z = 1} S f g" + using assms by (auto simp: sphere_def) + moreover have "continuous_on {0..1} (exp \ (\t. 2 * of_real pi * of_real t * \))" + by (intro continuous_intros) + moreover have "(exp \ (\t. 2 * of_real pi * of_real t * \)) ` {0..1} \ {z. cmod z = 1}" + by (auto simp: norm_mult) + ultimately + show ?thesis + apply (simp add: homotopic_loops_def comp_assoc) + apply (rule homotopic_with_compose_continuous_right) + apply (auto simp: pathstart_def pathfinish_def) + done +qed + +lemma homotopic_loops_imp_homotopic_circlemaps: + assumes "homotopic_loops S p q" + shows "homotopic_with_canon (\h. True) (sphere 0 1) S + (p \ (\z. (Arg2pi z / (2 * pi)))) + (q \ (\z. (Arg2pi z / (2 * pi))))" +proof - + obtain h where conth: "continuous_on ({0..1::real} \ {0..1}) h" + and him: "h ` ({0..1} \ {0..1}) \ S" + and h0: "(\x. h (0, x) = p x)" + and h1: "(\x. h (1, x) = q x)" + and h01: "(\t\{0..1}. h (t, 1) = h (t, 0)) " + using assms + by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def) + define j where "j \ \z. if 0 \ Im (snd z) + then h (fst z, Arg2pi (snd z) / (2 * pi)) + else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))" + have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \ Arg2pi y = 0 \ Arg2pi (cnj y) = 0" if "cmod y = 1" for y + using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj field_split_simps) + show ?thesis + proof (simp add: homotopic_with; intro conjI ballI exI) + show "continuous_on ({0..1} \ sphere 0 1) (\w. h (fst w, Arg2pi (snd w) / (2 * pi)))" + proof (rule continuous_on_eq) + show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \ {0..1} \ sphere 0 1" for x + using Arg2pi_eq that h01 by (force simp: j_def) + have eq: "S = S \ (UNIV \ {z. 0 \ Im z}) \ S \ (UNIV \ {z. Im z \ 0})" for S :: "(real*complex)set" + by auto + have c1: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. 0 \ Im z}) (\x. h (fst x, Arg2pi (snd x) / (2 * pi)))" + apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) + apply (auto simp: Arg2pi) + apply (meson Arg2pi_lt_2pi linear not_le) + done + have c2: "continuous_on ({0..1} \ sphere 0 1 \ UNIV \ {z. Im z \ 0}) (\x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))" + apply (intro continuous_intros continuous_on_compose2 [OF conth] continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi]) + apply (auto simp: Arg2pi) + apply (meson Arg2pi_lt_2pi linear not_le) + done + show "continuous_on ({0..1} \ sphere 0 1) j" + apply (simp add: j_def) + apply (subst eq) + apply (rule continuous_on_cases_local) + apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2) + using Arg2pi_eq h01 + by force + qed + have "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ h ` ({0..1} \ {0..1})" + by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le) + also have "... \ S" + using him by blast + finally show "(\w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \ sphere 0 1) \ S" . + qed (auto simp: h0 h1) +qed + +lemma simply_connected_homotopic_loops: + "simply_connected S \ + (\p q. homotopic_loops S p p \ homotopic_loops S q q \ homotopic_loops S p q)" +unfolding simply_connected_def using homotopic_loops_refl by metis + + +lemma simply_connected_eq_homotopic_circlemaps1: + fixes f :: "complex \ 'a::topological_space" and g :: "complex \ 'a" + assumes S: "simply_connected S" + and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \ S" + and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \ S" + shows "homotopic_with_canon (\h. True) (sphere 0 1) S f g" +proof - + have "homotopic_loops S (f \ exp \ (\t. of_real(2 * pi * t) * \)) (g \ exp \ (\t. of_real(2 * pi * t) * \))" + apply (rule S [unfolded simply_connected_homotopic_loops, rule_format]) + apply (simp add: homotopic_circlemaps_imp_homotopic_loops homotopic_with_refl contf fim contg gim) + done + then show ?thesis + apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps]) + apply (auto simp: o_def complex_norm_eq_1_exp mult.commute) + done +qed + +lemma simply_connected_eq_homotopic_circlemaps2a: + fixes h :: "complex \ 'a::topological_space" + assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \ S" + and hom: "\f g::complex \ 'a. + \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; + continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ + \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" + shows "\a. homotopic_with_canon (\h. True) (sphere 0 1) S h (\x. a)" + apply (rule_tac x="h 1" in exI) + apply (rule hom) + using assms + by (auto simp: continuous_on_const) + +lemma simply_connected_eq_homotopic_circlemaps2b: + fixes S :: "'a::real_normed_vector set" + assumes "\f g::complex \ 'a. + \continuous_on (sphere 0 1) f; f ` (sphere 0 1) \ S; + continuous_on (sphere 0 1) g; g ` (sphere 0 1) \ S\ + \ homotopic_with_canon (\h. True) (sphere 0 1) S f g" + shows "path_connected S" +proof (clarsimp simp add: path_connected_eq_homotopic_points) + fix a b + assume "a \ S" "b \ S" + then show "homotopic_loops S (linepath a a) (linepath b b)" + using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\x. a" "\x. b"]] + by (auto simp: o_def continuous_on_const linepath_def) +qed + +lemma simply_connected_eq_homotopic_circlemaps3: + fixes h :: "complex \ 'a::real_normed_vector" + assumes "path_connected S" + and hom: "\f::complex \ 'a. + \continuous_on (sphere 0 1) f; f `(sphere 0 1) \ S\ + \ \a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)" + shows "simply_connected S" +proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms) + fix p + assume p: "path p" "path_image p \ S" "pathfinish p = pathstart p" + then have "homotopic_loops S p p" + by (simp add: homotopic_loops_refl) + then obtain a where homp: "homotopic_with_canon (\h. True) (sphere 0 1) S (p \ (\z. Arg2pi z / (2 * pi))) (\x. a)" + by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom) + show "\a. a \ S \ homotopic_loops S p (linepath a a)" + proof (intro exI conjI) + show "a \ S" + using homotopic_with_imp_subset2 [OF homp] + by (metis dist_0_norm image_subset_iff mem_sphere norm_one) + have teq: "\t. \0 \ t; t \ 1\ + \ t = Arg2pi (exp (2 * of_real pi * of_real t * \)) / (2 * pi) \ t=1 \ Arg2pi (exp (2 * of_real pi * of_real t * \)) = 0" + apply (rule disjCI) + using Arg2pi_of_real [of 1] apply (auto simp: Arg2pi_exp) + done + have "homotopic_loops S p (p \ (\z. Arg2pi z / (2 * pi)) \ exp \ (\t. 2 * complex_of_real pi * complex_of_real t * \))" + apply (rule homotopic_loops_eq [OF p]) + using p teq apply (fastforce simp: pathfinish_def pathstart_def) + done + then + show "homotopic_loops S p (linepath a a)" + by (simp add: linepath_refl homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]]) + qed +qed + + +proposition simply_connected_eq_homotopic_circlemaps: + fixes S :: "'a::real_normed_vector set" + shows "simply_connected S \ + (\f g::complex \ 'a. + continuous_on (sphere 0 1) f \ f ` (sphere 0 1) \ S \ + continuous_on (sphere 0 1) g \ g ` (sphere 0 1) \ S + \ homotopic_with_canon (\h. True) (sphere 0 1) S f g)" + apply (rule iffI) + apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1) + by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3) + +proposition simply_connected_eq_contractible_circlemap: + fixes S :: "'a::real_normed_vector set" + shows "simply_connected S \ + path_connected S \ + (\f::complex \ 'a. + continuous_on (sphere 0 1) f \ f `(sphere 0 1) \ S + \ (\a. homotopic_with_canon (\h. True) (sphere 0 1) S f (\x. a)))" + apply (rule iffI) + apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b) + using simply_connected_eq_homotopic_circlemaps3 by blast + +corollary homotopy_eqv_simple_connectedness: + fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" + shows "S homotopy_eqv T \ simply_connected S \ simply_connected T" + by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality) + + +subsection\Homeomorphism of simple closed curves to circles\ + +proposition homeomorphic_simple_path_image_circle: + fixes a :: complex and \ :: "real \ 'a::t2_space" + assumes "simple_path \" and loop: "pathfinish \ = pathstart \" and "0 < r" + shows "(path_image \) homeomorphic sphere a r" +proof - + have "homotopic_loops (path_image \) \ \" + by (simp add: assms homotopic_loops_refl simple_path_imp_path) + then have hom: "homotopic_with_canon (\h. True) (sphere 0 1) (path_image \) + (\ \ (\z. Arg2pi z / (2*pi))) (\ \ (\z. Arg2pi z / (2*pi)))" + by (rule homotopic_loops_imp_homotopic_circlemaps) + have "\g. homeomorphism (sphere 0 1) (path_image \) (\ \ (\z. Arg2pi z / (2*pi))) g" + proof (rule homeomorphism_compact) + show "continuous_on (sphere 0 1) (\ \ (\z. Arg2pi z / (2*pi)))" + using hom homotopic_with_imp_continuous by blast + show "inj_on (\ \ (\z. Arg2pi z / (2*pi))) (sphere 0 1)" + proof + fix x y + assume xy: "x \ sphere 0 1" "y \ sphere 0 1" + and eq: "(\ \ (\z. Arg2pi z / (2*pi))) x = (\ \ (\z. Arg2pi z / (2*pi))) y" + then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))" + proof - + have "(Arg2pi x / (2*pi)) \ {0..1}" "(Arg2pi y / (2*pi)) \ {0..1}" + using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+ + with eq show ?thesis + using \simple_path \\ Arg2pi_lt_2pi unfolding simple_path_def o_def + by (metis eq_divide_eq_1 not_less_iff_gr_or_eq) + qed + with xy show "x = y" + by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere) + qed + have "\z. cmod z = 1 \ \x\{0..1}. \ (Arg2pi z / (2*pi)) = \ x" + by (metis Arg2pi_ge_0 Arg2pi_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral) + moreover have "\z\sphere 0 1. \ x = \ (Arg2pi z / (2*pi))" if "0 \ x" "x \ 1" for x + proof (cases "x=1") + case True + with Arg2pi_of_real [of 1] loop show ?thesis + by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \0 \ x\) + next + case False + then have *: "(Arg2pi (exp (\*(2* of_real pi* of_real x))) / (2*pi)) = x" + using that by (auto simp: Arg2pi_exp field_split_simps) + show ?thesis + by (rule_tac x="exp(\ * of_real(2*pi*x))" in bexI) (auto simp: *) + qed + ultimately show "(\ \ (\z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \" + by (auto simp: path_image_def image_iff) + qed auto + then have "path_image \ homeomorphic sphere (0::complex) 1" + using homeomorphic_def homeomorphic_sym by blast + also have "... homeomorphic sphere a r" + by (simp add: assms homeomorphic_spheres) + finally show ?thesis . +qed + +lemma homeomorphic_simple_path_images: + fixes \1 :: "real \ 'a::t2_space" and \2 :: "real \ 'b::t2_space" + assumes "simple_path \1" and loop: "pathfinish \1 = pathstart \1" + assumes "simple_path \2" and loop: "pathfinish \2 = pathstart \2" + shows "(path_image \1) homeomorphic (path_image \2)" + by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero) + subsection\Dimension-based conditions for various homeomorphisms\ lemma homeomorphic_subspaces_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "subspace S" "subspace T" shows "S homeomorphic T \ dim S = dim T" proof assume "S homeomorphic T" then obtain f g where hom: "homeomorphism S T f g" using homeomorphic_def by blast show "dim S = dim T" proof (rule order_antisym) show "dim S \ dim T" by (metis assms dual_order.refl inj_onI homeomorphism_cont1 [OF hom] homeomorphism_apply1 [OF hom] homeomorphism_image1 [OF hom] continuous_injective_image_subspace_dim_le) show "dim T \ dim S" by (metis assms dual_order.refl inj_onI homeomorphism_cont2 [OF hom] homeomorphism_apply2 [OF hom] homeomorphism_image2 [OF hom] continuous_injective_image_subspace_dim_le) qed next assume "dim S = dim T" then show "S homeomorphic T" by (simp add: assms homeomorphic_subspaces) qed lemma homeomorphic_affine_sets_eq: fixes S :: "'a::euclidean_space set" and T :: "'b::euclidean_space set" assumes "affine S" "affine T" shows "S homeomorphic T \ aff_dim S = aff_dim T" proof (cases "S = {} \ T = {}") case True then show ?thesis using assms homeomorphic_affine_sets by force next case False then obtain a b where "a \ S" "b \ T" by blast then have "subspace ((+) (- a) ` S)" "subspace ((+) (- b) ` T)" using affine_diffs_subspace assms by blast+ then show ?thesis by (metis affine_imp_convex assms homeomorphic_affine_sets homeomorphic_convex_sets) qed lemma homeomorphic_hyperplanes_eq: fixes a :: "'a::euclidean_space" and c :: "'b::euclidean_space" assumes "a \ 0" "c \ 0" shows "({x. a \ x = b} homeomorphic {x. c \ x = d} \ DIM('a) = DIM('b))" apply (auto simp: homeomorphic_affine_sets_eq affine_hyperplane assms) by (metis DIM_positive Suc_pred) lemma homeomorphic_UNIV_UNIV: shows "(UNIV::'a set) homeomorphic (UNIV::'b set) \ DIM('a::euclidean_space) = DIM('b::euclidean_space)" by (simp add: homeomorphic_subspaces_eq) lemma simply_connected_sphere_gen: assumes "convex S" "bounded S" and 3: "3 \ aff_dim S" shows "simply_connected(rel_frontier S)" proof - have pa: "path_connected (rel_frontier S)" using assms by (simp add: path_connected_sphere_gen) show ?thesis proof (clarsimp simp add: simply_connected_eq_contractible_circlemap pa) fix f assume f: "continuous_on (sphere (0::complex) 1) f" "f ` sphere 0 1 \ rel_frontier S" have eq: "sphere (0::complex) 1 = rel_frontier(cball 0 1)" by simp have "convex (cball (0::complex) 1)" by (rule convex_cball) then obtain c where "homotopic_with_canon (\z. True) (sphere (0::complex) 1) (rel_frontier S) f (\x. c)" apply (rule inessential_spheremap_lowdim_gen [OF _ bounded_cball \convex S\ \bounded S\, where f=f]) using f 3 apply (auto simp: aff_dim_cball) done then show "\a. homotopic_with_canon (\h. True) (sphere 0 1) (rel_frontier S) f (\x. a)" by blast qed qed subsection\more invariance of domain\(*FIX ME title? *) proposition invariance_of_domain_sphere_affine_set_gen: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and U: "bounded U" "convex U" and "affine T" and affTU: "aff_dim T < aff_dim U" and ope: "openin (top_of_set (rel_frontier U)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "rel_frontier U = {}") case True then show ?thesis using ope openin_subset by force next case False obtain b c where b: "b \ rel_frontier U" and c: "c \ rel_frontier U" and "b \ c" using \bounded U\ rel_frontier_not_sing [of U] subset_singletonD False by fastforce obtain V :: "'a set" where "affine V" and affV: "aff_dim V = aff_dim U - 1" proof (rule choose_affine_subset [OF affine_UNIV]) show "- 1 \ aff_dim U - 1" by (metis aff_dim_empty aff_dim_geq aff_dim_negative_iff affTU diff_0 diff_right_mono not_le) show "aff_dim U - 1 \ aff_dim (UNIV::'a set)" by (metis aff_dim_UNIV aff_dim_le_DIM le_cases not_le zle_diff1_eq) qed auto have SU: "S \ rel_frontier U" using ope openin_imp_subset by auto have homb: "rel_frontier U - {b} homeomorphic V" and homc: "rel_frontier U - {c} homeomorphic V" using homeomorphic_punctured_sphere_affine_gen [of U _ V] by (simp_all add: \affine V\ affV U b c) then obtain g h j k where gh: "homeomorphism (rel_frontier U - {b}) V g h" and jk: "homeomorphism (rel_frontier U - {c}) V j k" by (auto simp: homeomorphic_def) with SU have hgsub: "(h ` g ` (S - {b})) \ S" and kjsub: "(k ` j ` (S - {c})) \ S" by (simp_all add: homeomorphism_def subset_eq) have [simp]: "aff_dim T \ aff_dim V" by (simp add: affTU affV) have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (g ` (S - {b}))" apply (rule homeomorphism_imp_open_map [OF gh]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (g ` (S - {b})) (f \ h)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets gh set_eq_subset) using contf continuous_on_subset hgsub by blast show "inj_on (f \ h) (g ` (S - {b}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU b homeomorphism_def inj_onD injf insert_Diff insert_iff gh rev_subsetD) show "(f \ h) ` g ` (S - {b}) \ T" by (metis fim image_comp image_mono hgsub subset_trans) qed (auto simp: assms) moreover have "openin (top_of_set T) ((f \ k) ` j ` (S - {c}))" proof (rule invariance_of_domain_affine_sets [OF _ \affine V\]) show "openin (top_of_set V) (j ` (S - {c}))" apply (rule homeomorphism_imp_open_map [OF jk]) by (meson Diff_mono Diff_subset SU ope openin_delete openin_subset_trans order_refl) show "continuous_on (j ` (S - {c})) (f \ k)" apply (rule continuous_on_compose) apply (meson Diff_mono SU homeomorphism_def homeomorphism_of_subsets jk set_eq_subset) using contf continuous_on_subset kjsub by blast show "inj_on (f \ k) (j ` (S - {c}))" using kjsub apply (clarsimp simp add: inj_on_def) by (metis SU c homeomorphism_def inj_onD injf insert_Diff insert_iff jk rev_subsetD) show "(f \ k) ` j ` (S - {c}) \ T" by (metis fim image_comp image_mono kjsub subset_trans) qed (auto simp: assms) ultimately have "openin (top_of_set T) ((f \ h) ` g ` (S - {b}) \ ((f \ k) ` j ` (S - {c})))" by (rule openin_Un) moreover have "(f \ h) ` g ` (S - {b}) = f ` (S - {b})" proof - have "h ` g ` (S - {b}) = (S - {b})" proof show "h ` g ` (S - {b}) \ S - {b}" using homeomorphism_apply1 [OF gh] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {b} \ h ` g ` (S - {b})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF gh] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "(f \ k) ` j ` (S - {c}) = f ` (S - {c})" proof - have "k ` j ` (S - {c}) = (S - {c})" proof show "k ` j ` (S - {c}) \ S - {c}" using homeomorphism_apply1 [OF jk] SU by (fastforce simp add: image_iff image_subset_iff) show "S - {c} \ k ` j ` (S - {c})" apply clarify by (metis SU subsetD homeomorphism_apply1 [OF jk] image_iff member_remove remove_def) qed then show ?thesis by (metis image_comp) qed moreover have "f ` (S - {b}) \ f ` (S - {c}) = f ` (S)" using \b \ c\ by blast ultimately show ?thesis by simp qed lemma invariance_of_domain_sphere_affine_set: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on S f" and injf: "inj_on f S" and fim: "f ` S \ T" and "r \ 0" "affine T" and affTU: "aff_dim T < DIM('a)" and ope: "openin (top_of_set (sphere a r)) S" shows "openin (top_of_set T) (f ` S)" proof (cases "sphere a r = {}") case True then show ?thesis using ope openin_subset by force next case False show ?thesis proof (rule invariance_of_domain_sphere_affine_set_gen [OF contf injf fim bounded_cball convex_cball \affine T\]) show "aff_dim T < aff_dim (cball a r)" by (metis False affTU aff_dim_cball assms(4) linorder_cases sphere_empty) show "openin (top_of_set (rel_frontier (cball a r))) S" by (simp add: \r \ 0\ ope) qed qed lemma no_embedding_sphere_lowdim: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes contf: "continuous_on (sphere a r) f" and injf: "inj_on f (sphere a r)" and "r > 0" shows "DIM('a) \ DIM('b)" proof - have "False" if "DIM('a) > DIM('b)" proof - have "compact (f ` sphere a r)" using compact_continuous_image by (simp add: compact_continuous_image contf) then have "\ open (f ` sphere a r)" using compact_open by (metis assms(3) image_is_empty not_less_iff_gr_or_eq sphere_eq_empty) then show False using invariance_of_domain_sphere_affine_set [OF contf injf subset_UNIV] \r > 0\ by (metis aff_dim_UNIV affine_UNIV less_irrefl of_nat_less_iff open_openin openin_subtopology_self subtopology_UNIV that) qed then show ?thesis using not_less by blast qed lemma simply_connected_sphere: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(sphere a r)" proof (cases rule: linorder_cases [of r 0]) case less then show ?thesis by simp next case equal then show ?thesis by (auto simp: convex_imp_simply_connected) next case greater then show ?thesis using simply_connected_sphere_gen [of "cball a r"] assms by (simp add: aff_dim_cball) qed lemma simply_connected_sphere_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(sphere a r) \ 3 \ DIM('a) \ r \ 0" (is "?lhs = ?rhs") proof (cases "r \ 0") case True have "simply_connected (sphere a r)" apply (rule convex_imp_simply_connected) using True less_eq_real_def by auto with True show ?thesis by auto next case False show ?thesis proof assume L: ?lhs have "False" if "DIM('a) = 1 \ DIM('a) = 2" using that proof assume "DIM('a) = 1" with L show False using connected_sphere_eq simply_connected_imp_connected by (metis False Suc_1 not_less_eq_eq order_refl) next assume "DIM('a) = 2" then have "sphere a r homeomorphic sphere (0::complex) 1" by (metis DIM_complex False homeomorphic_spheres_gen not_less zero_less_one) then have "simply_connected(sphere (0::complex) 1)" using L homeomorphic_simply_connected_eq by blast then obtain a::complex where "homotopic_with_canon (\h. True) (sphere 0 1) (sphere 0 1) id (\x. a)" apply (simp add: simply_connected_eq_contractible_circlemap) by (metis continuous_on_id' id_apply image_id subset_refl) then show False using contractible_sphere contractible_def not_one_le_zero by blast qed with False show ?rhs apply simp by (metis DIM_ge_Suc0 le_antisym not_less_eq_eq numeral_2_eq_2 numeral_3_eq_3) next assume ?rhs with False show ?lhs by (simp add: simply_connected_sphere) qed qed lemma simply_connected_punctured_universe_eq: fixes a :: "'a::euclidean_space" shows "simply_connected(- {a}) \ 3 \ DIM('a)" proof - have [simp]: "a \ rel_interior (cball a 1)" by (simp add: rel_interior_nonempty_interior) have [simp]: "affine hull cball a 1 - {a} = -{a}" by (metis Compl_eq_Diff_UNIV aff_dim_cball aff_dim_lt_full not_less_iff_gr_or_eq zero_less_one) have "simply_connected(- {a}) \ simply_connected(sphere a 1)" apply (rule sym) apply (rule homotopy_eqv_simple_connectedness) using homotopy_eqv_rel_frontier_punctured_affine_hull [of "cball a 1" a] apply auto done also have "... \ 3 \ DIM('a)" by (simp add: simply_connected_sphere_eq) finally show ?thesis . qed lemma not_simply_connected_circle: fixes a :: complex shows "0 < r \ \ simply_connected(sphere a r)" by (simp add: simply_connected_sphere_eq) proposition simply_connected_punctured_convex: fixes a :: "'a::euclidean_space" assumes "convex S" and 3: "3 \ aff_dim S" shows "simply_connected(S - {a})" proof (cases "a \ rel_interior S") case True then obtain e where "a \ S" "0 < e" and e: "cball a e \ affine hull S \ S" by (auto simp: rel_interior_cball) have con: "convex (cball a e \ affine hull S)" by (simp add: convex_Int) have bo: "bounded (cball a e \ affine hull S)" by (simp add: bounded_Int) have "affine hull S \ interior (cball a e) \ {}" using \0 < e\ \a \ S\ hull_subset by fastforce then have "3 \ aff_dim (affine hull S \ cball a e)" by (simp add: 3 aff_dim_convex_Int_nonempty_interior [OF convex_affine_hull]) also have "... = aff_dim (cball a e \ affine hull S)" by (simp add: Int_commute) finally have "3 \ aff_dim (cball a e \ affine hull S)" . moreover have "rel_frontier (cball a e \ affine hull S) homotopy_eqv S - {a}" proof (rule homotopy_eqv_rel_frontier_punctured_convex) show "a \ rel_interior (cball a e \ affine hull S)" by (meson IntI Int_mono \a \ S\ \0 < e\ e \cball a e \ affine hull S \ S\ ball_subset_cball centre_in_cball dual_order.strict_implies_order hull_inc hull_mono mem_rel_interior_ball) have "closed (cball a e \ affine hull S)" by blast then show "rel_frontier (cball a e \ affine hull S) \ S" apply (simp add: rel_frontier_def) using e by blast show "S \ affine hull (cball a e \ affine hull S)" by (metis (no_types, lifting) IntI \a \ S\ \0 < e\ affine_hull_convex_Int_nonempty_interior centre_in_ball convex_affine_hull empty_iff hull_subset inf_commute interior_cball subsetCE subsetI) qed (auto simp: assms con bo) ultimately show ?thesis using homotopy_eqv_simple_connectedness simply_connected_sphere_gen [OF con bo] by blast next case False show ?thesis apply (rule contractible_imp_simply_connected) apply (rule contractible_convex_tweak_boundary_points [OF \convex S\]) apply (simp add: False rel_interior_subset subset_Diff_insert) by (meson Diff_subset closure_subset subset_trans) qed corollary simply_connected_punctured_universe: fixes a :: "'a::euclidean_space" assumes "3 \ DIM('a)" shows "simply_connected(- {a})" proof - have [simp]: "affine hull cball a 1 = UNIV" apply auto by (metis UNIV_I aff_dim_cball aff_dim_lt_full zero_less_one not_less_iff_gr_or_eq) have "simply_connected (rel_frontier (cball a 1)) = simply_connected (affine hull cball a 1 - {a})" apply (rule homotopy_eqv_simple_connectedness) apply (rule homotopy_eqv_rel_frontier_punctured_affine_hull) apply (force simp: rel_interior_cball intro: homotopy_eqv_simple_connectedness homotopy_eqv_rel_frontier_punctured_affine_hull)+ done then show ?thesis using simply_connected_sphere [of a 1, OF assms] by (auto simp: Compl_eq_Diff_UNIV) qed subsection\The power, squaring and exponential functions as covering maps\ proposition covering_space_power_punctured_plane: assumes "0 < n" shows "covering_space (- {0}) (\z::complex. z^n) (- {0})" proof - consider "n = 1" | "2 \ n" using assms by linarith then obtain e where "0 < e" and e: "\w z. cmod(w - z) < e * cmod z \ (w^n = z^n \ w = z)" proof cases assume "n = 1" then show ?thesis by (rule_tac e=1 in that) auto next assume "2 \ n" have eq_if_pow_eq: "w = z" if lt: "cmod (w - z) < 2 * sin (pi / real n) * cmod z" and eq: "w^n = z^n" for w z proof (cases "z = 0") case True with eq assms show ?thesis by (auto simp: power_0_left) next case False then have "z \ 0" by auto have "(w/z)^n = 1" by (metis False divide_self_if eq power_divide power_one) then obtain j where j: "w / z = exp (2 * of_real pi * \ * j / n)" and "j < n" using Suc_leI assms \2 \ n\ complex_roots_unity [THEN eqset_imp_iff, of n "w/z"] by force have "cmod (w/z - 1) < 2 * sin (pi / real n)" using lt assms \z \ 0\ by (simp add: field_split_simps norm_divide) then have "cmod (exp (\ * of_real (2 * pi * j / n)) - 1) < 2 * sin (pi / real n)" by (simp add: j field_simps) then have "2 * \sin((2 * pi * j / n) / 2)\ < 2 * sin (pi / real n)" by (simp only: dist_exp_i_1) then have sin_less: "sin((pi * j / n)) < sin (pi / real n)" by (simp add: field_simps) then have "w / z = 1" proof (cases "j = 0") case True then show ?thesis by (auto simp: j) next case False then have "sin (pi / real n) \ sin((pi * j / n))" proof (cases "j / n \ 1/2") case True show ?thesis apply (rule sin_monotone_2pi_le) using \j \ 0 \ \j < n\ True apply (auto simp: field_simps intro: order_trans [of _ 0]) done next case False then have seq: "sin(pi * j / n) = sin(pi * (n - j) / n)" using \j < n\ by (simp add: algebra_simps diff_divide_distrib of_nat_diff) show ?thesis apply (simp only: seq) apply (rule sin_monotone_2pi_le) using \j < n\ False apply (auto simp: field_simps intro: order_trans [of _ 0]) done qed with sin_less show ?thesis by force qed then show ?thesis by simp qed show ?thesis apply (rule_tac e = "2 * sin(pi / n)" in that) apply (force simp: \2 \ n\ sin_pi_divide_n_gt_0) apply (meson eq_if_pow_eq) done qed have zn1: "continuous_on (- {0}) (\z::complex. z^n)" by (rule continuous_intros)+ have zn2: "(\z::complex. z^n) ` (- {0}) = - {0}" using assms by (auto simp: image_def elim: exists_complex_root_nonzero [where n = n]) have zn3: "\T. z^n \ T \ open T \ 0 \ T \ (\v. \v = -{0} \ (\z. z ^ n) -` T \ (\u\v. open u \ 0 \ u) \ pairwise disjnt v \ (\u\v. Ex (homeomorphism u T (\z. z^n))))" if "z \ 0" for z::complex proof - define d where "d \ min (1/2) (e/4) * norm z" have "0 < d" by (simp add: d_def \0 < e\ \z \ 0\) have iff_x_eq_y: "x^n = y^n \ x = y" if eq: "w^n = z^n" and x: "x \ ball w d" and y: "y \ ball w d" for w x y proof - have [simp]: "norm z = norm w" using that by (simp add: assms power_eq_imp_eq_norm) show ?thesis proof (cases "w = 0") case True with \z \ 0\ assms eq show ?thesis by (auto simp: power_0_left) next case False have "cmod (x - y) < 2*d" using x y by (simp add: dist_norm [symmetric]) (metis dist_commute mult_2 dist_triangle_less_add) also have "... \ 2 * e / 4 * norm w" using \e > 0\ by (simp add: d_def min_mult_distrib_right) also have "... = e * (cmod w / 2)" by simp also have "... \ e * cmod y" apply (rule mult_left_mono) using \e > 0\ y apply (simp_all add: dist_norm d_def min_mult_distrib_right del: divide_const_simps) apply (metis dist_0_norm dist_complex_def dist_triangle_half_l linorder_not_less order_less_irrefl) done finally have "cmod (x - y) < e * cmod y" . then show ?thesis by (rule e) qed qed then have inj: "inj_on (\w. w^n) (ball z d)" by (simp add: inj_on_def) have cont: "continuous_on (ball z d) (\w. w ^ n)" by (intro continuous_intros) have noncon: "\ (\w::complex. w^n) constant_on UNIV" by (metis UNIV_I assms constant_on_def power_one zero_neq_one zero_power) have im_eq: "(\w. w^n) ` ball z' d = (\w. w^n) ` ball z d" if z': "z'^n = z^n" for z' proof - have nz': "norm z' = norm z" using that assms power_eq_imp_eq_norm by blast have "(w \ (\w. w^n) ` ball z' d) = (w \ (\w. w^n) ` ball z d)" for w proof (cases "w=0") case True with assms show ?thesis by (simp add: image_def ball_def nz') next case False have "z' \ 0" using \z \ 0\ nz' by force have [simp]: "(z*x / z')^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have [simp]: "cmod (z - z * x / z') = cmod (z' - x)" if "x \ 0" for x proof - have "cmod (z - z * x / z') = cmod z * cmod (1 - x / z')" by (metis (no_types) ab_semigroup_mult_class.mult_ac(1) divide_complex_def mult.right_neutral norm_mult right_diff_distrib') also have "... = cmod z' * cmod (1 - x / z')" by (simp add: nz') also have "... = cmod (z' - x)" by (simp add: \z' \ 0\ diff_divide_eq_iff norm_divide) finally show ?thesis . qed have [simp]: "(z'*x / z)^n = x^n" if "x \ 0" for x using z' that by (simp add: field_simps \z \ 0\) have [simp]: "cmod (z' - z' * x / z) = cmod (z - x)" if "x \ 0" for x proof - have "cmod (z * (1 - x * inverse z)) = cmod (z - x)" by (metis \z \ 0\ diff_divide_distrib divide_complex_def divide_self_if nonzero_eq_divide_eq semiring_normalization_rules(7)) then show ?thesis by (metis (no_types) mult.assoc divide_complex_def mult.right_neutral norm_mult nz' right_diff_distrib') qed show ?thesis unfolding image_def ball_def apply safe apply simp_all apply (rule_tac x="z/z' * x" in exI) using assms False apply (simp add: dist_norm) apply (rule_tac x="z'/z * x" in exI) using assms False apply (simp add: dist_norm) done qed then show ?thesis by blast qed have ex_ball: "\B. (\z'. B = ball z' d \ z'^n = z^n) \ x \ B" if "x \ 0" and eq: "x^n = w^n" and dzw: "dist z w < d" for x w proof - have "w \ 0" by (metis assms power_eq_0_iff that(1) that(2)) have [simp]: "cmod x = cmod w" using assms power_eq_imp_eq_norm eq by blast have [simp]: "cmod (x * z / w - x) = cmod (z - w)" proof - have "cmod (x * z / w - x) = cmod x * cmod (z / w - 1)" by (metis (no_types) mult.right_neutral norm_mult right_diff_distrib' times_divide_eq_right) also have "... = cmod w * cmod (z / w - 1)" by simp also have "... = cmod (z - w)" by (simp add: \w \ 0\ divide_diff_eq_iff nonzero_norm_divide) finally show ?thesis . qed show ?thesis apply (rule_tac x="ball (z / w * x) d" in exI) using \d > 0\ that apply (simp add: ball_eq_ball_iff) apply (simp add: \z \ 0\ \w \ 0\ field_simps) apply (simp add: dist_norm) done qed show ?thesis proof (rule exI, intro conjI) show "z ^ n \ (\w. w ^ n) ` ball z d" using \d > 0\ by simp show "open ((\w. w ^ n) ` ball z d)" by (rule invariance_of_domain [OF cont open_ball inj]) show "0 \ (\w. w ^ n) ` ball z d" using \z \ 0\ assms by (force simp: d_def) show "\v. \v = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d \ (\u\v. open u \ 0 \ u) \ disjoint v \ (\u\v. Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n)))" proof (rule exI, intro ballI conjI) show "\{ball z' d |z'. z'^n = z^n} = - {0} \ (\z. z ^ n) -` (\w. w ^ n) ` ball z d" (is "?l = ?r") proof show "?l \ ?r" apply auto apply (simp add: assms d_def power_eq_imp_eq_norm that) by (metis im_eq image_eqI mem_ball) show "?r \ ?l" by auto (meson ex_ball) qed show "\u. u \ {ball z' d |z'. z' ^ n = z ^ n} \ 0 \ u" by (force simp add: assms d_def power_eq_imp_eq_norm that) show "disjoint {ball z' d |z'. z' ^ n = z ^ n}" proof (clarsimp simp add: pairwise_def disjnt_iff) fix \ \ x assume "\^n = z^n" "\^n = z^n" "ball \ d \ ball \ d" and "dist \ x < d" "dist \ x < d" then have "dist \ \ < d+d" using dist_triangle_less_add by blast then have "cmod (\ - \) < 2*d" by (simp add: dist_norm) also have "... \ e * cmod z" using mult_right_mono \0 < e\ that by (auto simp: d_def) finally have "cmod (\ - \) < e * cmod z" . with e have "\ = \" by (metis \\^n = z^n\ \\^n = z^n\ assms power_eq_imp_eq_norm) then show "False" using \ball \ d \ ball \ d\ by blast qed show "Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" if "u \ {ball z' d |z'. z' ^ n = z ^ n}" for u proof (rule invariance_of_domain_homeomorphism [of "u" "\z. z^n"]) show "open u" using that by auto show "continuous_on u (\z. z ^ n)" by (intro continuous_intros) show "inj_on (\z. z ^ n) u" using that by (auto simp: iff_x_eq_y inj_on_def) show "\g. homeomorphism u ((\z. z ^ n) ` u) (\z. z ^ n) g \ Ex (homeomorphism u ((\w. w ^ n) ` ball z d) (\z. z ^ n))" using im_eq that by clarify metis qed auto qed auto qed qed show ?thesis using assms apply (simp add: covering_space_def zn1 zn2) apply (subst zn2 [symmetric]) apply (simp add: openin_open_eq open_Compl) apply (blast intro: zn3) done qed corollary covering_space_square_punctured_plane: "covering_space (- {0}) (\z::complex. z^2) (- {0})" by (simp add: covering_space_power_punctured_plane) proposition covering_space_exp_punctured_plane: "covering_space UNIV (\z::complex. exp z) (- {0})" proof (simp add: covering_space_def, intro conjI ballI) show "continuous_on UNIV (\z::complex. exp z)" by (rule continuous_on_exp [OF continuous_on_id]) show "range exp = - {0::complex}" by auto (metis exp_Ln range_eqI) show "\T. z \ T \ openin (top_of_set (- {0})) T \ (\v. \v = exp -` T \ (\u\v. open u) \ disjoint v \ (\u\v. \q. homeomorphism u T exp q))" if "z \ - {0::complex}" for z proof - have "z \ 0" using that by auto have inj_exp: "inj_on exp (ball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: ball_subset_ball_iff) define \ where "\ \ range (\n. (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1))" show ?thesis proof (intro exI conjI) show "z \ exp ` (ball(Ln z) 1)" by (metis \z \ 0\ centre_in_ball exp_Ln rev_image_eqI zero_less_one) have "open (- {0::complex})" by blast moreover have "inj_on exp (ball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: ball_subset_ball_iff) ultimately show "openin (top_of_set (- {0})) (exp ` ball (Ln z) 1)" by (auto simp: openin_open_eq invariance_of_domain continuous_on_exp [OF continuous_on_id]) show "\\ = exp -` exp ` ball (Ln z) 1" by (force simp: \_def Complex_Transcendental.exp_eq image_iff) show "\V\\. open V" by (auto simp: \_def inj_on_def continuous_intros invariance_of_domain) have xy: "2 \ cmod (2 * of_int x * of_real pi * \ - 2 * of_int y * of_real pi * \)" if "x < y" for x y proof - have "1 \ abs (x - y)" using that by linarith then have "1 \ cmod (of_int x - of_int y) * 1" by (metis mult.right_neutral norm_of_int of_int_1_le_iff of_int_abs of_int_diff) also have "... \ cmod (of_int x - of_int y) * of_real pi" apply (rule mult_left_mono) using pi_ge_two by auto also have "... \ cmod ((of_int x - of_int y) * of_real pi * \)" by (simp add: norm_mult) also have "... \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" by (simp add: algebra_simps) finally have "1 \ cmod (of_int x * of_real pi * \ - of_int y * of_real pi * \)" . then have "2 * 1 \ cmod (2 * (of_int x * of_real pi * \ - of_int y * of_real pi * \))" by (metis mult_le_cancel_left_pos norm_mult_numeral1 zero_less_numeral) then show ?thesis by (simp add: algebra_simps) qed show "disjoint \" apply (clarsimp simp add: \_def pairwise_def disjnt_def add.commute [of _ "x*y" for x y] image_add_ball ball_eq_ball_iff) apply (rule disjoint_ballI) apply (auto simp: dist_norm neq_iff) by (metis norm_minus_commute xy)+ show "\u\\. \q. homeomorphism u (exp ` ball (Ln z) 1) exp q" proof fix u assume "u \ \" then obtain n where n: "u = (\x. x + of_real (2 * of_int n * pi) * \) ` (ball(Ln z) 1)" by (auto simp: \_def) have "compact (cball (Ln z) 1)" by simp moreover have "continuous_on (cball (Ln z) 1) exp" by (rule continuous_on_exp [OF continuous_on_id]) moreover have "inj_on exp (cball (Ln z) 1)" apply (rule inj_on_subset [OF inj_on_exp_pi [of "Ln z"]]) using pi_ge_two by (simp add: cball_subset_ball_iff) ultimately obtain \ where hom: "homeomorphism (cball (Ln z) 1) (exp ` cball (Ln z) 1) exp \" using homeomorphism_compact by blast have eq1: "exp ` u = exp ` ball (Ln z) 1" unfolding n apply (auto simp: algebra_simps) apply (rename_tac w) apply (rule_tac x = "w + \ * (of_int n * (of_real pi * 2))" in image_eqI) apply (auto simp: image_iff) done have \exp: "\ (exp x) + 2 * of_int n * of_real pi * \ = x" if "x \ u" for x proof - have "exp x = exp (x - 2 * of_int n * of_real pi * \)" by (simp add: exp_eq) then have "\ (exp x) = \ (exp (x - 2 * of_int n * of_real pi * \))" by simp also have "... = x - 2 * of_int n * of_real pi * \" apply (rule homeomorphism_apply1 [OF hom]) using \x \ u\ by (auto simp: n) finally show ?thesis by simp qed have exp2n: "exp (\ (exp x) + 2 * of_int n * complex_of_real pi * \) = exp x" if "dist (Ln z) x < 1" for x using that by (auto simp: exp_eq homeomorphism_apply1 [OF hom]) have cont: "continuous_on (exp ` ball (Ln z) 1) (\x. \ x + 2 * of_int n * complex_of_real pi * \)" apply (intro continuous_intros) apply (rule continuous_on_subset [OF homeomorphism_cont2 [OF hom]]) apply (force simp:) done show "\q. homeomorphism u (exp ` ball (Ln z) 1) exp q" apply (rule_tac x="(\x. x + of_real(2 * n * pi) * \) \ \" in exI) unfolding homeomorphism_def apply (intro conjI ballI eq1 continuous_on_exp [OF continuous_on_id]) apply (auto simp: \exp exp2n cont n) apply (simp add: homeomorphism_apply1 [OF hom]) using hom homeomorphism_apply1 apply (force simp: image_iff) done qed qed qed qed subsection\Hence the Borsukian results about mappings into circles\(*FIX ME title *) lemma inessential_eq_continuous_logarithm: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x)))" (is "?lhs \ ?rhs") proof assume ?lhs thus ?rhs by (metis covering_space_lift_inessential_function covering_space_exp_punctured_plane) next assume ?rhs then obtain g where contg: "continuous_on S g" and f: "\x. x \ S \ f x = exp(g x)" by metis obtain a where "homotopic_with_canon (\h. True) S (- {of_real 0}) (exp \ g) (\x. a)" proof (rule nullhomotopic_through_contractible [OF contg subset_UNIV _ _ contractible_UNIV]) show "continuous_on (UNIV::complex set) exp" by (intro continuous_intros) show "range exp \ - {0}" by auto qed force then have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using f homotopic_with_eq by fastforce then show ?lhs .. qed corollary inessential_imp_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" assumes "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" obtains g where "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" proof - have "homotopic_with_canon (\h. True) S (- {0}) f (\t. a)" using assms homotopic_with_subset_right by fastforce then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma inessential_eq_continuous_logarithm_circle: fixes f :: "'a::real_normed_vector \ complex" shows "(\a. homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)) \ (\g. continuous_on S g \ (\x \ S. f x = exp(\ * of_real(g x))))" (is "?lhs \ ?rhs") proof assume L: ?lhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(g x)" using inessential_imp_continuous_logarithm_circle by blast have "f ` S \ sphere 0 1" by (metis L homotopic_with_imp_subset1) then have "\x. x \ S \ Re (g x) = 0" using g by auto then show ?rhs apply (rule_tac x="Im \ g" in exI) apply (intro conjI contg continuous_intros) apply (auto simp: Euler g) done next assume ?rhs then obtain g where contg: "continuous_on S g" and g: "\x. x \ S \ f x = exp(\* of_real(g x))" by metis obtain a where "homotopic_with_canon (\h. True) S (sphere 0 1) ((exp \ (\z. \*z)) \ (of_real \ g)) (\x. a)" proof (rule nullhomotopic_through_contractible) show "continuous_on S (complex_of_real \ g)" by (intro conjI contg continuous_intros) show "(complex_of_real \ g) ` S \ \" by auto show "continuous_on \ (exp \ (*)\)" by (intro continuous_intros) show "(exp \ (*)\) ` \ \ sphere 0 1" by (auto simp: complex_is_Real_iff) qed (auto simp: convex_Reals convex_imp_contractible) moreover have "\x. x \ S \ (exp \ (*)\ \ (complex_of_real \ g)) x = f x" by (simp add: g) ultimately have "homotopic_with_canon (\h. True) S (sphere 0 1) f (\t. a)" using homotopic_with_eq by force then show ?lhs .. qed proposition homotopic_with_sphere_times: fixes f :: "'a::real_normed_vector \ complex" assumes hom: "homotopic_with_canon (\x. True) S (sphere 0 1) f g" and conth: "continuous_on S h" and hin: "\x. x \ S \ h x \ sphere 0 1" shows "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x * h x) (\x. g x * h x)" proof - obtain k where contk: "continuous_on ({0..1::real} \ S) k" and kim: "k ` ({0..1} \ S) \ sphere 0 1" and k0: "\x. k(0, x) = f x" and k1: "\x. k(1, x) = g x" using hom by (auto simp: homotopic_with_def) show ?thesis apply (simp add: homotopic_with) apply (rule_tac x="\z. k z*(h \ snd)z" in exI) apply (intro conjI contk continuous_intros) apply (simp add: conth) using kim hin apply (force simp: norm_mult k0 k1)+ done qed proposition homotopic_circlemaps_divide: fixes f :: "'a::real_normed_vector \ complex" shows "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ (\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c))" proof - have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" if "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)" for c proof - have "S = {} \ path_component (sphere 0 1) 1 c" using homotopic_with_imp_subset2 [OF that] path_connected_sphere [of "0::complex" 1] by (auto simp: path_connected_component) then have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. 1) (\x. c)" by (simp add: homotopic_constant_maps) then show ?thesis using homotopic_with_symD homotopic_with_trans that by blast qed then have *: "(\c. homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. c)) \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" by auto have "homotopic_with_canon (\x. True) S (sphere 0 1) f g \ continuous_on S f \ f ` S \ sphere 0 1 \ continuous_on S g \ g ` S \ sphere 0 1 \ homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" (is "?lhs \ ?rhs") proof assume L: ?lhs have geq1 [simp]: "\x. x \ S \ cmod (g x) = 1" using homotopic_with_imp_subset2 [OF L] by (simp add: image_subset_iff) have cont: "continuous_on S (inverse \ g)" apply (rule continuous_intros) using homotopic_with_imp_continuous [OF L] apply blast apply (rule continuous_on_subset [of "sphere 0 1", OF continuous_on_inverse]) apply (auto simp: continuous_on_id) done have "homotopic_with_canon (\x. True) S (sphere 0 1) (\x. f x / g x) (\x. 1)" using homotopic_with_sphere_times [OF L cont] apply (rule homotopic_with_eq) apply (auto simp: division_ring_class.divide_inverse norm_inverse) by (metis geq1 norm_zero right_inverse zero_neq_one) with L show ?rhs by (auto simp: homotopic_with_imp_continuous dest: homotopic_with_imp_subset1 homotopic_with_imp_subset2) next assume ?rhs then show ?lhs - by (force simp: elim: homotopic_with_eq dest: homotopic_with_sphere_times [where h=g])+ + by (force simp: elim: homotopic_with_eq dest: homotopic_with_sphere_times [where h=g]) qed then show ?thesis by (simp add: *) qed subsection\Upper and lower hemicontinuous functions\ text\And relation in the case of preimage map to open and closed maps, and fact that upper and lower hemicontinuity together imply continuity in the sense of the Hausdorff metric (at points where the function gives a bounded and nonempty set).\ text\Many similar proofs below.\ lemma upper_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}) \ (\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ T - U}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ T - U} = {x \ S. f x \ U \ {}}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ (T - U) \ {}} = S - {x \ S. f x \ U}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma lower_hemicontinuous: assumes "\x. x \ S \ f x \ T" shows "((\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}) \ (\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "openin (top_of_set T) U" then have "closedin (top_of_set T) (T - U)" by (simp add: closedin_diff) then have "closedin (top_of_set S) {x \ S. f x \ T-U}" using * [of "T-U"] by blast moreover have "{x \ S. f x \ T-U} = S - {x \ S. f x \ U \ {}}" using assms by auto ultimately show "openin (top_of_set S) {x \ S. f x \ U \ {}}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and "closedin (top_of_set T) U" then have "openin (top_of_set T) (T - U)" by (simp add: openin_diff) then have "openin (top_of_set S) {x \ S. f x \ (T - U) \ {}}" using * [of "T-U"] by blast moreover have "S - {x \ S. f x \ (T - U) \ {}} = {x \ S. f x \ U}" using assms by blast ultimately show "closedin (top_of_set S) {x \ S. f x \ U}" by (simp add: openin_closedin_eq) qed lemma open_map_iff_lower_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)) \ (\U. closedin (top_of_set S) U \ closedin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "T - (f ` (S - U)) = {y \ T. {x \ S. f x = y} \ U}" using assms by blast ultimately show "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ U}" by (simp add: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "{y \ T. {x \ S. f x = y} \ S - U} = T - (f ` U)" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) (f ` U)" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) qed lemma closed_map_iff_upper_hemicontinuous_preimage: assumes "f ` S \ T" shows "((\U. closedin (top_of_set S) U \ closedin (top_of_set T) (f ` U)) \ (\U. openin (top_of_set S) U \ openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}))" (is "?lhs = ?rhs") proof (intro iffI allI impI) fix U assume * [rule_format]: ?lhs and opeSU: "openin (top_of_set S) U" then have "closedin (top_of_set S) (S - U)" by (simp add: closedin_diff) then have "closedin (top_of_set T) (f ` (S - U))" using * [of "S-U"] by blast moreover have "f ` (S - U) = T - {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by auto ultimately show "openin (top_of_set T) {y \ T. {x. x \ S \ f x = y} \ U}" using assms openin_imp_subset [OF opeSU] by (force simp: openin_closedin_eq) next fix U assume * [rule_format]: ?rhs and cloSU: "closedin (top_of_set S) U" then have "openin (top_of_set S) (S - U)" by (simp add: openin_diff) then have "openin (top_of_set T) {y \ T. {x \ S. f x = y} \ S - U}" using * [of "S-U"] by blast moreover have "(f ` U) = T - {y \ T. {x \ S. f x = y} \ S - U}" using assms closedin_imp_subset [OF cloSU] by auto ultimately show "closedin (top_of_set T) (f ` U)" by (simp add: openin_closedin_eq) qed proposition upper_lower_hemicontinuous_explicit: fixes T :: "('b::{real_normed_vector,heine_borel}) set" assumes fST: "\x. x \ S \ f x \ T" and ope: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U}" and clo: "\U. closedin (top_of_set T) U \ closedin (top_of_set S) {x \ S. f x \ U}" and "x \ S" "0 < e" and bofx: "bounded(f x)" and fx_ne: "f x \ {}" obtains d where "0 < d" "\x'. \x' \ S; dist x x' < d\ \ (\y \ f x. \y'. y' \ f x' \ dist y y' < e) \ (\y' \ f x'. \y. y \ f x \ dist y' y < e)" proof - have "openin (top_of_set T) (T \ (\a\f x. \b\ball 0 e. {a + b}))" by (auto simp: open_sums openin_open_Int) with ope have "openin (top_of_set S) {u \ S. f u \ T \ (\a\f x. \b\ball 0 e. {a + b})}" by blast with \0 < e\ \x \ S\ obtain d1 where "d1 > 0" and d1: "\x'. \x' \ S; dist x' x < d1\ \ f x' \ T \ f x' \ (\a \ f x. \b \ ball 0 e. {a + b})" by (force simp: openin_euclidean_subtopology_iff dest: fST) have oo: "\U. openin (top_of_set T) U \ openin (top_of_set S) {x \ S. f x \ U \ {}}" apply (rule lower_hemicontinuous [THEN iffD1, rule_format]) using fST clo by auto have "compact (closure(f x))" by (simp add: bofx) moreover have "closure(f x) \ (\a \ f x. ball a (e/2))" using \0 < e\ by (force simp: closure_approachable simp del: divide_const_simps) ultimately obtain C where "C \ f x" "finite C" "closure(f x) \ (\a \ C. ball a (e/2))" apply (rule compactE, force) by (metis finite_subset_image) then have fx_cover: "f x \ (\a \ C. ball a (e/2))" by (meson closure_subset order_trans) with fx_ne have "C \ {}" by blast have xin: "x \ (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" using \x \ S\ \0 < e\ fST \C \ f x\ by force have "openin (top_of_set S) {x \ S. f x \ (T \ ball a (e/2)) \ {}}" for a by (simp add: openin_open_Int oo) then have "openin (top_of_set S) (\a \ C. {x \ S. f x \ T \ ball a (e/2) \ {}})" by (simp add: Int_assoc openin_INT2 [OF \finite C\ \C \ {}\]) with xin obtain d2 where "d2>0" and d2: "\u v. \u \ S; dist u x < d2; v \ C\ \ f u \ T \ ball v (e/2) \ {}" unfolding openin_euclidean_subtopology_iff using xin by fastforce show ?thesis proof (intro that conjI ballI) show "0 < min d1 d2" using \0 < d1\ \0 < d2\ by linarith next fix x' y assume "x' \ S" "dist x x' < min d1 d2" "y \ f x" then have dd2: "dist x' x < d2" by (auto simp: dist_commute) obtain a where "a \ C" "y \ ball a (e/2)" using fx_cover \y \ f x\ by auto then show "\y'. y' \ f x' \ dist y y' < e" using d2 [OF \x' \ S\ dd2] dist_triangle_half_r by fastforce next fix x' y' assume "x' \ S" "dist x x' < min d1 d2" "y' \ f x'" then have "dist x' x < d1" by (auto simp: dist_commute) then have "y' \ (\a\f x. \b\ball 0 e. {a + b})" using d1 [OF \x' \ S\] \y' \ f x'\ by force then show "\y. y \ f x \ dist y' y < e" apply auto by (metis add_diff_cancel_left' dist_norm) qed qed subsection\Complex logs exist on various "well-behaved" sets\ lemma continuous_logarithm_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" proof - obtain c where hom: "homotopic_with_canon (\h. True) S (-{0}) f (\x. c)" using nullhomotopic_from_contractible assms by (metis imageE subset_Compl_singleton) then show ?thesis by (metis inessential_eq_continuous_logarithm that) qed lemma continuous_logarithm_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = exp(g x)" using covering_space_lift [OF covering_space_exp_punctured_plane S contf] by (metis (full_types) f imageE subset_Compl_singleton) lemma continuous_logarithm_on_cball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (cball a r) f" and "\z. z \ cball a r \ f z \ 0" obtains h where "continuous_on (cball a r) h" "\z. z \ cball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_logarithm_on_ball: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on (ball a r) f" and "\z. z \ ball a r \ f z \ 0" obtains h where "continuous_on (ball a r) h" "\z. z \ ball a r \ f z = exp(h z)" using assms continuous_logarithm_on_contractible convex_imp_contractible by blast lemma continuous_sqrt_on_contractible: fixes f :: "'a::real_normed_vector \ complex" assumes "continuous_on S f" "contractible S" and "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_contractible [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed lemma continuous_sqrt_on_simply_connected: fixes f :: "'a::real_normed_vector \ complex" assumes contf: "continuous_on S f" and S: "simply_connected S" "locally path_connected S" and f: "\z. z \ S \ f z \ 0" obtains g where "continuous_on S g" "\x. x \ S \ f x = (g x) ^ 2" proof - obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp(g x)" using continuous_logarithm_on_simply_connected [OF assms] by blast show ?thesis proof show "continuous_on S (\z. exp (g z / 2))" by (rule continuous_on_compose2 [of UNIV exp]; intro continuous_intros contg subset_UNIV) auto show "\x. x \ S \ f x = (exp (g x / 2))\<^sup>2" by (metis exp_double feq nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral) qed qed subsection\Another simple case where sphere maps are nullhomotopic\ lemma inessential_spheremap_2_aux: fixes f :: "'a::euclidean_space \ complex" assumes 2: "2 < DIM('a)" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere 0 1)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) f (\x. c)" proof - obtain g where contg: "continuous_on (sphere a r) g" and feq: "\x. x \ sphere a r \ f x = exp(g x)" proof (rule continuous_logarithm_on_simply_connected [OF contf]) show "simply_connected (sphere a r)" using 2 by (simp add: simply_connected_sphere_eq) show "locally path_connected (sphere a r)" by (simp add: locally_path_connected_sphere) show "\z. z \ sphere a r \ f z \ 0" using fim by force qed auto have "\g. continuous_on (sphere a r) g \ (\x\sphere a r. f x = exp (\ * complex_of_real (g x)))" proof (intro exI conjI) show "continuous_on (sphere a r) (Im \ g)" by (intro contg continuous_intros continuous_on_compose) show "\x\sphere a r. f x = exp (\ * complex_of_real ((Im \ g) x))" using exp_eq_polar feq fim norm_exp_eq_Re by auto qed with inessential_eq_continuous_logarithm_circle that show ?thesis by metis qed lemma inessential_spheremap_2: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes a2: "2 < DIM('a)" and b2: "DIM('b) = 2" and contf: "continuous_on (sphere a r) f" and fim: "f `(sphere a r) \ (sphere b s)" obtains c where "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. c)" proof (cases "s \ 0") case True then show ?thesis using contf contractible_sphere fim nullhomotopic_into_contractible that by blast next case False then have "sphere b s homeomorphic sphere (0::complex) 1" using assms by (simp add: homeomorphic_spheres_gen) then obtain h k where hk: "homeomorphism (sphere b s) (sphere (0::complex) 1) h k" by (auto simp: homeomorphic_def) then have conth: "continuous_on (sphere b s) h" and contk: "continuous_on (sphere 0 1) k" and him: "h ` sphere b s \ sphere 0 1" and kim: "k ` sphere 0 1 \ sphere b s" by (simp_all add: homeomorphism_def) obtain c where "homotopic_with_canon (\z. True) (sphere a r) (sphere 0 1) (h \ f) (\x. c)" proof (rule inessential_spheremap_2_aux [OF a2]) show "continuous_on (sphere a r) (h \ f)" by (meson continuous_on_compose [OF contf] conth continuous_on_subset fim) show "(h \ f) ` sphere a r \ sphere 0 1" using fim him by force qed auto then have "homotopic_with_canon (\f. True) (sphere a r) (sphere b s) (k \ (h \ f)) (k \ (\x. c))" by (rule homotopic_compose_continuous_left [OF _ contk kim]) then have "homotopic_with_canon (\z. True) (sphere a r) (sphere b s) f (\x. k c)" apply (rule homotopic_with_eq, auto) by (metis fim hk homeomorphism_def image_subset_iff mem_sphere) then show ?thesis by (metis that) qed subsection\Holomorphic logarithms and square roots\ lemma contractible_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_contractible [OF contf S fnz]) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" apply (subst Lim_at_zero) apply (simp add: DERIV_D cong: if_cong Lim_cong_within) done qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" apply (rule eventually_mono) apply (auto simp: exp_eq dist_norm norm_mult) done then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed (*Identical proofs*) lemma simply_connected_imp_holomorphic_log: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = exp(g z)" proof - have contf: "continuous_on S f" by (simp add: holf holomorphic_on_imp_continuous_on) obtain g where contg: "continuous_on S g" and feq: "\x. x \ S \ f x = exp (g x)" by (metis continuous_logarithm_on_simply_connected [OF contf S fnz]) have "g field_differentiable at z within S" if "f field_differentiable at z within S" "z \ S" for z proof - obtain f' where f': "((\y. (f y - f z) / (y - z)) \ f') (at z within S)" using \f field_differentiable at z within S\ by (auto simp: field_differentiable_def has_field_derivative_iff) then have ee: "((\x. (exp(g x) - exp(g z)) / (x - z)) \ f') (at z within S)" by (simp add: feq \z \ S\ Lim_transform_within [OF _ zero_less_one]) have "(((\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \ g) \ exp (g z)) (at z within S)" proof (rule tendsto_compose_at) show "(g \ g z) (at z within S)" using contg continuous_on \z \ S\ by blast show "(\y. if y = g z then exp (g z) else (exp y - exp (g z)) / (y - g z)) \g z\ exp (g z)" apply (subst Lim_at_zero) apply (simp add: DERIV_D cong: if_cong Lim_cong_within) done qed auto then have dd: "((\x. if g x = g z then exp(g z) else (exp(g x) - exp(g z)) / (g x - g z)) \ exp(g z)) (at z within S)" by (simp add: o_def) have "continuous (at z within S) g" using contg continuous_on_eq_continuous_within \z \ S\ by blast then have "(\\<^sub>F x in at z within S. dist (g x) (g z) < 2*pi)" by (simp add: continuous_within tendsto_iff) then have "\\<^sub>F x in at z within S. exp (g x) = exp (g z) \ g x \ g z \ x = z" apply (rule eventually_mono) apply (auto simp: exp_eq dist_norm norm_mult) done then have "((\y. (g y - g z) / (y - z)) \ f' / exp (g z)) (at z within S)" by (auto intro!: Lim_transform_eventually [OF tendsto_divide [OF ee dd]]) then show ?thesis by (auto simp: field_differentiable_def has_field_derivative_iff) qed then have "g holomorphic_on S" using holf holomorphic_on_def by auto then show ?thesis using feq that by auto qed lemma contractible_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "contractible S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using contractible_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" apply (auto simp: feq) by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) qed qed lemma simply_connected_imp_holomorphic_sqrt: assumes holf: "f holomorphic_on S" and S: "simply_connected S" "locally path_connected S" and fnz: "\z. z \ S \ f z \ 0" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = g z ^ 2" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = exp(g z)" using simply_connected_imp_holomorphic_log [OF assms] by blast show ?thesis proof show "exp \ (\z. z / 2) \ g holomorphic_on S" by (intro holomorphic_on_compose holg holomorphic_intros) auto show "\z. z \ S \ f z = ((exp \ (\z. z / 2) \ g) z)\<^sup>2" apply (auto simp: feq) by (metis eq_divide_eq_numeral1(1) exp_double mult.commute zero_neq_numeral) qed qed text\ Related theorems about holomorphic inverse cosines.\ lemma contractible_imp_holomorphic_arccos: assumes holf: "f holomorphic_on S" and S: "contractible S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "\z. z \ S \ f z = cos(g z)" proof - have hol1f: "(\z. 1 - f z ^ 2) holomorphic_on S" by (intro holomorphic_intros holf) obtain g where holg: "g holomorphic_on S" and eq: "\z. z \ S \ 1 - (f z)\<^sup>2 = (g z)\<^sup>2" using contractible_imp_holomorphic_sqrt [OF hol1f S] by (metis eq_iff_diff_eq_0 non1 power2_eq_1_iff) have holfg: "(\z. f z + \*g z) holomorphic_on S" by (intro holf holg holomorphic_intros) have "\z. z \ S \ f z + \*g z \ 0" by (metis Arccos_body_lemma eq add.commute add.inverse_unique complex_i_mult_minus power2_csqrt power2_eq_iff) then obtain h where holh: "h holomorphic_on S" and fgeq: "\z. z \ S \ f z + \*g z = exp (h z)" using contractible_imp_holomorphic_log [OF holfg S] by metis show ?thesis proof show "(\z. -\*h z) holomorphic_on S" by (intro holh holomorphic_intros) show "f z = cos (- \*h z)" if "z \ S" for z proof - have "(f z + \*g z)*(f z - \*g z) = 1" using that eq by (auto simp: algebra_simps power2_eq_square) then have "f z - \*g z = inverse (f z + \*g z)" using inverse_unique by force also have "... = exp (- h z)" by (simp add: exp_minus fgeq that) finally have "f z = exp (- h z) + \*g z" by (simp add: diff_eq_eq) then show ?thesis apply (simp add: cos_exp_eq) by (metis fgeq add.assoc mult_2_right that) qed qed qed lemma contractible_imp_holomorphic_arccos_bounded: assumes holf: "f holomorphic_on S" and S: "contractible S" and "a \ S" and non1: "\z. z \ S \ f z \ 1 \ f z \ -1" obtains g where "g holomorphic_on S" "norm(g a) \ pi + norm(f a)" "\z. z \ S \ f z = cos(g z)" proof - obtain g where holg: "g holomorphic_on S" and feq: "\z. z \ S \ f z = cos (g z)" using contractible_imp_holomorphic_arccos [OF holf S non1] by blast obtain b where "cos b = f a" "norm b \ pi + norm (f a)" using cos_Arccos norm_Arccos_bounded by blast then have "cos b = cos (g a)" by (simp add: \a \ S\ feq) then consider n where "n \ \" "b = g a + of_real(2*n*pi)" | n where "n \ \" "b = -g a + of_real(2*n*pi)" by (auto simp: complex_cos_eq) then show ?thesis proof cases case 1 show ?thesis proof show "(\z. g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "1" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed next case 2 show ?thesis proof show "(\z. -g z + of_real(2*n*pi)) holomorphic_on S" by (intro holomorphic_intros holg) show "cmod (-g a + of_real(2*n*pi)) \ pi + cmod (f a)" using "2" \cmod b \ pi + cmod (f a)\ by blast show "\z. z \ S \ f z = cos (-g z + complex_of_real (2*n*pi))" by (metis \n \ \\ complex_cos_eq feq) qed qed qed subsection\The "Borsukian" property of sets\ text\This doesn't have a standard name. Kuratowski uses ``contractible with respect to \[S\<^sup>1]\'' while Whyburn uses ``property b''. It's closely related to unicoherence.\ definition\<^marker>\tag important\ Borsukian where "Borsukian S \ \f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\a. homotopic_with_canon (\h. True) S (- {0}) f (\x. a))" lemma Borsukian_retraction_gen: assumes "Borsukian S" "continuous_on S h" "h ` S = T" "continuous_on T k" "k ` T \ S" "\y. y \ T \ h(k y) = y" shows "Borsukian T" proof - interpret R: Retracts S h T k using assms by (simp add: Retracts.intro) show ?thesis using assms apply (simp add: Borsukian_def, clarify) apply (rule R.cohomotopically_trivial_retraction_null_gen [OF TrueI TrueI refl, of "-{0}"], auto) done qed lemma retract_of_Borsukian: "\Borsukian T; S retract_of T\ \ Borsukian S" apply (auto simp: retract_of_def retraction_def) apply (erule (1) Borsukian_retraction_gen) apply (meson retraction retraction_def) apply (auto simp: continuous_on_id) done lemma homeomorphic_Borsukian: "\Borsukian S; S homeomorphic T\ \ Borsukian T" using Borsukian_retraction_gen order_refl by (fastforce simp add: homeomorphism_def homeomorphic_def) lemma homeomorphic_Borsukian_eq: "S homeomorphic T \ Borsukian S \ Borsukian T" by (meson homeomorphic_Borsukian homeomorphic_sym) lemma Borsukian_translation: fixes S :: "'a::real_normed_vector set" shows "Borsukian (image (\x. a + x) S) \ Borsukian S" apply (rule homeomorphic_Borsukian_eq) using homeomorphic_translation homeomorphic_sym by blast lemma Borsukian_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "Borsukian(f ` S) \ Borsukian S" apply (rule homeomorphic_Borsukian_eq) using assms homeomorphic_sym linear_homeomorphic_image by blast lemma homotopy_eqv_Borsukianness: fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set" assumes "S homotopy_eqv T" shows "(Borsukian S \ Borsukian T)" by (meson Borsukian_def assms homotopy_eqv_cohomotopic_triviality_null) lemma Borsukian_alt: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f g. continuous_on S f \ f ` S \ -{0} \ continuous_on S g \ g ` S \ -{0} \ homotopic_with_canon (\h. True) S (- {0::complex}) f g)" unfolding Borsukian_def homotopic_triviality by (simp add: path_connected_punctured_universe) lemma Borsukian_continuous_logarithm: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ (- {0::complex}) \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" by (simp add: Borsukian_def inessential_eq_continuous_logarithm) lemma Borsukian_continuous_logarithm_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S g \ (\x \ S. f x = exp(g x))))" (is "?lhs = ?rhs") proof assume ?lhs then show ?rhs by (force simp: Borsukian_continuous_logarithm) next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on S f" and 0: "0 \ f ` S" then have "continuous_on S (\x. f x / complex_of_real (cmod (f x)))" by (intro continuous_intros) auto moreover have "(\x. f x / complex_of_real (cmod (f x))) ` S \ sphere 0 1" using 0 by (auto simp: norm_divide) ultimately obtain g where contg: "continuous_on S g" and fg: "\x \ S. f x / complex_of_real (cmod (f x)) = exp(g x)" using RHS [of "\x. f x / of_real(norm(f x))"] by auto show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" proof (intro exI ballI conjI) show "continuous_on S (\x. (Ln \ of_real \ norm \ f)x + g x)" by (intro continuous_intros contf contg conjI) (use "0" in auto) show "f x = exp ((Ln \ complex_of_real \ cmod \ f) x + g x)" if "x \ S" for x using 0 that apply (clarsimp simp: exp_add) apply (subst exp_Ln, force) by (metis eq_divide_eq exp_not_eq_zero fg mult.commute) qed qed qed lemma Borsukian_continuous_logarithm_circle_real: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\g. continuous_on S (complex_of_real \ g) \ (\x \ S. f x = exp(\ * of_real(g x)))))" (is "?lhs = ?rhs") proof assume LHS: ?lhs show ?rhs proof (clarify) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S g" and "\x. x \ S \ f x = exp(g x)" using LHS by (auto simp: Borsukian_continuous_logarithm_circle) then have "\x\S. f x = exp (\ * complex_of_real ((Im \ g) x))" using f01 apply (simp add: image_iff subset_iff) by (metis cis_conv_exp exp_eq_polar mult.left_neutral norm_exp_eq_Re of_real_1) then show "\g. continuous_on S (complex_of_real \ g) \ (\x\S. f x = exp (\ * complex_of_real (g x)))" by (rule_tac x="Im \ g" in exI) (force intro: continuous_intros contg) qed next assume RHS [rule_format]: ?rhs show ?lhs proof (clarsimp simp: Borsukian_continuous_logarithm_circle) fix f :: "'a \ complex" assume "continuous_on S f" and f01: "f ` S \ sphere 0 1" then obtain g where contg: "continuous_on S (complex_of_real \ g)" and "\x. x \ S \ f x = exp(\ * of_real(g x))" by (metis RHS) then show "\g. continuous_on S g \ (\x\S. f x = exp (g x))" by (rule_tac x="\x. \* of_real(g x)" in exI) (auto simp: continuous_intros contg) qed qed lemma Borsukian_circle: fixes S :: "'a::real_normed_vector set" shows "Borsukian S \ (\f. continuous_on S f \ f ` S \ sphere (0::complex) 1 \ (\a. homotopic_with_canon (\h. True) S (sphere (0::complex) 1) f (\x. a)))" by (simp add: inessential_eq_continuous_logarithm_circle Borsukian_continuous_logarithm_circle_real) lemma contractible_imp_Borsukian: "contractible S \ Borsukian S" by (meson Borsukian_def nullhomotopic_from_contractible) lemma simply_connected_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "\simply_connected S; locally path_connected S\ \ Borsukian S" apply (simp add: Borsukian_continuous_logarithm) by (metis (no_types, lifting) continuous_logarithm_on_simply_connected image_iff) lemma starlike_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "starlike S \ Borsukian S" by (simp add: contractible_imp_Borsukian starlike_imp_contractible) lemma Borsukian_empty: "Borsukian {}" by (auto simp: contractible_imp_Borsukian) lemma Borsukian_UNIV: "Borsukian (UNIV :: 'a::real_normed_vector set)" by (auto simp: contractible_imp_Borsukian) lemma convex_imp_Borsukian: fixes S :: "'a::real_normed_vector set" shows "convex S \ Borsukian S" by (meson Borsukian_def convex_imp_contractible nullhomotopic_from_contractible) proposition Borsukian_sphere: fixes a :: "'a::euclidean_space" shows "3 \ DIM('a) \ Borsukian (sphere a r)" apply (rule simply_connected_imp_Borsukian) using simply_connected_sphere apply blast using ENR_imp_locally_path_connected ENR_sphere by blast proposition Borsukian_open_Un: fixes S :: "'a::real_normed_vector set" assumes opeS: "openin (top_of_set (S \ T)) S" and opeT: "openin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" apply (rule continuous_on_cases_local_open [OF opeS opeT contg conth]) using True by blast show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) apply (meson contg continuous_on_subset inf_le1) by (meson conth continuous_on_subset inf_sup_ord(2)) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by (auto simp: constant_on_def) with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) apply (intro continuous_on_cases_local_open opeS opeT contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) done qed qed text\The proof is a duplicate of that of \Borsukian_open_Un\.\ lemma Borsukian_closed_Un: fixes S :: "'a::real_normed_vector set" assumes cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" and BS: "Borsukian S" and BT: "Borsukian T" and ST: "connected(S \ T)" shows "Borsukian(S \ T)" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix f :: "'a \ complex" assume contf: "continuous_on (S \ T) f" and 0: "0 \ f ` (S \ T)" then have contfS: "continuous_on S f" and contfT: "continuous_on T f" using continuous_on_subset by auto have "\continuous_on S f; f ` S \ -{0}\ \ \g. continuous_on S g \ (\x \ S. f x = exp(g x))" using BS by (auto simp: Borsukian_continuous_logarithm) then obtain g where contg: "continuous_on S g" and fg: "\x. x \ S \ f x = exp(g x)" using "0" contfS by blast have "\continuous_on T f; f ` T \ -{0}\ \ \g. continuous_on T g \ (\x \ T. f x = exp(g x))" using BT by (auto simp: Borsukian_continuous_logarithm) then obtain h where conth: "continuous_on T h" and fh: "\x. x \ T \ f x = exp(h x)" using "0" contfT by blast show "\g. continuous_on (S \ T) g \ (\x\S \ T. f x = exp (g x))" proof (cases "S \ T = {}") case True show ?thesis proof (intro exI conjI) show "continuous_on (S \ T) (\x. if x \ S then g x else h x)" apply (rule continuous_on_cases_local [OF cloS cloT contg conth]) using True by blast show "\x\S \ T. f x = exp (if x \ S then g x else h x)" using fg fh by auto qed next case False have "(\x. g x - h x) constant_on S \ T" proof (rule continuous_discrete_range_constant [OF ST]) show "continuous_on (S \ T) (\x. g x - h x)" apply (intro continuous_intros) apply (meson contg continuous_on_subset inf_le1) by (meson conth continuous_on_subset inf_sup_ord(2)) show "\e>0. \y. y \ S \ T \ g y - h y \ g x - h x \ e \ cmod (g y - h y - (g x - h x))" if "x \ S \ T" for x proof - have "g y - g x = h y - h x" if "y \ S" "y \ T" "cmod (g y - g x - (h y - h x)) < 2 * pi" for y proof (rule exp_complex_eqI) have "\Im (g y - g x) - Im (h y - h x)\ \ cmod (g y - g x - (h y - h x))" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (g y - g x) - Im (h y - h x)\ < 2 * pi" using that by linarith have "exp (g x) = exp (h x)" "exp (g y) = exp (h y)" using fg fh that \x \ S \ T\ by fastforce+ then show "exp (g y - g x) = exp (h y - h x)" by (simp add: exp_diff) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed then obtain a where a: "\x. x \ S \ T \ g x - h x = a" by (auto simp: constant_on_def) with False have "exp a = 1" by (metis IntI disjoint_iff_not_equal divide_self_if exp_diff exp_not_eq_zero fg fh) with a show ?thesis apply (rule_tac x="\x. if x \ S then g x else a + h x" in exI) apply (intro continuous_on_cases_local cloS cloT contg conth continuous_intros conjI) apply (auto simp: algebra_simps fg fh exp_add) done qed qed lemma Borsukian_separation_compact: fixes S :: "complex set" assumes "compact S" shows "Borsukian S \ connected(- S)" by (simp add: Borsuk_separation_theorem Borsukian_circle assms) lemma Borsukian_monotone_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and conn: "\y. y \ T \ connected {x. x \ S \ f x = y}" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm) fix g :: "'b \ complex" assume contg: "continuous_on T g" and 0: "0 \ g ` T" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ -{0}" using fim 0 by auto ultimately obtain h where conth: "continuous_on S h" and gfh: "\x. x \ S \ (g \ f) x = exp(h x)" using \Borsukian S\ by (auto simp: Borsukian_continuous_logarithm) have "\y. \x. y \ T \ x \ S \ f x = y" using fim by auto then obtain f' where f': "\y. y \ T \ f' y \ S \ f (f' y) = y" by metis have *: "(\x. h x - h(f' y)) constant_on {x. x \ S \ f x = y}" if "y \ T" for y proof (rule continuous_discrete_range_constant [OF conn [OF that], of "\x. h x - h (f' y)"], simp_all add: algebra_simps) show "continuous_on {x \ S. f x = y} (\x. h x - h (f' y))" by (intro continuous_intros continuous_on_subset [OF conth]) auto show "\e>0. \u. u \ S \ f u = y \ h u \ h x \ e \ cmod (h u - h x)" if x: "x \ S \ f x = y" for x proof - have "h u = h x" if "u \ S" "f u = y" "cmod (h u - h x) < 2 * pi" for u proof (rule exp_complex_eqI) have "\Im (h u) - Im (h x)\ \ cmod (h u - h x)" by (metis abs_Im_le_cmod minus_complex.simps(2)) then show "\Im (h u) - Im (h x)\ < 2 * pi" using that by linarith show "exp (h u) = exp (h x)" by (simp add: gfh [symmetric] x that) qed then show ?thesis by (rule_tac x="2*pi" in exI) (fastforce simp add: algebra_simps) qed qed have "h x = h (f' (f x))" if "x \ S" for x using * [of "f x"] fim that unfolding constant_on_def by clarsimp (metis f' imageI right_minus_eq) moreover have "\x. x \ T \ \u. u \ S \ x = f u \ h (f' x) = h u" using f' by fastforce ultimately have eq: "((\x. (x, (h \ f') x)) ` T) = {p. \x. x \ S \ (x, p) \ (S \ UNIV) \ ((\z. snd z - ((f \ fst) z, (h \ fst) z)) -` {0})}" using fim by (auto simp: image_iff) show "\h. continuous_on T h \ (\x\T. g x = exp (h x))" proof (intro exI conjI) show "continuous_on T (h \ f')" proof (rule continuous_from_closed_graph [of "h ` S"]) show "compact (h ` S)" by (simp add: \compact S\ compact_continuous_image conth) show "(h \ f') ` T \ h ` S" by (auto simp: f') show "closed ((\x. (x, (h \ f') x)) ` T)" apply (subst eq) apply (intro closed_compact_projection [OF \compact S\] continuous_closed_preimage continuous_intros continuous_on_subset [OF contf] continuous_on_subset [OF conth]) apply (auto simp: \compact S\ closed_Times compact_imp_closed) done qed qed (use f' gfh in fastforce) qed lemma Borsukian_open_map_image_compact: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "Borsukian S" and contf: "continuous_on S f" and fim: "f ` S = T" and "compact S" and ope: "\U. openin (top_of_set S) U \ openin (top_of_set T) (f ` U)" shows "Borsukian T" proof (clarsimp simp add: Borsukian_continuous_logarithm_circle_real) fix g :: "'b \ complex" assume contg: "continuous_on T g" and gim: "g ` T \ sphere 0 1" have "continuous_on S (g \ f)" using contf contg continuous_on_compose fim by blast moreover have "(g \ f) ` S \ sphere 0 1" using fim gim by auto ultimately obtain h where cont_cxh: "continuous_on S (complex_of_real \ h)" and gfh: "\x. x \ S \ (g \ f) x = exp(\ * of_real(h x))" using \Borsukian S\ Borsukian_continuous_logarithm_circle_real by metis then have conth: "continuous_on S h" by simp have "\x. x \ S \ f x = y \ (\x' \ S. f x' = y \ h x \ h x')" if "y \ T" for y proof - have 1: "compact (h ` {x \ S. f x = y})" proof (rule compact_continuous_image) show "continuous_on {x \ S. f x = y} h" by (rule continuous_on_subset [OF conth]) auto have "compact (S \ f -` {y})" by (rule proper_map_from_compact [OF contf _ \compact S\, of T]) (simp_all add: fim that) then show "compact {x \ S. f x = y}" by (auto simp: vimage_def Int_def) qed have 2: "h ` {x \ S. f x = y} \ {}" using fim that by auto have "\s \ h ` {x \ S. f x = y}. \t \ h ` {x \ S. f x = y}. s \ t" using compact_attains_inf [OF 1 2] by blast then show ?thesis by auto qed then obtain k where kTS: "\y. y \ T \ k y \ S" and fk: "\y. y \ T \ f (k y) = y " and hle: "\x' y. \y \ T; x' \ S; f x' = y\ \ h (k y) \ h x'" by metis have "continuous_on T (h \ k)" proof (clarsimp simp add: continuous_on_iff) fix y and e::real assume "y \ T" "0 < e" moreover have "uniformly_continuous_on S (complex_of_real \ h)" using \compact S\ cont_cxh compact_uniformly_continuous by blast ultimately obtain d where "0 < d" and d: "\x x'. \x\S; x'\S; dist x' x < d\ \ dist (h x') (h x) < e" by (force simp: uniformly_continuous_on_def) obtain \ where "0 < \" and \: "\x'. \x' \ T; dist y x' < \\ \ (\v \ {z \ S. f z = y}. \v'. v' \ {z \ S. f z = x'} \ dist v v' < d) \ (\v' \ {z \ S. f z = x'}. \v. v \ {z \ S. f z = y} \ dist v' v < d)" proof (rule upper_lower_hemicontinuous_explicit [of T "\y. {z \ S. f z = y}" S]) show "\U. openin (top_of_set S) U \ openin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using closed_map_iff_upper_hemicontinuous_preimage [OF fim [THEN equalityD1]] by (simp add: Abstract_Topology_2.continuous_imp_closed_map \compact S\ contf fim) show "\U. closedin (top_of_set S) U \ closedin (top_of_set T) {x \ T. {z \ S. f z = x} \ U}" using ope open_map_iff_lower_hemicontinuous_preimage [OF fim [THEN equalityD1]] by meson show "bounded {z \ S. f z = y}" by (metis (no_types, lifting) compact_imp_bounded [OF \compact S\] bounded_subset mem_Collect_eq subsetI) qed (use \y \ T\ \0 < d\ fk kTS in \force+\) have "dist (h (k y')) (h (k y)) < e" if "y' \ T" "dist y y' < \" for y' proof - have k1: "k y \ S" "f (k y) = y" and k2: "k y' \ S" "f (k y') = y'" by (auto simp: \y \ T\ \y' \ T\ kTS fk) have 1: "\v. \v \ S; f v = y\ \ \v'. v' \ {z \ S. f z = y'} \ dist v v' < d" and 2: "\v'. \v' \ S; f v' = y'\ \ \v. v \ {z \ S. f z = y} \ dist v' v < d" using \ [OF that] by auto then obtain w' w where "w' \ S" "f w' = y'" "dist (k y) w' < d" and "w \ S" "f w = y" "dist (k y') w < d" using 1 [OF k1] 2 [OF k2] by auto then show ?thesis using d [of w "k y'"] d [of w' "k y"] k1 k2 \y' \ T\ \y \ T\ hle by (fastforce simp: dist_norm abs_diff_less_iff algebra_simps) qed then show "\d>0. \x'\T. dist x' y < d \ dist (h (k x')) (h (k y)) < e" using \0 < \\ by (auto simp: dist_commute) qed then show "\h. continuous_on T h \ (\x\T. g x = exp (\ * complex_of_real (h x)))" using fk gfh kTS by force qed text\If two points are separated by a closed set, there's a minimal one.\ proposition closed_irreducible_separator: fixes a :: "'a::real_normed_vector" assumes "closed S" and ab: "\ connected_component (- S) a b" obtains T where "T \ S" "closed T" "T \ {}" "\ connected_component (- T) a b" "\U. U \ T \ connected_component (- U) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis proof assume *: "a \ S" show ?thesis proof show "{a} \ S" using * by blast show "\ connected_component (- {a}) a b" using connected_component_in by auto show "\U. U \ {a} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto next assume *: "b \ S" show ?thesis proof show "{b} \ S" using * by blast show "\ connected_component (- {b}) a b" using connected_component_in by auto show "\U. U \ {b} \ connected_component (- U) a b" by (metis connected_component_UNIV UNIV_I compl_bot_eq connected_component_eq_eq less_le_not_le subset_singletonD) qed auto qed next case False define A where "A \ connected_component_set (- S) a" define B where "B \ connected_component_set (- (closure A)) b" have "a \ A" using False A_def by auto have "b \ B" unfolding A_def B_def closure_Un_frontier using ab False \closed S\ frontier_complement frontier_of_connected_component_subset frontier_subset_closed by force have "frontier B \ frontier (connected_component_set (- closure A) b)" using B_def by blast also have frsub: "... \ frontier A" proof - have "\A. closure (- closure (- A)) \ closure A" by (metis (no_types) closure_mono closure_subset compl_le_compl_iff double_compl) then show ?thesis by (metis (no_types) closure_closure double_compl frontier_closures frontier_of_connected_component_subset le_inf_iff subset_trans) qed finally have frBA: "frontier B \ frontier A" . show ?thesis proof show "frontier B \ S" proof - have "frontier S \ S" by (simp add: \closed S\ frontier_subset_closed) then show ?thesis using frsub frontier_complement frontier_of_connected_component_subset unfolding A_def B_def by blast qed show "closed (frontier B)" by simp show "\ connected_component (- frontier B) a b" unfolding connected_component_def proof clarify fix T assume "connected T" and TB: "T \ - frontier B" and "a \ T" and "b \ T" have "a \ B" by (metis A_def B_def ComplD \a \ A\ assms(1) closed_open connected_component_subset in_closure_connected_component subsetD) have "T \ B \ {}" using \b \ B\ \b \ T\ by blast moreover have "T - B \ {}" using \a \ B\ \a \ T\ by blast ultimately show "False" using connected_Int_frontier [of T B] TB \connected T\ by blast qed moreover have "connected_component (- frontier B) a b" if "frontier B = {}" apply (simp add: that) using connected_component_eq_UNIV by blast ultimately show "frontier B \ {}" by blast show "connected_component (- U) a b" if "U \ frontier B" for U proof - obtain p where Usub: "U \ frontier B" and p: "p \ frontier B" "p \ U" using \U \ frontier B\ by blast show ?thesis unfolding connected_component_def proof (intro exI conjI) have "connected ((insert p A) \ (insert p B))" proof (rule connected_Un) show "connected (insert p A)" by (metis A_def IntD1 frBA \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subsetCE subset_insertI) show "connected (insert p B)" by (metis B_def IntD1 \p \ frontier B\ closure_insert closure_subset connected_connected_component connected_intermediate_closure frontier_closures insert_absorb subset_insertI) qed blast then show "connected (insert p (B \ A))" by (simp add: sup.commute) have "A \ - U" using A_def Usub \frontier B \ S\ connected_component_subset by fastforce moreover have "B \ - U" using B_def Usub connected_component_subset frBA frontier_closures by fastforce ultimately show "insert p (B \ A) \ - U" using p by auto qed (auto simp: \a \ A\ \b \ B\) qed qed qed lemma frontier_minimal_separating_closed_pointwise: fixes S :: "'a::real_normed_vector set" assumes S: "closed S" "a \ S" and nconn: "\ connected_component (- S) a b" and conn: "\T. \closed T; T \ S\ \ connected_component (- T) a b" shows "frontier(connected_component_set (- S) a) = S" (is "?F = S") proof - have "?F \ S" by (simp add: S componentsI frontier_of_components_closed_complement) moreover have False if "?F \ S" proof - have "connected_component (- ?F) a b" by (simp add: conn that) then obtain T where "connected T" "T \ -?F" "a \ T" "b \ T" by (auto simp: connected_component_def) moreover have "T \ ?F \ {}" proof (rule connected_Int_frontier [OF \connected T\]) show "T \ connected_component_set (- S) a \ {}" using \a \ S\ \a \ T\ by fastforce show "T - connected_component_set (- S) a \ {}" using \b \ T\ nconn by blast qed ultimately show ?thesis by blast qed ultimately show ?thesis by blast qed subsection\Unicoherence (closed)\ definition\<^marker>\tag important\ unicoherent where "unicoherent U \ \S T. connected S \ connected T \ S \ T = U \ closedin (top_of_set U) S \ closedin (top_of_set U) T \ connected (S \ T)" lemma unicoherentI [intro?]: assumes "\S T. \connected S; connected T; U = S \ T; closedin (top_of_set U) S; closedin (top_of_set U) T\ \ connected (S \ T)" shows "unicoherent U" using assms unfolding unicoherent_def by blast lemma unicoherentD: assumes "unicoherent U" "connected S" "connected T" "U = S \ T" "closedin (top_of_set U) S" "closedin (top_of_set U) T" shows "connected (S \ T)" using assms unfolding unicoherent_def by blast proposition homeomorphic_unicoherent: assumes ST: "S homeomorphic T" and S: "unicoherent S" shows "unicoherent T" proof - obtain f g where gf: "\x. x \ S \ g (f x) = x" and fim: "T = f ` S" and gfim: "g ` f ` S = S" and contf: "continuous_on S f" and contg: "continuous_on (f ` S) g" using ST by (auto simp: homeomorphic_def homeomorphism_def) show ?thesis proof fix U V assume "connected U" "connected V" and T: "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" have "f ` (g ` U \ g ` V) \ U" "f ` (g ` U \ g ` V) \ V" using gf fim T by auto (metis UnCI image_iff)+ moreover have "U \ V \ f ` (g ` U \ g ` V)" using gf fim by (force simp: image_iff T) ultimately have "U \ V = f ` (g ` U \ g ` V)" by blast moreover have "connected (f ` (g ` U \ g ` V))" proof (rule connected_continuous_image) show "continuous_on (g ` U \ g ` V) f" apply (rule continuous_on_subset [OF contf]) using T fim gfim by blast show "connected (g ` U \ g ` V)" proof (intro conjI unicoherentD [OF S]) show "connected (g ` U)" "connected (g ` V)" using \connected U\ cloU \connected V\ cloV by (metis Topological_Spaces.connected_continuous_image closedin_imp_subset contg continuous_on_subset fim)+ show "S = g ` U \ g ` V" using T fim gfim by auto have hom: "homeomorphism T S g f" by (simp add: contf contg fim gf gfim homeomorphism_def) have "closedin (top_of_set T) U" "closedin (top_of_set T) V" by (simp_all add: cloU cloV) then show "closedin (top_of_set S) (g ` U)" "closedin (top_of_set S) (g ` V)" by (blast intro: homeomorphism_imp_closed_map [OF hom])+ qed qed ultimately show "connected (U \ V)" by metis qed qed lemma homeomorphic_unicoherent_eq: "S homeomorphic T \ (unicoherent S \ unicoherent T)" by (meson homeomorphic_sym homeomorphic_unicoherent) lemma unicoherent_translation: fixes S :: "'a::real_normed_vector set" shows "unicoherent (image (\x. a + x) S) \ unicoherent S" using homeomorphic_translation homeomorphic_unicoherent_eq by blast lemma unicoherent_injective_linear_image: fixes f :: "'a::euclidean_space \ 'b::euclidean_space" assumes "linear f" "inj f" shows "(unicoherent(f ` S) \ unicoherent S)" using assms homeomorphic_unicoherent_eq linear_homeomorphic_image by blast lemma Borsukian_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "Borsukian U" shows "unicoherent U" unfolding unicoherent_def proof clarify fix S T assume "connected S" "connected T" "U = S \ T" and cloS: "closedin (top_of_set (S \ T)) S" and cloT: "closedin (top_of_set (S \ T)) T" show "connected (S \ T)" unfolding connected_closedin_eq proof clarify fix V W assume "closedin (top_of_set (S \ T)) V" and "closedin (top_of_set (S \ T)) W" and VW: "V \ W = S \ T" "V \ W = {}" and "V \ {}" "W \ {}" then have cloV: "closedin (top_of_set U) V" and cloW: "closedin (top_of_set U) W" using \U = S \ T\ cloS cloT closedin_trans by blast+ obtain q where contq: "continuous_on U q" and q01: "\x. x \ U \ q x \ {0..1::real}" and qV: "\x. x \ V \ q x = 0" and qW: "\x. x \ W \ q x = 1" by (rule Urysohn_local [OF cloV cloW \V \ W = {}\, of 0 1]) (fastforce simp: closed_segment_eq_real_ivl) let ?h = "\x. if x \ S then exp(pi * \ * q x) else 1 / exp(pi * \ * q x)" have eqST: "exp(pi * \ * q x) = 1 / exp(pi * \ * q x)" if "x \ S \ T" for x proof - have "x \ V \ W" using that \V \ W = S \ T\ by blast with qV qW show ?thesis by force qed obtain g where contg: "continuous_on U g" and circle: "g ` U \ sphere 0 1" and S: "\x. x \ S \ g x = exp(pi * \ * q x)" and T: "\x. x \ T \ g x = 1 / exp(pi * \ * q x)" proof show "continuous_on U ?h" unfolding \U = S \ T\ proof (rule continuous_on_cases_local [OF cloS cloT]) show "continuous_on S (\x. exp (pi * \ * q x))" apply (intro continuous_intros) using \U = S \ T\ continuous_on_subset contq by blast show "continuous_on T (\x. 1 / exp (pi * \ * q x))" apply (intro continuous_intros) using \U = S \ T\ continuous_on_subset contq by auto qed (use eqST in auto) qed (use eqST in \auto simp: norm_divide\) then obtain h where conth: "continuous_on U h" and heq: "\x. x \ U \ g x = exp (h x)" by (metis Borsukian_continuous_logarithm_circle assms) obtain v w where "v \ V" "w \ W" using \V \ {}\ \W \ {}\ by blast then have vw: "v \ S \ T" "w \ S \ T" using VW by auto have iff: "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 1 \ abs (m - n)" for m n proof - have "2 * pi \ cmod (2 * of_int m * of_real pi * \ - 2 * of_int n * of_real pi * \) \ 2 * pi \ cmod ((2 * pi * \) * (of_int m - of_int n))" by (simp add: algebra_simps) also have "... \ 2 * pi \ 2 * pi * cmod (of_int m - of_int n)" by (simp add: norm_mult) also have "... \ 1 \ abs (m - n)" by simp (metis norm_of_int of_int_1_le_iff of_int_abs of_int_diff) finally show ?thesis . qed have *: "\n::int. h x - (pi * \ * q x) = (of_int(2*n) * pi) * \" if "x \ S" for x using that S \U = S \ T\ heq exp_eq [symmetric] by (simp add: algebra_simps) moreover have "(\x. h x - (pi * \ * q x)) constant_on S" proof (rule continuous_discrete_range_constant [OF \connected S\]) have "continuous_on S h" "continuous_on S q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on S (\x. h x - (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" if "x \ S" "y \ S" and ne: "h y - (pi * \ * q y) \ h x - (pi * \ * q x)" for x y using * [OF \x \ S\] * [OF \y \ S\] ne by (auto simp: iff) then show "\x. x \ S \ \e>0. \y. y \ S \ h y - (pi * \ * q y) \ h x - (pi * \ * q x) \ e \ cmod (h y - (pi * \ * q y) - (h x - (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain m where m: "\x. x \ S \ h x - (pi * \ * q x) = (of_int(2*m) * pi) * \" using vw by (force simp: constant_on_def) have *: "\n::int. h x = - (pi * \ * q x) + (of_int(2*n) * pi) * \" if "x \ T" for x unfolding exp_eq [symmetric] using that T \U = S \ T\ by (simp add: exp_minus field_simps heq [symmetric]) moreover have "(\x. h x + (pi * \ * q x)) constant_on T" proof (rule continuous_discrete_range_constant [OF \connected T\]) have "continuous_on T h" "continuous_on T q" using \U = S \ T\ continuous_on_subset conth contq by blast+ then show "continuous_on T (\x. h x + (pi * \ * q x))" by (intro continuous_intros) have "2*pi \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" if "x \ T" "y \ T" and ne: "h y + (pi * \ * q y) \ h x + (pi * \ * q x)" for x y using * [OF \x \ T\] * [OF \y \ T\] ne by (auto simp: iff) then show "\x. x \ T \ \e>0. \y. y \ T \ h y + (pi * \ * q y) \ h x + (pi * \ * q x) \ e \ cmod (h y + (pi * \ * q y) - (h x + (pi * \ * q x)))" by (rule_tac x="2*pi" in exI) auto qed ultimately obtain n where n: "\x. x \ T \ h x + (pi * \ * q x) = (of_int(2*n) * pi) * \" using vw by (force simp: constant_on_def) show "False" using m [of v] m [of w] n [of v] n [of w] vw by (auto simp: algebra_simps \v \ V\ \w \ W\ qV qW) qed qed corollary contractible_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "contractible U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms contractible_imp_Borsukian) corollary convex_imp_unicoherent: fixes U :: "'a::euclidean_space set" assumes "convex U" shows "unicoherent U" by (simp add: Borsukian_imp_unicoherent assms convex_imp_Borsukian) text\If the type class constraint can be relaxed, I don't know how!\ corollary unicoherent_UNIV: "unicoherent (UNIV :: 'a :: euclidean_space set)" by (simp add: convex_imp_unicoherent) lemma unicoherent_monotone_image_compact: fixes T :: "'b :: t2_space set" assumes S: "unicoherent S" "compact S" and contf: "continuous_on S f" and fim: "f ` S = T" and conn: "\y. y \ T \ connected (S \ f -` {y})" shows "unicoherent T" proof fix U V assume UV: "connected U" "connected V" "T = U \ V" and cloU: "closedin (top_of_set T) U" and cloV: "closedin (top_of_set T) V" moreover have "compact T" using \compact S\ compact_continuous_image contf fim by blast ultimately have "closed U" "closed V" by (auto simp: closedin_closed_eq compact_imp_closed) let ?SUV = "(S \ f -` U) \ (S \ f -` V)" have UV_eq: "f ` ?SUV = U \ V" using \T = U \ V\ fim by force+ have "connected (f ` ?SUV)" proof (rule connected_continuous_image) show "continuous_on ?SUV f" by (meson contf continuous_on_subset inf_le1) show "connected ?SUV" proof (rule unicoherentD [OF \unicoherent S\, of "S \ f -` U" "S \ f -` V"]) have "\C. closedin (top_of_set S) C \ closedin (top_of_set T) (f ` C)" by (metis \compact S\ closed_subset closedin_compact closedin_imp_subset compact_continuous_image compact_imp_closed contf continuous_on_subset fim image_mono) then show "connected (S \ f -` U)" "connected (S \ f -` V)" using UV by (auto simp: conn intro: connected_closed_monotone_preimage [OF contf fim]) show "S = (S \ f -` U) \ (S \ f -` V)" using UV fim by blast show "closedin (top_of_set S) (S \ f -` U)" "closedin (top_of_set S) (S \ f -` V)" by (auto simp: continuous_on_imp_closedin cloU cloV contf fim) qed qed with UV_eq show "connected (U \ V)" by simp qed subsection\Several common variants of unicoherence\ lemma connected_frontier_simple: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected(- S)" shows "connected(frontier S)" unfolding frontier_closures apply (rule unicoherentD [OF unicoherent_UNIV]) apply (simp_all add: assms connected_imp_connected_closure) by (simp add: closure_def) lemma connected_frontier_component_complement: fixes S :: "'a :: euclidean_space set" assumes "connected S" and C: "C \ components(- S)" shows "connected(frontier C)" apply (rule connected_frontier_simple) using C in_components_connected apply blast by (metis Compl_eq_Diff_UNIV connected_UNIV assms top_greatest component_complement_connected) lemma connected_frontier_disjoint: fixes S :: "'a :: euclidean_space set" assumes "connected S" "connected T" "disjnt S T" and ST: "frontier S \ frontier T" shows "connected(frontier S)" proof (cases "S = UNIV") case True then show ?thesis by simp next case False then have "-S \ {}" by blast then obtain C where C: "C \ components(- S)" and "T \ C" by (metis ComplI disjnt_iff subsetI exists_component_superset \disjnt S T\ \connected T\) moreover have "frontier S = frontier C" proof - have "frontier C \ frontier S" using C frontier_complement frontier_of_components_subset by blast moreover have "x \ frontier C" if "x \ frontier S" for x proof - have "x \ closure C" using that unfolding frontier_def by (metis (no_types) Diff_eq ST \T \ C\ closure_mono contra_subsetD frontier_def le_inf_iff that) moreover have "x \ interior C" using that unfolding frontier_def by (metis C Compl_eq_Diff_UNIV Diff_iff subsetD in_components_subset interior_diff interior_mono) ultimately show ?thesis by (auto simp: frontier_def) qed ultimately show ?thesis by blast qed ultimately show ?thesis using \connected S\ connected_frontier_component_complement by auto qed subsection\Some separation results\ lemma separation_by_component_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected_component (- S) a b" obtains C where "C \ components S" "\ connected_component(- C) a b" proof (cases "a \ S \ b \ S") case True then show ?thesis using connected_component_in componentsI that by fastforce next case False obtain T where "T \ S" "closed T" "T \ {}" and nab: "\ connected_component (- T) a b" and conn: "\U. U \ T \ connected_component (- U) a b" using closed_irreducible_separator [OF assms] by metis moreover have "connected T" proof - have ab: "frontier(connected_component_set (- T) a) = T" "frontier(connected_component_set (- T) b) = T" using frontier_minimal_separating_closed_pointwise by (metis False \T \ S\ \closed T\ connected_component_sym conn connected_component_eq_empty connected_component_intermediate_subset empty_subsetI nab)+ have "connected (frontier (connected_component_set (- T) a))" proof (rule connected_frontier_disjoint) show "disjnt (connected_component_set (- T) a) (connected_component_set (- T) b)" unfolding disjnt_iff by (metis connected_component_eq connected_component_eq_empty connected_component_idemp mem_Collect_eq nab) show "frontier (connected_component_set (- T) a) \ frontier (connected_component_set (- T) b)" by (simp add: ab) qed auto with ab \closed T\ show ?thesis by simp qed ultimately obtain C where "C \ components S" "T \ C" using exists_component_superset [of T S] by blast then show ?thesis by (meson Compl_anti_mono connected_component_of_subset nab that) qed lemma separation_by_component_closed: fixes S :: "'a :: euclidean_space set" assumes "closed S" "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain x y where "closed S" "x \ S" "y \ S" and "\ connected_component (- S) x y" using assms by (auto simp: connected_iff_connected_component) then obtain C where "C \ components S" "\ connected_component(- C) x y" using separation_by_component_closed_pointwise by metis then show "thesis" apply (clarify elim!: componentsE) by (metis Compl_iff \C \ components S\ \x \ S\ \y \ S\ connected_component_eq connected_component_eq_eq connected_iff_connected_component that) qed lemma separation_by_Un_closed_pointwise: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected_component (- S) a b" and conT: "connected_component (- T) a b" shows "connected_component (- (S \ T)) a b" proof (rule ccontr) have "a \ S" "b \ S" "a \ T" "b \ T" using conS conT connected_component_in by auto assume "\ connected_component (- (S \ T)) a b" then obtain C where "C \ components (S \ T)" and C: "\ connected_component(- C) a b" using separation_by_component_closed_pointwise assms by blast then have "C \ S \ C \ T" proof - have "connected C" "C \ S \ T" using \C \ components (S \ T)\ in_components_subset by (blast elim: componentsE)+ moreover then have "C \ T = {} \ C \ S = {}" by (metis Int_empty_right ST inf.commute connected_closed) ultimately show ?thesis by blast qed then show False by (meson Compl_anti_mono C conS conT connected_component_of_subset) qed lemma separation_by_Un_closed: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and conS: "connected(- S)" and conT: "connected(- T)" shows "connected(- (S \ T))" using assms separation_by_Un_closed_pointwise by (fastforce simp add: connected_iff_connected_component) lemma open_unicoherent_UNIV: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "connected S" "connected T" "S \ T = UNIV" shows "connected(S \ T)" proof - have "connected(- (-S \ -T))" by (metis closed_Compl compl_sup compl_top_eq double_compl separation_by_Un_closed assms) then show ?thesis by simp qed lemma separation_by_component_open_aux: fixes S :: "'a :: euclidean_space set" assumes ST: "closed S" "closed T" "S \ T = {}" and "S \ {}" "T \ {}" obtains C where "C \ components(-(S \ T))" "C \ {}" "frontier C \ S \ {}" "frontier C \ T \ {}" proof (rule ccontr) let ?S = "S \ \{C \ components(- (S \ T)). frontier C \ S}" let ?T = "T \ \{C \ components(- (S \ T)). frontier C \ T}" assume "\ thesis" with that have *: "frontier C \ S = {} \ frontier C \ T = {}" if C: "C \ components (- (S \ T))" "C \ {}" for C using C by blast have "\A B::'a set. closed A \ closed B \ UNIV \ A \ B \ A \ B = {} \ A \ {} \ B \ {}" proof (intro exI conjI) have "frontier (\{C \ components (- S \ - T). frontier C \ S}) \ S" apply (rule subset_trans [OF frontier_Union_subset_closure]) by (metis (no_types, lifting) SUP_least \closed S\ closure_minimal mem_Collect_eq) then have "frontier ?S \ S" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?S" using frontier_subset_eq by fastforce have "frontier (\{C \ components (- S \ - T). frontier C \ T}) \ T" apply (rule subset_trans [OF frontier_Union_subset_closure]) by (metis (no_types, lifting) SUP_least \closed T\ closure_minimal mem_Collect_eq) then have "frontier ?T \ T" by (simp add: frontier_subset_eq assms subset_trans [OF frontier_Un_subset]) then show "closed ?T" using frontier_subset_eq by fastforce have "UNIV \ (S \ T) \ \(components(- (S \ T)))" using Union_components by blast also have "... \ ?S \ ?T" proof - have "C \ components (-(S \ T)) \ frontier C \ S \ C \ components (-(S \ T)) \ frontier C \ T" if "C \ components (- (S \ T))" "C \ {}" for C using * [OF that] that by clarify (metis (no_types, lifting) UnE \closed S\ \closed T\ closed_Un disjoint_iff_not_equal frontier_of_components_closed_complement subsetCE) then show ?thesis by blast qed finally show "UNIV \ ?S \ ?T" . have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} \ - (S \ T)" using in_components_subset by fastforce moreover have "\{C \ components (- (S \ T)). frontier C \ S} \ \{C \ components (- (S \ T)). frontier C \ T} = {}" proof - have "C \ C' = {}" if "C \ components (- (S \ T))" "frontier C \ S" "C' \ components (- (S \ T))" "frontier C' \ T" for C C' proof - have NUN: "- S \ - T \ UNIV" using \T \ {}\ by blast have "C \ C'" proof assume "C = C'" with that have "frontier C' \ S \ T" by simp also have "... = {}" using \S \ T = {}\ by blast finally have "C' = {} \ C' = UNIV" using frontier_eq_empty by auto then show False using \C = C'\ NUN that by (force simp: dest: in_components_nonempty in_components_subset) qed with that show ?thesis by (simp add: components_nonoverlap [of _ "-(S \ T)"]) qed then show ?thesis by blast qed ultimately show "?S \ ?T = {}" using ST by blast show "?S \ {}" "?T \ {}" using \S \ {}\ \T \ {}\ by blast+ qed then show False by (metis Compl_disjoint connected_UNIV compl_bot_eq compl_unique connected_closedD inf_sup_absorb sup_compl_top_left1 top.extremum_uniqueI) qed proposition separation_by_component_open: fixes S :: "'a :: euclidean_space set" assumes "open S" and non: "\ connected(- S)" obtains C where "C \ components S" "\ connected(- C)" proof - obtain T U where "closed T" "closed U" and TU: "T \ U = - S" "T \ U = {}" "T \ {}" "U \ {}" using assms by (auto simp: connected_closed_set closed_def) then obtain C where C: "C \ components(-(T \ U))" "C \ {}" and "frontier C \ T \ {}" "frontier C \ U \ {}" using separation_by_component_open_aux [OF \closed T\ \closed U\ \T \ U = {}\] by force show "thesis" proof show "C \ components S" using C(1) TU(1) by auto show "\ connected (- C)" proof assume "connected (- C)" then have "connected (frontier C)" using connected_frontier_simple [of C] \C \ components S\ in_components_connected by blast then show False unfolding connected_closed by (metis C(1) TU(2) \closed T\ \closed U\ \frontier C \ T \ {}\ \frontier C \ U \ {}\ closed_Un frontier_of_components_closed_complement inf_bot_right inf_commute) qed qed qed lemma separation_by_Un_open: fixes S :: "'a :: euclidean_space set" assumes "open S" "open T" "S \ T = {}" and cS: "connected(-S)" and cT: "connected(-T)" shows "connected(- (S \ T))" using assms unicoherent_UNIV unfolding unicoherent_def by force lemma nonseparation_by_component_eq: fixes S :: "'a :: euclidean_space set" assumes "open S \ closed S" shows "((\C \ components S. connected(-C)) \ connected(- S))" (is "?lhs = ?rhs") proof assume ?lhs with assms show ?rhs by (meson separation_by_component_closed separation_by_component_open) next assume ?rhs with assms show ?lhs using component_complement_connected by force qed text\Another interesting equivalent of an inessential mapping into C-{0}\ proposition inessential_eq_extensible: fixes f :: "'a::euclidean_space \ complex" assumes "closed S" shows "(\a. homotopic_with_canon (\h. True) S (-{0}) f (\t. a)) \ (\g. continuous_on UNIV g \ (\x \ S. g x = f x) \ (\x. g x \ 0))" (is "?lhs = ?rhs") proof assume ?lhs then obtain a where a: "homotopic_with_canon (\h. True) S (-{0}) f (\t. a)" .. show ?rhs proof (cases "S = {}") case True with a show ?thesis using continuous_on_const by force next case False have anr: "ANR (-{0::complex})" by (simp add: ANR_delete open_Compl open_imp_ANR) obtain g where contg: "continuous_on UNIV g" and gim: "g ` UNIV \ -{0}" and gf: "\x. x \ S \ g x = f x" proof (rule Borsuk_homotopy_extension_homotopic [OF _ _ continuous_on_const _ homotopic_with_symD [OF a]]) show "closedin (top_of_set UNIV) S" using assms by auto show "range (\t. a) \ - {0}" using a homotopic_with_imp_subset2 False by blast qed (use anr that in \force+\) then show ?thesis by force qed next assume ?rhs then obtain g where contg: "continuous_on UNIV g" and gf: "\x. x \ S \ g x = f x" and non0: "\x. g x \ 0" by metis obtain h k::"'a\'a" where hk: "homeomorphism (ball 0 1) UNIV h k" using homeomorphic_ball01_UNIV homeomorphic_def by blast then have "continuous_on (ball 0 1) (g \ h)" by (meson contg continuous_on_compose continuous_on_subset homeomorphism_cont1 top_greatest) then obtain j where contj: "continuous_on (ball 0 1) j" and j: "\z. z \ ball 0 1 \ exp(j z) = (g \ h) z" by (metis (mono_tags, hide_lams) continuous_logarithm_on_ball comp_apply non0) have [simp]: "\x. x \ S \ h (k x) = x" using hk homeomorphism_apply2 by blast have "\\. continuous_on S \\ (\x\S. f x = exp (\ x))" proof (intro exI conjI ballI) show "continuous_on S (j \ k)" proof (rule continuous_on_compose) show "continuous_on S k" by (meson continuous_on_subset hk homeomorphism_cont2 top_greatest) show "continuous_on (k ` S) j" apply (rule continuous_on_subset [OF contj]) using homeomorphism_image2 [OF hk] continuous_on_subset [OF contj] by blast qed show "f x = exp ((j \ k) x)" if "x \ S" for x proof - have "f x = (g \ h) (k x)" by (simp add: gf that) also have "... = exp (j (k x))" by (metis rangeI homeomorphism_image2 [OF hk] j) finally show ?thesis by simp qed qed then show ?lhs by (simp add: inessential_eq_continuous_logarithm) qed lemma inessential_on_clopen_Union: fixes \ :: "'a::euclidean_space set set" assumes T: "path_connected T" and "\S. S \ \ \ closedin (top_of_set (\\)) S" and "\S. S \ \ \ openin (top_of_set (\\)) S" and hom: "\S. S \ \ \ \a. homotopic_with_canon (\x. True) S T f (\x. a)" obtains a where "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (cases "\\ = {}") case True with that show ?thesis by force next case False then obtain C where "C \ \" "C \ {}" by blast then obtain a where clo: "closedin (top_of_set (\\)) C" and ope: "openin (top_of_set (\\)) C" and "homotopic_with_canon (\x. True) C T f (\x. a)" using assms by blast with \C \ {}\ have "f ` C \ T" "a \ T" using homotopic_with_imp_subset1 homotopic_with_imp_subset2 by blast+ have "homotopic_with_canon (\x. True) (\\) T f (\x. a)" proof (rule homotopic_on_clopen_Union) show "\S. S \ \ \ closedin (top_of_set (\\)) S" "\S. S \ \ \ openin (top_of_set (\\)) S" by (simp_all add: assms) show "homotopic_with_canon (\x. True) S T f (\x. a)" if "S \ \" for S proof (cases "S = {}") case True then show ?thesis by auto next case False then obtain b where "b \ S" by blast obtain c where c: "homotopic_with_canon (\x. True) S T f (\x. c)" using \S \ \\ hom by blast then have "c \ T" using \b \ S\ homotopic_with_imp_subset2 by blast then have "homotopic_with_canon (\x. True) S T (\x. a) (\x. c)" using T \a \ T\ homotopic_constant_maps path_connected_component by (simp add: homotopic_constant_maps path_connected_component) then show ?thesis using c homotopic_with_symD homotopic_with_trans by blast qed qed then show ?thesis .. qed proposition Janiszewski_dual: fixes S :: "complex set" assumes "compact S" "compact T" "connected S" "connected T" "connected(- (S \ T))" shows "connected(S \ T)" proof - have ST: "compact (S \ T)" by (simp add: assms compact_Un) with Borsukian_imp_unicoherent [of "S \ T"] ST assms show ?thesis by (auto simp: closed_subset compact_imp_closed Borsukian_separation_compact unicoherent_def) qed end