# HG changeset patch # User paulson # Date 1456408728 0 # Node ID 86f27b264d3da58b3bda4aeee579b1e042d3a002 # Parent 8383b126b0a9e868629d2e85da6b7364aa299dad Conformal_mappings: a big development in complex analysis (+ some lemmas) diff -r 8383b126b0a9 -r 86f27b264d3d NEWS --- a/NEWS Thu Feb 25 00:36:44 2016 +0100 +++ b/NEWS Thu Feb 25 13:58:48 2016 +0000 @@ -38,6 +38,9 @@ * Compound constants INFIMUM and SUPREMUM are mere abbreviations now. INCOMPATIBILITY. +* More complex analysis including Cauchy's inequality, Liouville theorem, +open mapping theorem, maximum modulus principle, Schwarz Lemma. + * "Gcd (f ` A)" and "Lcm (f ` A)" are printed with optional comprehension-like syntax analogously to "Inf (f ` A)" and "Sup (f ` A)". diff -r 8383b126b0a9 -r 86f27b264d3d src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy --- a/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Thu Feb 25 00:36:44 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm.thy Thu Feb 25 13:58:48 2016 +0000 @@ -197,6 +197,24 @@ subsubsection\The concept of continuously differentiable\ +text \ +John Harrison writes as follows: + +``The usual assumption in complex analysis texts is that a path \ should be piecewise +continuously differentiable, which ensures that the path integral exists at least for any continuous +f, since all piecewise continuous functions are integrable. However, our notion of validity is +weaker, just piecewise differentiability... [namely] continuity plus differentiability except on a +finite set ... [Our] underlying theory of integration is the Kurzweil-Henstock theory. In contrast to +the Riemann or Lebesgue theory (but in common with a simple notion based on antiderivatives), this +can integrate all derivatives.'' + +"Formalizing basic complex analysis." From Insight to Proof: Festschrift in Honour of Andrzej Trybulec. +Studies in Logic, Grammar and Rhetoric 10.23 (2007): 151-165. + +And indeed he does not assume that his derivatives are continuous, but the penalty is unreasonably +difficult proofs concerning winding numbers. We need a self-contained and straightforward theorem +asserting that all derivatives can be integrated before we can adopt Harrison's choice.\ + definition C1_differentiable_on :: "(real \ 'a::real_normed_vector) \ real set \ bool" (infix "C1'_differentiable'_on" 50) where @@ -549,6 +567,63 @@ lemma closed_valid_path_image: "valid_path g \ closed(path_image g)" by (metis closed_path_image valid_path_imp_path) +lemma valid_path_compose: + assumes "valid_path g" + and der:"\x. x \ path_image g \ \f'. (f has_field_derivative f') (at x)" + and con: "continuous_on {0..1} (\x. deriv f (g x))" + shows "valid_path (f o g)" +proof - + obtain s where "finite s" and g_diff: "g C1_differentiable_on {0..1} - s" + using `valid_path g` unfolding valid_path_def piecewise_C1_differentiable_on_def by auto + have "f \ g differentiable at t" if "t \ {0..1} - s" for t + proof (rule differentiable_chain_at) + show "g differentiable at t" using `valid_path g` + by (meson C1_differentiable_on_eq \g C1_differentiable_on {0..1} - s\ that) + next + have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis + then obtain f' where "(f has_field_derivative f') (at (g t))" + using der by auto + then have " (f has_derivative op * f') (at (g t))" + using has_field_derivative_imp_has_derivative[of f f' "at (g t)"] by auto + then show "f differentiable at (g t)" using differentiableI by auto + qed + moreover have "continuous_on ({0..1} - s) (\x. vector_derivative (f \ g) (at x))" + proof (rule continuous_on_eq [where f = "\x. vector_derivative g (at x) * deriv f (g x)"], + rule continuous_intros) + show "continuous_on ({0..1} - s) (\x. vector_derivative g (at x))" + using g_diff C1_differentiable_on_eq by auto + next + show "continuous_on ({0..1} - s) (\x. deriv f (g x))" + using con continuous_on_subset by blast + next + show "vector_derivative g (at t) * deriv f (g t) = vector_derivative (f \ g) (at t)" + if "t \ {0..1} - s" for t + proof (rule vector_derivative_chain_at_general[symmetric]) + show "g differentiable at t" by (meson C1_differentiable_on_eq g_diff that) + next + have "g t\path_image g" using that DiffD1 image_eqI path_image_def by metis + then obtain f' where "(f has_field_derivative f') (at (g t))" + using der by auto + then show "\g'. (f has_field_derivative g') (at (g t))" by auto + qed + qed + ultimately have "f o g C1_differentiable_on {0..1} - s" + using C1_differentiable_on_eq by blast + moreover have "path (f o g)" + proof - + have "isCont f x" if "x \ path_image g" for x + proof - + obtain f' where "(f has_field_derivative f') (at x)" + using der `x\path_image g` by auto + thus ?thesis using DERIV_isCont by auto + qed + then have "continuous_on (path_image g) f" using continuous_at_imp_continuous_on by auto + then show ?thesis using path_continuous_image `valid_path g` valid_path_imp_path by auto + qed + ultimately show ?thesis unfolding valid_path_def piecewise_C1_differentiable_on_def path_def + using `finite s` by auto +qed + subsection\Contour Integrals along a path\ diff -r 8383b126b0a9 -r 86f27b264d3d src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy Thu Feb 25 00:36:44 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Complex_Analysis_Basics.thy Thu Feb 25 13:58:48 2016 +0000 @@ -393,7 +393,7 @@ unfolding complex_differentiable_def by (metis at_within_open) -subsection\Caratheodory characterization.\ +subsection\Caratheodory characterization\ lemma complex_differentiable_caratheodory_at: "f complex_differentiable (at z) \ @@ -750,7 +750,6 @@ by (metis e e' min_less_iff_conj) qed - lemma analytic_on_divide: assumes f: "f analytic_on s" and g: "g analytic_on s" @@ -767,7 +766,30 @@ "(\i. i \ I \ (f i) analytic_on s) \ (\x. setsum (\i. f i x) I) analytic_on s" by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add) -subsection\analyticity at a point.\ +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" + apply (rule complex_derivative_transform_within_open [where s=S]) + apply (rule assms holomorphic_on_compose_gen holomorphic_intros)+ + apply simp + done + also have "... = 1" + by simp + finally show ?thesis . +qed + +subsection\analyticity at a point\ lemma analytic_at_ball: "f analytic_on {z} \ (\e. 0 f holomorphic_on ball z e)" @@ -972,7 +994,7 @@ apply fact done -subsection\Inverse function theorem for complex derivatives.\ +subsection\Inverse function theorem for complex derivatives\ lemma has_complex_derivative_inverse_basic: fixes f :: "complex \ complex" @@ -1118,7 +1140,7 @@ finally show ?thesis . qed -text\Something more like the traditional MVT for real components.\ +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)" diff -r 8383b126b0a9 -r 86f27b264d3d src/HOL/Multivariate_Analysis/Conformal_Mappings.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Multivariate_Analysis/Conformal_Mappings.thy Thu Feb 25 13:58:48 2016 +0000 @@ -0,0 +1,1728 @@ +section \Conformal Mappings. Consequences of Cauchy's integral theorem.\ + +text\By John Harrison et al. Ported from HOL Light by L C Paulson (2016)\ + +theory Conformal_Mappings +imports "~~/src/HOL/Multivariate_Analysis/Cauchy_Integral_Thm" + +begin + +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 "0 < r" and "0 < n" + and fin : "\w. w \ ball z r \ f w \ ball y B0" + 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) + def 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 divide_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 * ii)/(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 + +proposition 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 * ii)/(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 + +proposition 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 + def 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: divide_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 add:) + 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: divide_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.\ + +proposition 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 + def b \ "\i. a (i+m) / a m" + def g \ "\x. suminf (\i. b i * (x - \) ^ i)" + 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 + +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" "ball \ r \ S" "\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 S: "open S" "connected S" + and "U \ S" "\ \ 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 + def 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 + + +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 + def g \ "\z. inverse (f 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 add:) + 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 add:) + 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" "connected S" + and "open U" "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 def \ \ "norm (f w - f \) / 3" + ultimately have "0 < \" + using \0 < r\ dist_complex_def r sne by auto + have "ball (f \) \ \ f ` U" + proof + fix \ + assume \: "\ \ ball (f \) \" + have *: "cmod (\ - f \) < cmod (\ - f z)" if "cmod (\ - z) = r" for z + proof - + have lt: "cmod (f w - f \) / 3 < cmod (\ - f z)" + using w [OF that] \ + using dist_triangle2 [of "f \" "\" "f z"] dist_triangle2 [of "f \" "f z" \] + by (simp add: \_def dist_norm norm_minus_commute) + show ?thesis + by (metis \_def dist_commute dist_norm less_trans lt mem_ball \) + qed + have "continuous_on (cball \ r) (\z. \ - f z)" + apply (rule continuous_intros)+ + using \cball \ r \ U\ \f holomorphic_on U\ + apply (blast intro: continuous_on_subset holomorphic_on_imp_continuous_on) + done + moreover have "(\z. \ - f z) holomorphic_on ball \ r" + apply (rule holomorphic_intros)+ + apply (metis \cball \ r \ U\ \f holomorphic_on U\ holomorphic_on_subset interior_cball interior_subset) + done + ultimately obtain z where "z \ ball \ r" "\ - f z = 0" + apply (rule holomorphic_contract_to_zero) + apply (blast intro!: \0 < r\ *)+ + done + then show "\ \ f ` U" + using \cball \ r \ U\ by fastforce + qed + then show ?thesis using \0 < \\ by blast + qed + qed + have "open (f ` X)" if "X \ components U" for X + proof - + have holfU: "f holomorphic_on U" + using \U \ S\ holf holomorphic_on_subset by blast + have "X \ {}" + using that by (simp add: in_components_nonempty) + moreover have "open X" + using that \open U\ open_components by auto + moreover have "connected X" + using that in_components_maximal by blast + moreover have "f holomorphic_on X" + by (meson that holfU holomorphic_on_subset in_components_maximal) + moreover have "\y\X. f y \ x" for x + proof (rule ccontr) + assume not: "\ (\y\X. f y \ x)" + have "X \ S" + using \U \ S\ in_components_subset that by blast + obtain w where w: "w \ X" using \X \ {}\ by blast + have wis: "w islimpt X" + using w \open X\ interior_eq by auto + have hol: "(\z. f z - x) holomorphic_on S" + by (simp add: holf holomorphic_on_diff) + with fne [unfolded constant_on_def] analytic_continuation [OF hol S \X \ S\ _ wis] + not \X \ S\ w + show False by auto + qed + ultimately show ?thesis + by (rule *) + qed + then show ?thesis + by (subst Union_components [of U]) (auto simp: image_Union) +qed + + +text\No need for @{term S} to be connected. But the nonconstant condition is stronger.\ +corollary 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 (fact Union_components) + with \U \ S\ have "U = (\C \ components S. C \ U)" by auto + then have "f ` U = (\C \ components S. f ` (C \ U))" + by auto + 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 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" "connected S" + and "open U" "U \ S" "\ \ 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 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\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 + def g \ "\w. suminf (\i. (deriv ^^ (i + n)) f \ / (fact(i + n)) * (w - \)^i)" + 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 - + def 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.setsum.setdiff_irrelevant [symmetric]) + apply (simp add:) + 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 add:) + 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) complex_differentiable at x" + apply (rule derivative_intros)+ + using holg mem_ball apply (blast intro: holomorphic_deriv holomorphic_on_imp_differentiable_at) + apply (metis Topology_Euclidean_Space.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_complex_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 add:) + 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: "\c. \z \ S. f z \ c" + 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 auto +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 + def 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 - + def 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 + def 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 - + def h \ "\z. (z - \)^2 * f z" + have h0: "(h has_field_derivative 0) (at \)" + apply (simp add: h_def Derivative.DERIV_within_iff) + apply (rule Lim_transform_within [OF that, of 1]) + apply (auto simp: divide_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 complex_differentiable at z within S" + proof (cases "z = \") + case True then show ?thesis + using complex_differentiable_at_within complex_differentiable_def h0 by blast + next + case False + then have "f complex_differentiable at z within S" + using holomorphic_onD [OF holf, of z] \z \ S\ + unfolding complex_differentiable_def DERIV_within_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 + def g \ "\z. if z = \ then deriv h \ else (h z - h \) / (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 + + +proposition 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 add:) + 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 divide_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: divide_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]) + apply (simp add:) + done + next + case False + then have fi0: "\r. r > 0 \ \z\ball 0 r - {0}. f (inverse z) = 0" + by simp + have fz0: "f z = 0" if "0 < r" and lt1: "\x. x \ 0 \ cmod x < r \ inverse (cmod (f (inverse x))) < 1" + for z r + proof - + have f0: "(f \ 0) at_infinity" + proof - + have DIM_complex[intro]: "2 \ DIM(complex)" --\should not be necessary!\ + by simp + have "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 "\w \ 0; cmod w < r\ \ f (inverse w) = 0" for w + apply (rule disjE [OF connected_closedD [where A = "{0}" and B = "- ball 0 1"]], auto) + apply (metis (mono_tags, hide_lams) not_less_iff_gr_or_eq one_less_inverse lt1 zero_less_norm_iff) + using False \0 < r\ apply fastforce + by (metis (no_types, hide_lams) Compl_iff IntI comp_apply empty_iff image_eqI insert_Diff_single insert_iff mem_ball_0 not_less_iff_gr_or_eq one_less_inverse that(2) zero_less_norm_iff) + then show ?thesis + apply (simp add: lim_at_infinity_0) + apply (rule Lim_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\Entire proper functions are precisely the non-trivial polynomials\ + +proposition 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 (rule allI continuous_closed_preimage_univ continuous_intros)+ + using \compact K\ compact_eq_bounded_closed by blast + ultimately show ?thesis + using closed_inter_compact [OF \closed S\] compact_eq_bounded_closed by blast +qed + +corollary 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 + + +proposition 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 setsum_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 + def m \ "GREATEST k. k\n \ a k \ 0" + have m: "m\n \ a m \ 0" + unfolding m_def + apply (rule GreatestI [where b = "Suc n"]) + using k apply auto + done + have [simp]: "a i = 0" if "m < i" "i \ n" for i + using Greatest_le [where b = "Suc 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 [where b = "Suc 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.setsum.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 add:) + 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\ + +proposition 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 add:) + 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 (op * (deriv f \))" + using dnz by simp + then obtain g' where g': "linear g'" "g' \ op * (deriv f \) = id" + using linear_injective_left_inverse [of "op * (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_complex_differentiable has_field_derivative_imp_has_derivative holf + holomorphic_on_imp_differentiable_at \open S\ apply blast + done +qed + + +proposition 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 + + +proposition 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 + 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_complex_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 Topology_Euclidean_Space.open_ball centre_in_ball) + def 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 * ii * (0/n)) \ cball 0 \" + "\ * exp(2 * of_real pi * ii * (1/n)) \ cball 0 \" + by (auto simp: norm_mult) + with \ have "\ * exp(2 * of_real pi * ii * (0/n)) \ U" + "\ * exp(2 * of_real pi * ii * (1/n)) \ U" by blast+ + then obtain y0 y1 where "y0 \ T" and y0: "(y0 - \) * g y0 = \ * exp(2 * of_real pi * ii * (0/n))" + and "y1 \ T" and y1: "(y1 - \) * g y1 = \ * exp(2 * of_real pi * ii * (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] + 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\ + +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_complex_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_complex_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 - + def h \ "\z. if z = 0 then deriv f 0 else f z / z" + have d0: "deriv f 0 = h 0" + by (simp add: h_def) + moreover have "h holomorphic_on (ball 0 r)" + by (rule pole_theorem_open_0 [OF holf, of 0]) (auto simp: h_def) + moreover have "norm z < r \ f z = z * h z" for z + by (simp add: h_def) + ultimately show ?thesis + using that by blast +qed + + +proposition Schwarz_Lemma: + assumes holf: "f holomorphic_on (ball 0 1)" and [simp]: "f 0 = 0" + and no: "\z. norm z < 1 \ norm (f z) < 1" + and \: "norm \ < 1" + shows "norm (f \) \ norm \" and "norm(deriv f 0) \ 1" + and "((\z. norm z < 1 \ z \ 0 \ norm(f z) = norm z) \ norm(deriv f 0) = 1) + \ \\. (\z. norm z < 1 \ f z = \ * z) \ norm \ = 1" (is "?P \ ?Q") +proof - + obtain h where holh: "h holomorphic_on (ball 0 1)" + and fz_eq: "\z. norm z < 1 \ f z = z * (h z)" and df0: "deriv f 0 = h 0" + by (rule Schwarz3 [OF holf]) auto + have noh_le: "norm (h z) \ 1" if z: "norm z < 1" for z + proof - + have "norm (h z) < a" if a: "1 < a" for a + proof - + have "max (inverse a) (norm z) < 1" + using z a by (simp_all add: inverse_less_1_iff) + then obtain r where r: "max (inverse a) (norm z) < r" and "r < 1" + using Rats_dense_in_real by blast + then have nzr: "norm z < r" and ira: "inverse r < a" + using z a less_imp_inverse_less by force+ + then have "0 < r" + by (meson norm_not_less_zero not_le order.strict_trans2) + have holh': "h holomorphic_on ball 0 r" + by (meson holh \r < 1\ holomorphic_on_subset less_eq_real_def subset_ball) + have conth': "continuous_on (cball 0 r) h" + by (meson \r < 1\ dual_order.trans holh holomorphic_on_imp_continuous_on holomorphic_on_subset mem_ball_0 mem_cball_0 not_less subsetI) + obtain w where w: "norm w = r" and lenw: "\z. norm z < r \ norm(h z) \ norm(h w)" + apply (rule Schwarz1 [OF holh']) using conth' \0 < r\ by auto + have "h w = f w / w" using fz_eq \r < 1\ nzr w by auto + then have "cmod (h z) < inverse r" + by (metis \0 < r\ \r < 1\ divide_strict_right_mono inverse_eq_divide + le_less_trans lenw no norm_divide nzr w) + then show ?thesis using ira by linarith + qed + then show "norm (h z) \ 1" + using not_le by blast + qed + show "cmod (f \) \ cmod \" + proof (cases "\ = 0") + case True then show ?thesis by auto + next + case False + then show ?thesis + by (simp add: noh_le fz_eq \ mult_left_le norm_mult) + qed + show no_df0: "norm(deriv f 0) \ 1" + by (simp add: \\z. cmod z < 1 \ cmod (h z) \ 1\ df0) + show "?Q" if "?P" + using that + proof + assume "\z. cmod z < 1 \ z \ 0 \ cmod (f z) = cmod z" + then obtain \ where \: "cmod \ < 1" "\ \ 0" "cmod (f \) = cmod \" by blast + then have [simp]: "norm (h \) = 1" + by (simp add: fz_eq norm_mult) + have "ball \ (1 - cmod \) \ ball 0 1" + by (simp add: ball_subset_ball_iff) + moreover have "\z. cmod (\ - z) < 1 - cmod \ \ cmod (h z) \ cmod (h \)" + apply (simp add: algebra_simps) + by (metis add_diff_cancel_left' diff_diff_eq2 le_less_trans noh_le norm_triangle_ineq4) + ultimately obtain c where c: "\z. norm z < 1 \ h z = c" + using Schwarz2 [OF holh, of "1 - norm \" \, unfolded constant_on_def] \ by auto + moreover then have "norm c = 1" + using \ by force + ultimately 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 + +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 "f contour_integrable_on linepath b' a'" + by (meson a'b' contf continuous_on_subset contour_integrable_continuous_linepath + convex_contains_segment \convex S\) + then 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 complex_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 complex_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 complex_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 complex_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 + +proposition 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 complex_differentiable at x within S \ {z. Im z < 0}" + if "x \ S" "Im x < 0" "f complex_differentiable at (cnj x) within S \ {z. 0 < Im z}" for x + using that + apply (simp add: complex_differentiable_def Derivative.DERIV_within_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 "-ii" _ 0]) + using 1 2 3 + apply auto + done +qed + +end diff -r 8383b126b0a9 -r 86f27b264d3d src/HOL/Multivariate_Analysis/Derivative.thy --- a/src/HOL/Multivariate_Analysis/Derivative.thy Thu Feb 25 00:36:44 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Derivative.thy Thu Feb 25 13:58:48 2016 +0000 @@ -2506,6 +2506,22 @@ vector_derivative f (at x) *\<^sub>R vector_derivative g (at (f x))" by (metis vector_diff_chain_at vector_derivative_at vector_derivative_works assms) +lemma field_vector_diff_chain_at: (*thanks to Wenda Li*) + assumes Df: "(f has_vector_derivative f') (at x)" + and Dg: "(g has_field_derivative g') (at (f x))" + shows "((g \ f) has_vector_derivative (f' * g')) (at x)" +using diff_chain_at[OF Df[unfolded has_vector_derivative_def] + Dg [unfolded has_field_derivative_def]] + by (auto simp: o_def mult.commute has_vector_derivative_def) + +lemma vector_derivative_chain_at_general: (*thanks to Wenda Li*) + assumes "f differentiable at x" "(\g'. (g has_field_derivative g') (at (f x)))" + shows "vector_derivative (g \ f) (at x) = + vector_derivative f (at x) * deriv g (f x)" +apply (rule vector_derivative_at [OF field_vector_diff_chain_at]) +using assms +by (auto simp: vector_derivative_works DERIV_deriv_iff_has_field_derivative) + subsection \Relation between convexity and derivative\ diff -r 8383b126b0a9 -r 86f27b264d3d src/HOL/Multivariate_Analysis/Determinants.thy --- a/src/HOL/Multivariate_Analysis/Determinants.thy Thu Feb 25 00:36:44 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Determinants.thy Thu Feb 25 13:58:48 2016 +0000 @@ -577,7 +577,7 @@ qed -lemma eq_id_iff[simp]: "(\x. f x = x) \ f = id" +lemma eq_id_iff: "(\x. f x = x) \ f = id" by auto lemma det_linear_rows_setsum_lemma: @@ -592,7 +592,7 @@ have th0: "\x y. (\ i. if i \ {} then x i else y i) = (\ i. y i)" by vector from empty.prems show ?case - unfolding th0 by simp + unfolding th0 by (simp add: eq_id_iff) next case (insert z T a c) let ?F = "\T. {f. (\i \ T. f i \ S) \ (\i. i \ T \ f i = i)}" diff -r 8383b126b0a9 -r 86f27b264d3d src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy --- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Thu Feb 25 00:36:44 2016 +0100 +++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy Thu Feb 25 13:58:48 2016 +0000 @@ -4,10 +4,9 @@ Extended_Real_Limits Determinants Ordered_Euclidean_Space - Complex_Analysis_Basics Bounded_Continuous_Function Weierstrass - Cauchy_Integral_Thm + Conformal_Mappings Generalised_Binomial_Theorem Gamma begin