# HG changeset patch # User paulson # Date 1675960193 0 # Node ID 268c8184288349b4e755c9d78575051edd57ecc3 # Parent b6f3eb537d912c913dfc3eac7c927d14434e530a# Parent 8c093a4b8ccf9a9628cfd771504963c1944e1c7c merged diff -r b6f3eb537d91 -r 268c81842883 src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Thu Feb 09 08:35:50 2023 +0000 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Thu Feb 09 16:29:53 2023 +0000 @@ -417,14 +417,24 @@ (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) + by (meson analytic_imp_holomorphic analytic_on_def holomorphic_on_subset openE) 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) + using analytic_on_def holomorphic_on_imp_differentiable_at by auto + +lemma analytic_at_imp_isCont: + assumes "f analytic_on {z}" + shows "isCont f z" + using assms by (meson analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at insertI1) + +lemma analytic_at_neq_imp_eventually_neq: + assumes "f analytic_on {x}" "f x \ c" + shows "eventually (\y. f y \ c) (at x)" +proof (intro tendsto_imp_eventually_ne) + show "f \x\ f x" + using assms by (simp add: analytic_at_imp_isCont isContD) +qed (use assms in auto) lemma analytic_on_subset: "f analytic_on S \ T \ S \ f analytic_on T" by (auto simp: analytic_on_def) @@ -652,15 +662,20 @@ lemma analytic_at_ball: "f analytic_on {z} \ (\e. 0 f holomorphic_on ball z e)" -by (metis analytic_on_def singleton_iff) + 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) + "f analytic_on {z} \ (\s. open s \ z \ s \ f holomorphic_on s)" + by (metis analytic_on_holomorphic empty_subsetI insert_subset) + +lemma holomorphic_on_imp_analytic_at: + assumes "f holomorphic_on A" "open A" "z \ A" + shows "f analytic_on {z}" + using assms by (meson analytic_at) lemma analytic_on_analytic_at: - "f analytic_on s \ (\z \ s. f analytic_on {z})" -by (metis analytic_at_ball analytic_on_def) + "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} \ diff -r b6f3eb537d91 -r 268c81842883 src/HOL/Analysis/Isolated.thy --- a/src/HOL/Analysis/Isolated.thy Thu Feb 09 08:35:50 2023 +0000 +++ b/src/HOL/Analysis/Isolated.thy Thu Feb 09 16:29:53 2023 +0000 @@ -14,6 +14,12 @@ definition (in metric_space) uniform_discrete :: "'a set \ bool" where "uniform_discrete S \ (\e>0. \x\S. \y\S. dist x y < e \ x = y)" +lemma discreteI: "(\x. x \ X \ x isolated_in X ) \ discrete X" + unfolding discrete_def by auto + +lemma discreteD: "discrete X \ x \ X \ x isolated_in X " + unfolding discrete_def by auto + lemma uniformI1: assumes "e>0" "\x y. \x\S;y\S;dist x y \ x =y " shows "uniform_discrete S" @@ -43,6 +49,54 @@ shows "x isolated_in (insert a S) \ x isolated_in S \ (x=a \ \ (x islimpt S))" by (meson insert_iff islimpt_insert isolated_in_islimpt_iff) +lemma isolated_inI: + assumes "x\S" "open T" "T \ S = {x}" + shows "x isolated_in S" + using assms unfolding isolated_in_def by auto + +lemma isolated_inE: + assumes "x isolated_in S" + obtains T where "x \ S" "open T" "T \ S = {x}" + using assms that unfolding isolated_in_def by force + +lemma isolated_inE_dist: + assumes "x isolated_in S" + obtains d where "d > 0" "\y. y \ S \ dist x y < d \ y = x" + by (meson assms isolated_in_dist_Ex_iff) + +lemma isolated_in_altdef: + "x isolated_in S \ (x\S \ eventually (\y. y \ S) (at x))" +proof + assume "x isolated_in S" + from isolated_inE[OF this] + obtain T where "x \ S" and T:"open T" "T \ S = {x}" + by metis + have "\\<^sub>F y in nhds x. y \ T" + apply (rule eventually_nhds_in_open) + using T by auto + then have "eventually (\y. y \ T - {x}) (at x)" + unfolding eventually_at_filter by eventually_elim auto + then have "eventually (\y. y \ S) (at x)" + by eventually_elim (use T in auto) + then show " x \ S \ (\\<^sub>F y in at x. y \ S)" using \x \ S\ by auto +next + assume "x \ S \ (\\<^sub>F y in at x. y \ S)" + then have "\\<^sub>F y in at x. y \ S" "x\S" by auto + from this(1) have "eventually (\y. y \ S \ y = x) (nhds x)" + unfolding eventually_at_filter by eventually_elim auto + then obtain T where T:"open T" "x \ T" "(\y\T. y \ S \ y = x)" + unfolding eventually_nhds by auto + with \x \ S\ have "T \ S = {x}" + by fastforce + with \x\S\ \open T\ + show "x isolated_in S" + unfolding isolated_in_def by auto +qed + +lemma discrete_altdef: + "discrete S \ (\x\S. \\<^sub>F y in at x. y \ S)" + unfolding discrete_def isolated_in_altdef by auto + (* TODO. Other than @@ -194,4 +248,80 @@ ultimately show ?thesis by fastforce qed +definition sparse :: "real \ 'a :: metric_space set \ bool" + where "sparse \ X \ (\x\X. \y\X-{x}. dist x y > \)" + +lemma sparse_empty [simp, intro]: "sparse \ {}" + by (auto simp: sparse_def) + +lemma sparseI [intro?]: + "(\x y. x \ X \ y \ X \ x \ y \ dist x y > \) \ sparse \ X" + unfolding sparse_def by auto + +lemma sparseD: + "sparse \ X \ x \ X \ y \ X \ x \ y \ dist x y > \" + unfolding sparse_def by auto + +lemma sparseD': + "sparse \ X \ x \ X \ y \ X \ dist x y \ \ \ x = y" + unfolding sparse_def by force + +lemma sparse_singleton [simp, intro]: "sparse \ {x}" + by (auto simp: sparse_def) + +definition setdist_gt where "setdist_gt \ X Y \ (\x\X. \y\Y. dist x y > \)" + +lemma setdist_gt_empty [simp]: "setdist_gt \ {} Y" "setdist_gt \ X {}" + by (auto simp: setdist_gt_def) + +lemma setdist_gtI: "(\x y. x \ X \ y \ Y \ dist x y > \) \ setdist_gt \ X Y" + unfolding setdist_gt_def by auto + +lemma setdist_gtD: "setdist_gt \ X Y \ x \ X \ y \ Y \ dist x y > \" + unfolding setdist_gt_def by auto + +lemma setdist_gt_setdist: "\ < setdist A B \ setdist_gt \ A B" + unfolding setdist_gt_def using setdist_le_dist by fastforce + +lemma setdist_gt_mono: "setdist_gt \' A B \ \ \ \' \ A' \ A \ B' \ B \ setdist_gt \ A' B'" + by (force simp: setdist_gt_def) + +lemma setdist_gt_Un_left: "setdist_gt \ (A \ B) C \ setdist_gt \ A C \ setdist_gt \ B C" + by (auto simp: setdist_gt_def) + +lemma setdist_gt_Un_right: "setdist_gt \ C (A \ B) \ setdist_gt \ C A \ setdist_gt \ C B" + by (auto simp: setdist_gt_def) + +lemma compact_closed_imp_eventually_setdist_gt_at_right_0: + assumes "compact A" "closed B" "A \ B = {}" + shows "eventually (\\. setdist_gt \ A B) (at_right 0)" +proof (cases "A = {} \ B = {}") + case False + hence "setdist A B > 0" + by (metis IntI assms empty_iff in_closed_iff_infdist_zero order_less_le setdist_attains_inf setdist_pos_le setdist_sym) + hence "eventually (\\. \ < setdist A B) (at_right 0)" + using eventually_at_right_field by blast + thus ?thesis + by eventually_elim (auto intro: setdist_gt_setdist) +qed auto + +lemma setdist_gt_symI: "setdist_gt \ A B \ setdist_gt \ B A" + by (force simp: setdist_gt_def dist_commute) + +lemma setdist_gt_sym: "setdist_gt \ A B \ setdist_gt \ B A" + by (force simp: setdist_gt_def dist_commute) + +lemma eventually_setdist_gt_at_right_0_mult_iff: + assumes "c > 0" + shows "eventually (\\. setdist_gt (c * \) A B) (at_right 0) \ + eventually (\\. setdist_gt \ A B) (at_right 0)" +proof - + have "eventually (\\. setdist_gt (c * \) A B) (at_right 0) \ + eventually (\\. setdist_gt \ A B) (filtermap ((*) c) (at_right 0))" + by (simp add: eventually_filtermap) + also have "filtermap ((*) c) (at_right 0) = at_right 0" + by (subst filtermap_times_pos_at_right) (use assms in auto) + finally show ?thesis . +qed + end diff -r b6f3eb537d91 -r 268c81842883 src/HOL/Analysis/Product_Vector.thy --- a/src/HOL/Analysis/Product_Vector.thy Thu Feb 09 08:35:50 2023 +0000 +++ b/src/HOL/Analysis/Product_Vector.thy Thu Feb 09 16:29:53 2023 +0000 @@ -131,7 +131,10 @@ instance.. end -instantiation\<^marker>\tag unimportant\ prod :: (uniform_space, uniform_space) uniform_space begin +subsubsection \Uniform spaces\ + +instantiation\<^marker>\tag unimportant\ prod :: (uniform_space, uniform_space) uniform_space +begin instance proof standard fix U :: \('a \ 'b) set\ @@ -216,6 +219,133 @@ qed end + +lemma (in uniform_space) nhds_eq_comap_uniformity: "nhds x = filtercomap (\y. (x, y)) uniformity" +proof - + have *: "eventually P (filtercomap (\y. (x, y)) F) \ + eventually (\z. fst z = x \ P (snd z)) F" for P :: "'a \ bool" and F + unfolding eventually_filtercomap + by (smt (verit) eventually_elim2 fst_conv prod.collapse snd_conv) + thus ?thesis + unfolding filter_eq_iff + by (subst *) (auto simp: eventually_nhds_uniformity case_prod_unfold) +qed + +lemma uniformity_of_uniform_continuous_invariant: + fixes f :: "'a :: uniform_space \ 'a \ 'a" + assumes "filterlim (\((a,b),(c,d)). (f a c, f b d)) uniformity (uniformity \\<^sub>F uniformity)" + assumes "eventually P uniformity" + obtains Q where "eventually Q uniformity" "\a b c. Q (a, b) \ P (f a c, f b c)" + using eventually_compose_filterlim[OF assms(2,1)] uniformity_refl + by (fastforce simp: case_prod_unfold eventually_filtercomap eventually_prod_same) + +class uniform_topological_monoid_add = topological_monoid_add + uniform_space + + assumes uniformly_continuous_add': + "filterlim (\((a,b), (c,d)). (a + c, b + d)) uniformity (uniformity \\<^sub>F uniformity)" + +lemma uniformly_continuous_add: + "uniformly_continuous_on UNIV (\(x :: 'a :: uniform_topological_monoid_add,y). x + y)" + using uniformly_continuous_add'[where ?'a = 'a] + by (simp add: uniformly_continuous_on_uniformity case_prod_unfold uniformity_prod_def filterlim_filtermap) + +lemma filterlim_fst: "filterlim fst F (F \\<^sub>F G)" + by (simp add: filterlim_def filtermap_fst_prod_filter) + +lemma filterlim_snd: "filterlim snd G (F \\<^sub>F G)" + by (simp add: filterlim_def filtermap_snd_prod_filter) + +class uniform_topological_group_add = topological_group_add + uniform_topological_monoid_add + + assumes uniformly_continuous_uminus': "filterlim (\(a, b). (-a, -b)) uniformity uniformity" +begin + +lemma uniformly_continuous_minus': + "filterlim (\((a,b), (c,d)). (a - c, b - d)) uniformity (uniformity \\<^sub>F uniformity)" +proof - + have "filterlim ((\((a,b), (c,d)). (a + c, b + d)) \ (\((a,b), (c,d)). ((a, b), (-c, -d)))) + uniformity (uniformity \\<^sub>F uniformity)" + unfolding o_def using uniformly_continuous_uminus' + by (intro filterlim_compose[OF uniformly_continuous_add']) + (auto simp: case_prod_unfold intro!: filterlim_Pair + filterlim_fst filterlim_compose[OF _ filterlim_snd]) + thus ?thesis + by (simp add: o_def case_prod_unfold) +qed + +end + +lemma uniformly_continuous_uminus: + "uniformly_continuous_on UNIV (\x :: 'a :: uniform_topological_group_add. -x)" + using uniformly_continuous_uminus'[where ?'a = 'a] + by (simp add: uniformly_continuous_on_uniformity) + +lemma uniformly_continuous_minus: + "uniformly_continuous_on UNIV (\(x :: 'a :: uniform_topological_group_add,y). x - y)" + using uniformly_continuous_minus'[where ?'a = 'a] + by (simp add: uniformly_continuous_on_uniformity case_prod_unfold uniformity_prod_def filterlim_filtermap) + + + +lemma real_normed_vector_is_uniform_topological_group_add [Pure.intro]: + "OFCLASS('a :: real_normed_vector, uniform_topological_group_add_class)" +proof + show "filterlim (\((a::'a,b), (c,d)). (a + c, b + d)) uniformity (uniformity \\<^sub>F uniformity)" + unfolding filterlim_def le_filter_def eventually_filtermap case_prod_unfold + proof safe + fix P :: "'a \ 'a \ bool" + assume "eventually P uniformity" + then obtain \ where \: "\ > 0" "\x y. dist x y < \ \ P (x, y)" + by (auto simp: eventually_uniformity_metric) + define Q where "Q = (\(x::'a,y). dist x y < \ / 2)" + have Q: "eventually Q uniformity" + unfolding eventually_uniformity_metric Q_def using \\ > 0\ + by (meson case_prodI divide_pos_pos zero_less_numeral) + have "P (a + c, b + d)" if "Q (a, b)" "Q (c, d)" for a b c d + proof - + have "dist (a + c) (b + d) \ dist a b + dist c d" + by (simp add: dist_norm norm_diff_triangle_ineq) + also have "\ < \" + using that by (auto simp: Q_def) + finally show ?thesis + by (intro \) + qed + thus "\\<^sub>F x in uniformity \\<^sub>F uniformity. P (fst (fst x) + fst (snd x), snd (fst x) + snd (snd x))" + unfolding eventually_prod_filter by (intro exI[of _ Q] conjI Q) auto + qed +next + show "filterlim (\((a::'a), b). (-a, -b)) uniformity uniformity" + unfolding filterlim_def le_filter_def eventually_filtermap + proof safe + fix P :: "'a \ 'a \ bool" + assume "eventually P uniformity" + then obtain \ where \: "\ > 0" "\x y. dist x y < \ \ P (x, y)" + by (auto simp: eventually_uniformity_metric) + show "\\<^sub>F x in uniformity. P (case x of (a, b) \ (- a, - b))" + unfolding eventually_uniformity_metric + by (intro exI[of _ \]) (auto intro!: \ simp: dist_norm norm_minus_commute) + qed +qed + +instance real :: uniform_topological_group_add .. +instance complex :: uniform_topological_group_add .. + +lemma cauchy_seq_finset_iff_vanishing: + "uniformity = filtercomap (\(x,y). y - x :: 'a :: uniform_topological_group_add) (nhds 0)" +proof - + have "filtercomap (\x. (0, case x of (x, y) \ y - (x :: 'a))) uniformity \ uniformity" + apply (simp add: le_filter_def eventually_filtercomap) + using uniformity_of_uniform_continuous_invariant[OF uniformly_continuous_add'] + by (metis diff_self eq_diff_eq) + moreover + have "uniformity \ filtercomap (\x. (0, case x of (x, y) \ y - (x :: 'a))) uniformity" + apply (simp add: le_filter_def eventually_filtercomap) + using uniformity_of_uniform_continuous_invariant[OF uniformly_continuous_minus'] + by (metis (mono_tags) diff_self eventually_mono surjective_pairing) + ultimately show ?thesis + by (simp add: nhds_eq_comap_uniformity filtercomap_filtercomap) +qed + +subsubsection \Metric spaces\ + instantiation\<^marker>\tag unimportant\ prod :: (metric_space, metric_space) uniformity_dist begin instance proof @@ -422,7 +552,7 @@ using uniformly_continuous_on_prod_metric[of UNIV UNIV] by auto -text \This logically belong with the real vector spaces by we only have the necessary lemmas now.\ +text \This logically belong with the real vector spaces but we only have the necessary lemmas now.\ lemma isUCont_plus[simp]: shows \isUCont (\(x::'a::real_normed_vector,y). x+y)\ proof (rule isUCont_prod_metric[THEN iffD2], intro allI impI, simp) diff -r b6f3eb537d91 -r 268c81842883 src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy --- a/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy Thu Feb 09 08:35:50 2023 +0000 +++ b/src/HOL/Complex_Analysis/Cauchy_Integral_Formula.thy Thu Feb 09 16:29:53 2023 +0000 @@ -2567,6 +2567,11 @@ definition\<^marker>\tag important\ fps_expansion :: "(complex \ complex) \ complex \ complex fps" where "fps_expansion f z0 = Abs_fps (\n. (deriv ^^ n) f z0 / fact n)" +lemma fps_expansion_cong: + assumes "\\<^sub>F w in nhds x. f w =g w" + shows "fps_expansion f x = fps_expansion g x" + unfolding fps_expansion_def using assms higher_deriv_cong_ev by fastforce + lemma fixes r :: ereal assumes "f holomorphic_on eball z0 r" diff -r b6f3eb537d91 -r 268c81842883 src/HOL/Complex_Analysis/Complex_Singularities.thy --- a/src/HOL/Complex_Analysis/Complex_Singularities.thy Thu Feb 09 08:35:50 2023 +0000 +++ b/src/HOL/Complex_Analysis/Complex_Singularities.thy Thu Feb 09 16:29:53 2023 +0000 @@ -39,6 +39,11 @@ shows "NO_MATCH 0 z \ is_pole f z \ is_pole (\x. f (z + x)) 0" by (metis is_pole_shift_0) +lemma is_pole_compose_iff: + assumes "filtermap g (at x) = (at y)" + shows "is_pole (f \ g) x \ is_pole f y" + unfolding is_pole_def filterlim_def filtermap_compose assms .. + lemma is_pole_inverse_holomorphic: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" @@ -130,7 +135,106 @@ shows "is_pole (\w. f w / w ^ n) 0" using is_pole_basic[of f A 0] assms by simp -text \The proposition +lemma is_pole_compose: + assumes "is_pole f w" "g \z\ w" "eventually (\z. g z \ w) (at z)" + shows "is_pole (\x. f (g x)) z" + using assms(1) unfolding is_pole_def + by (rule filterlim_compose) (use assms in \auto simp: filterlim_at\) + +lemma is_pole_plus_const_iff: + "is_pole f z \ is_pole (\x. f x + c) z" +proof + assume "is_pole f z" + then have "filterlim f at_infinity (at z)" unfolding is_pole_def . + moreover have "((\_. c) \ c) (at z)" by auto + ultimately have " LIM x (at z). f x + c :> at_infinity" + using tendsto_add_filterlim_at_infinity'[of f "at z"] by auto + then show "is_pole (\x. f x + c) z" unfolding is_pole_def . +next + assume "is_pole (\x. f x + c) z" + then have "filterlim (\x. f x + c) at_infinity (at z)" + unfolding is_pole_def . + moreover have "((\_. -c) \ -c) (at z)" by auto + ultimately have " LIM x (at z). f x :> at_infinity" + using tendsto_add_filterlim_at_infinity'[of "(\x. f x + c)" + "at z" "(\_. - c)" "-c"] + by auto + then show "is_pole f z" unfolding is_pole_def . +qed + +lemma is_pole_minus_const_iff: + "is_pole (\x. f x - c) z \ is_pole f z" + using is_pole_plus_const_iff [of f z "-c"] by simp + +lemma is_pole_alt: + "is_pole f x = (\B>0. \U. open U \ x\U \ (\y\U. y\x \ norm (f y)\B))" + unfolding is_pole_def + unfolding filterlim_at_infinity[of 0,simplified] eventually_at_topological + by auto + +lemma is_pole_mult_analytic_nonzero1: + assumes "is_pole g x" "f analytic_on {x}" "f x \ 0" + shows "is_pole (\x. f x * g x) x" + unfolding is_pole_def +proof (rule tendsto_mult_filterlim_at_infinity) + show "f \x\ f x" + using assms by (simp add: analytic_at_imp_isCont isContD) +qed (use assms in \auto simp: is_pole_def\) + +lemma is_pole_mult_analytic_nonzero2: + assumes "is_pole f x" "g analytic_on {x}" "g x \ 0" + shows "is_pole (\x. f x * g x) x" + by (subst mult.commute, rule is_pole_mult_analytic_nonzero1) (use assms in auto) + +lemma is_pole_mult_analytic_nonzero1_iff: + assumes "f analytic_on {x}" "f x \ 0" + shows "is_pole (\x. f x * g x) x \ is_pole g x" +proof + assume "is_pole g x" + thus "is_pole (\x. f x * g x) x" + by (intro is_pole_mult_analytic_nonzero1 assms) +next + assume "is_pole (\x. f x * g x) x" + hence "is_pole (\x. inverse (f x) * (f x * g x)) x" + by (rule is_pole_mult_analytic_nonzero1) + (use assms in \auto intro!: analytic_intros\) + also have "?this \ is_pole g x" + proof (rule is_pole_cong) + have "eventually (\x. f x \ 0) (at x)" + using assms by (simp add: analytic_at_neq_imp_eventually_neq) + thus "eventually (\x. inverse (f x) * (f x * g x) = g x) (at x)" + by eventually_elim auto + qed auto + finally show "is_pole g x" . +qed + +lemma is_pole_mult_analytic_nonzero2_iff: + assumes "g analytic_on {x}" "g x \ 0" + shows "is_pole (\x. f x * g x) x \ is_pole f x" + by (subst mult.commute, rule is_pole_mult_analytic_nonzero1_iff) (fact assms)+ + +lemma frequently_const_imp_not_is_pole: + fixes z :: "'a::first_countable_topology" + assumes "frequently (\w. f w = c) (at z)" + shows "\ is_pole f z" +proof + assume "is_pole f z" + from assms have "z islimpt {w. f w = c}" + by (simp add: islimpt_conv_frequently_at) + then obtain g where g: "\n. g n \ {w. f w = c} - {z}" "g \ z" + unfolding islimpt_sequential by blast + then have "(f \ g) \ c" + by (simp add: tendsto_eventually) + moreover have *: "filterlim g (at z) sequentially" + using g by (auto simp: filterlim_at) + have "filterlim (f \ g) at_infinity sequentially" + unfolding o_def by (rule filterlim_compose [OF _ *]) + (use \is_pole f z\ in \simp add: is_pole_def\) + ultimately show False + using not_tendsto_and_filterlim_at_infinity trivial_limit_sequentially by blast +qed + + 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).\ @@ -140,6 +244,39 @@ definition isolated_singularity_at::"[complex \ complex, complex] \ bool" where "isolated_singularity_at f z = (\r>0. f analytic_on ball z r-{z})" +lemma not_essential_cong: + assumes "eventually (\x. f x = g x) (at z)" "z = z'" + shows "not_essential f z \ not_essential g z'" + unfolding not_essential_def using assms filterlim_cong is_pole_cong by fastforce + +lemma isolated_singularity_at_cong: + assumes "eventually (\x. f x = g x) (at z)" "z = z'" + shows "isolated_singularity_at f z \ isolated_singularity_at g z'" +proof - + have "isolated_singularity_at g z" + if "isolated_singularity_at f z" "eventually (\x. f x = g x) (at z)" for f g + proof - + from that(1) obtain r where r: "r > 0" "f analytic_on ball z r - {z}" + by (auto simp: isolated_singularity_at_def) + from that(2) obtain r' where r': "r' > 0" "\x\ball z r'-{z}. f x = g x" + unfolding eventually_at_filter eventually_nhds_metric by (auto simp: dist_commute) + + have "f holomorphic_on ball z r - {z}" + using r(2) by (subst (asm) analytic_on_open) auto + hence "f holomorphic_on ball z (min r r') - {z}" + by (rule holomorphic_on_subset) auto + also have "?this \ g holomorphic_on ball z (min r r') - {z}" + using r' by (intro holomorphic_cong) auto + also have "\ \ g analytic_on ball z (min r r') - {z}" + by (subst analytic_on_open) auto + finally show ?thesis + unfolding isolated_singularity_at_def + by (intro exI[of _ "min r r'"]) (use \r > 0\ \r' > 0\ in auto) + qed + from this[of f g] this[of g f] assms show ?thesis + by (auto simp: eq_commute) +qed + lemma removable_singularity: assumes "f holomorphic_on A - {x}" "open A" assumes "f \x\ c" @@ -795,6 +932,24 @@ using assms unfolding isolated_singularity_at_def by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff) +lemma isolated_singularity_at_altdef: + "isolated_singularity_at f z \ eventually (\z. f analytic_on {z}) (at z)" +proof + assume "isolated_singularity_at f z" + then obtain r where r: "r > 0" "f analytic_on ball z r - {z}" + unfolding isolated_singularity_at_def by blast + have "eventually (\w. w \ ball z r - {z}) (at z)" + using r(1) by (intro eventually_at_in_open) auto + thus "eventually (\z. f analytic_on {z}) (at z)" + by eventually_elim (use r analytic_on_subset in auto) +next + assume "eventually (\z. f analytic_on {z}) (at z)" + then obtain A where A: "open A" "z \ A" "\w. w \ A - {z} \ f analytic_on {w}" + unfolding eventually_at_topological by blast + then show "isolated_singularity_at f z" + by (meson analytic_imp_holomorphic analytic_on_analytic_at isolated_singularity_at_holomorphic) +qed + lemma isolated_singularity_at_shift: assumes "isolated_singularity_at (\x. f (x + w)) z" shows "isolated_singularity_at f (z + w)" @@ -863,6 +1018,20 @@ by (auto simp: not_essential_def) qed +lemma not_essential_analytic: + assumes "f analytic_on {z}" + shows "not_essential f z" + using analytic_at assms not_essential_holomorphic by blast + +lemma not_essential_id [singularity_intros]: "not_essential (\w. w) z" + by (simp add: not_essential_analytic) + +lemma is_pole_imp_not_essential [intro]: "is_pole f z \ not_essential f z" + by (auto simp: not_essential_def) + +lemma tendsto_imp_not_essential [intro]: "f \z\ c \ not_essential f z" + by (auto simp: not_essential_def) + lemma eventually_not_pole: assumes "isolated_singularity_at f z" shows "eventually (\w. \is_pole f w) (at z)" @@ -901,7 +1070,18 @@ thus ?thesis by simp qed -subsubsection \The order of non-essential singularities (i.e. removable singularities or poles)\ +lemma isolated_singularity_at_analytic: + assumes "f analytic_on {z}" + shows "isolated_singularity_at f z" +proof - + from assms obtain r where r: "r > 0" "f holomorphic_on ball z r" + by (auto simp: analytic_on_def) + show ?thesis + by (rule isolated_singularity_at_holomorphic[of f "ball z r"]) + (use \r > 0\ in \auto intro!: holomorphic_on_subset[OF r(2)]\) +qed + +subsection \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 @@ -1658,6 +1838,33 @@ using \z=z'\ unfolding P_def zorder_def zor_poly_def by auto qed +lemma zorder_times_analytic': + assumes "isolated_singularity_at f z" "not_essential f z" + assumes "g analytic_on {z}" "frequently (\z. f z * g z \ 0) (at z)" + shows "zorder (\x. f x * g x) z = zorder f z + zorder g z" +proof (rule zorder_times) + show "isolated_singularity_at g z" "not_essential g z" + by (intro isolated_singularity_at_analytic not_essential_analytic assms)+ +qed (use assms in auto) + +lemma zorder_cmult: + assumes "c \ 0" + shows "zorder (\z. c * f z) z = zorder f z" +proof - + define P where + "P = (\f n h r. 0 < r \ 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))" + have *: "P (\x. c * f x) n (\x. c * h x) r" if "P f n h r" "c \ 0" for f n h r c + using that unfolding P_def by (auto intro!: holomorphic_intros) + have "(\h r. P (\x. c * f x) n h r) \ (\h r. P f n h r)" for n + using *[of f n _ _ c] *[of "\x. c * f x" n _ _ "inverse c"] \c \ 0\ + by (fastforce simp: field_simps) + hence "(THE n. \h r. P (\x. c * f x) n h r) = (THE n. \h r. P f n h r)" + by simp + thus ?thesis + by (simp add: zorder_def P_def) +qed + lemma zorder_nonzero_div_power: assumes sz: "open s" "z \ s" "f holomorphic_on s" "f z \ 0" and "n > 0" shows "zorder (\w. f w / (w - z) ^ n) z = - n" @@ -2269,4 +2476,263 @@ qed qed + +lemma deriv_divide_is_pole: \\Generalises @{thm zorder_deriv}\ + fixes f g::"complex \ complex" and z::complex + assumes f_iso:"isolated_singularity_at f z" + and f_ness:"not_essential f z" + and fg_nconst: "\\<^sub>Fw in (at z). deriv f w * f w \ 0" + and f_ord:"zorder f z \0" + shows "is_pole (\z. deriv f z / f z) z" +proof (rule neg_zorder_imp_is_pole) + define ff where "ff=(\w. deriv f w / f w)" + show "isolated_singularity_at ff z" + using f_iso f_ness unfolding ff_def + by (auto intro:singularity_intros) + show "not_essential ff z" + unfolding ff_def using f_ness f_iso + by (auto intro:singularity_intros) + + have "zorder ff z = zorder (deriv f) z - zorder f z" + unfolding ff_def using f_iso f_ness fg_nconst + apply (rule_tac zorder_divide) + by (auto intro:singularity_intros) + moreover have "zorder (deriv f) z = zorder f z - 1" + proof (rule zorder_deriv_minus_1) + show " \\<^sub>F w in at z. f w \ 0" + using fg_nconst frequently_elim1 by fastforce + qed (use f_iso f_ness f_ord in auto) + ultimately show "zorder ff z < 0" by auto + + show "\\<^sub>F w in at z. ff w \ 0" + unfolding ff_def using fg_nconst by auto +qed + +lemma is_pole_deriv_divide_is_pole: + fixes f g::"complex \ complex" and z::complex + assumes f_iso:"isolated_singularity_at f z" + and "is_pole f z" + shows "is_pole (\z. deriv f z / f z) z" +proof (rule deriv_divide_is_pole[OF f_iso]) + show "not_essential f z" + using \is_pole f z\ unfolding not_essential_def by auto + show "\\<^sub>F w in at z. deriv f w * f w \ 0" + apply (rule isolated_pole_imp_nzero_times) + using assms by auto + show "zorder f z \ 0" + using isolated_pole_imp_neg_zorder assms by fastforce +qed + +subsection \Isolated zeroes\ + +definition isolated_zero :: "(complex \ complex) \ complex \ bool" where + "isolated_zero f z \ f z = 0 \ eventually (\z. f z \ 0) (at z)" + +lemma isolated_zero_altdef: "isolated_zero f z \ f z = 0 \ \z islimpt {z. f z = 0}" + unfolding isolated_zero_def eventually_at_filter eventually_nhds islimpt_def by blast + +lemma isolated_zero_mult1: + assumes "isolated_zero f x" "isolated_zero g x" + shows "isolated_zero (\x. f x * g x) x" +proof - + have "eventually (\x. f x \ 0) (at x)" "eventually (\x. g x \ 0) (at x)" + using assms unfolding isolated_zero_def by auto + hence "eventually (\x. f x * g x \ 0) (at x)" + by eventually_elim auto + with assms show ?thesis + by (auto simp: isolated_zero_def) +qed + +lemma isolated_zero_mult2: + assumes "isolated_zero f x" "g x \ 0" "g analytic_on {x}" + shows "isolated_zero (\x. f x * g x) x" +proof - + have "eventually (\x. f x \ 0) (at x)" + using assms unfolding isolated_zero_def by auto + moreover have "eventually (\x. g x \ 0) (at x)" + using analytic_at_neq_imp_eventually_neq[of g x 0] assms by auto + ultimately have "eventually (\x. f x * g x \ 0) (at x)" + by eventually_elim auto + thus ?thesis + using assms(1) by (auto simp: isolated_zero_def) +qed + +lemma isolated_zero_mult3: + assumes "isolated_zero f x" "g x \ 0" "g analytic_on {x}" + shows "isolated_zero (\x. g x * f x) x" + using isolated_zero_mult2[OF assms] by (simp add: mult_ac) + +lemma isolated_zero_prod: + assumes "\x. x \ I \ isolated_zero (f x) z" "I \ {}" "finite I" + shows "isolated_zero (\y. \x\I. f x y) z" + using assms(3,2,1) by (induction rule: finite_ne_induct) (auto intro: isolated_zero_mult1) + +lemma non_isolated_zero': + assumes "isolated_singularity_at f z" "not_essential f z" "f z = 0" "\isolated_zero f z" + shows "eventually (\z. f z = 0) (at z)" +proof (rule not_essential_frequently_0_imp_eventually_0) + from assms show "frequently (\z. f z = 0) (at z)" + by (auto simp: frequently_def isolated_zero_def) +qed fact+ + +lemma non_isolated_zero: + assumes "\isolated_zero f z" "f analytic_on {z}" "f z = 0" + shows "eventually (\z. f z = 0) (nhds z)" +proof - + have "eventually (\z. f z = 0) (at z)" + by (rule non_isolated_zero') + (use assms in \auto intro: not_essential_analytic isolated_singularity_at_analytic\) + with \f z = 0\ show ?thesis + unfolding eventually_at_filter by (auto elim!: eventually_mono) +qed + +lemma not_essential_compose: + assumes "not_essential f (g z)" "g analytic_on {z}" + shows "not_essential (\x. f (g x)) z" +proof (cases "isolated_zero (\w. g w - g z) z") + case False + hence "eventually (\w. g w - g z = 0) (nhds z)" + by (rule non_isolated_zero) (use assms in \auto intro!: analytic_intros\) + hence "not_essential (\x. f (g x)) z \ not_essential (\_. f (g z)) z" + by (intro not_essential_cong refl) + (auto elim!: eventually_mono simp: eventually_at_filter) + thus ?thesis + by (simp add: not_essential_const) +next + case True + hence ev: "eventually (\w. g w \ g z) (at z)" + by (auto simp: isolated_zero_def) + from assms consider c where "f \g z\ c" | "is_pole f (g z)" + by (auto simp: not_essential_def) + have "isCont g z" + by (rule analytic_at_imp_isCont) fact + hence lim: "g \z\ g z" + using isContD by blast + + from assms(1) consider c where "f \g z\ c" | "is_pole f (g z)" + unfolding not_essential_def by blast + thus ?thesis + proof cases + fix c assume "f \g z\ c" + hence "(\x. f (g x)) \z\ c" + by (rule filterlim_compose) (use lim ev in \auto simp: filterlim_at\) + thus ?thesis + by (auto simp: not_essential_def) + next + assume "is_pole f (g z)" + hence "is_pole (\x. f (g x)) z" + by (rule is_pole_compose) fact+ + thus ?thesis + by (auto simp: not_essential_def) + qed +qed + +subsection \Isolated points\ + +definition isolated_points_of :: "complex set \ complex set" where + "isolated_points_of A = {z\A. eventually (\w. w \ A) (at z)}" + +lemma isolated_points_of_altdef: "isolated_points_of A = {z\A. \z islimpt A}" + unfolding isolated_points_of_def islimpt_def eventually_at_filter eventually_nhds by blast + +lemma isolated_points_of_empty [simp]: "isolated_points_of {} = {}" + and isolated_points_of_UNIV [simp]: "isolated_points_of UNIV = {}" + by (auto simp: isolated_points_of_def) + +lemma isolated_points_of_open_is_empty [simp]: "open A \ isolated_points_of A = {}" + unfolding isolated_points_of_altdef + by (simp add: interior_limit_point interior_open) + +lemma isolated_points_of_subset: "isolated_points_of A \ A" + by (auto simp: isolated_points_of_def) + +lemma isolated_points_of_discrete: + assumes "discrete A" + shows "isolated_points_of A = A" + using assms by (auto simp: isolated_points_of_def discrete_altdef) + +lemmas uniform_discreteI1 = uniformI1 +lemmas uniform_discreteI2 = uniformI2 + +lemma isolated_singularity_at_compose: + assumes "isolated_singularity_at f (g z)" "g analytic_on {z}" + shows "isolated_singularity_at (\x. f (g x)) z" +proof (cases "isolated_zero (\w. g w - g z) z") + case False + hence "eventually (\w. g w - g z = 0) (nhds z)" + by (rule non_isolated_zero) (use assms in \auto intro!: analytic_intros\) + hence "isolated_singularity_at (\x. f (g x)) z \ isolated_singularity_at (\_. f (g z)) z" + by (intro isolated_singularity_at_cong refl) + (auto elim!: eventually_mono simp: eventually_at_filter) + thus ?thesis + by (simp add: isolated_singularity_at_const) +next + case True + from assms(1) obtain r where r: "r > 0" "f analytic_on ball (g z) r - {g z}" + by (auto simp: isolated_singularity_at_def) + hence holo_f: "f holomorphic_on ball (g z) r - {g z}" + by (subst (asm) analytic_on_open) auto + from assms(2) obtain r' where r': "r' > 0" "g holomorphic_on ball z r'" + by (auto simp: analytic_on_def) + + have "continuous_on (ball z r') g" + using holomorphic_on_imp_continuous_on r' by blast + hence "isCont g z" + using r' by (subst (asm) continuous_on_eq_continuous_at) auto + hence "g \z\ g z" + using isContD by blast + hence "eventually (\w. g w \ ball (g z) r) (at z)" + using \r > 0\ unfolding tendsto_def by force + moreover have "eventually (\w. g w \ g z) (at z)" using True + by (auto simp: isolated_zero_def elim!: eventually_mono) + ultimately have "eventually (\w. g w \ ball (g z) r - {g z}) (at z)" + by eventually_elim auto + then obtain r'' where r'': "r'' > 0" "\w\ball z r''-{z}. g w \ ball (g z) r - {g z}" + unfolding eventually_at_filter eventually_nhds_metric ball_def + by (auto simp: dist_commute) + have "f \ g holomorphic_on ball z (min r' r'') - {z}" + proof (rule holomorphic_on_compose_gen) + show "g holomorphic_on ball z (min r' r'') - {z}" + by (rule holomorphic_on_subset[OF r'(2)]) auto + show "f holomorphic_on ball (g z) r - {g z}" + by fact + show "g ` (ball z (min r' r'') - {z}) \ ball (g z) r - {g z}" + using r'' by force + qed + hence "f \ g analytic_on ball z (min r' r'') - {z}" + by (subst analytic_on_open) auto + thus ?thesis using \r' > 0\ \r'' > 0\ + by (auto simp: isolated_singularity_at_def o_def intro!: exI[of _ "min r' r''"]) +qed + +lemma is_pole_power_int_0: + assumes "f analytic_on {x}" "isolated_zero f x" "n < 0" + shows "is_pole (\x. f x powi n) x" +proof - + have "f \x\ f x" + using assms(1) by (simp add: analytic_at_imp_isCont isContD) + with assms show ?thesis + unfolding is_pole_def + by (intro filterlim_power_int_neg_at_infinity) (auto simp: isolated_zero_def) +qed + +lemma isolated_zero_imp_not_constant_on: + assumes "isolated_zero f x" "x \ A" "open A" + shows "\f constant_on A" +proof + assume "f constant_on A" + then obtain c where c: "\x. x \ A \ f x = c" + by (auto simp: constant_on_def) + from assms and c[of x] have [simp]: "c = 0" + by (auto simp: isolated_zero_def) + have "eventually (\x. f x \ 0) (at x)" + using assms by (auto simp: isolated_zero_def) + moreover have "eventually (\x. x \ A) (at x)" + using assms by (intro eventually_at_in_open') auto + ultimately have "eventually (\x. False) (at x)" + by eventually_elim (use c in auto) + thus False + by simp +qed + end diff -r b6f3eb537d91 -r 268c81842883 src/HOL/Complex_Analysis/Conformal_Mappings.thy --- a/src/HOL/Complex_Analysis/Conformal_Mappings.thy Thu Feb 09 08:35:50 2023 +0000 +++ b/src/HOL/Complex_Analysis/Conformal_Mappings.thy Thu Feb 09 16:29:53 2023 +0000 @@ -1079,7 +1079,7 @@ qed qed -text\Hence a nice clean inverse function theorem\ +subsubsection \Hence a nice clean inverse function theorem\ lemma has_field_derivative_inverse_strong: fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" @@ -1140,6 +1140,78 @@ qed qed +subsubsection \ Holomorphism of covering maps and lifts.\ + +lemma covering_space_lift_is_holomorphic: + assumes cov: "covering_space C p S" + and C: "open C" "p holomorphic_on C" + and holf: "f holomorphic_on U" and fim: "f ` U \ S" and gim: "g ` U \ C" + and contg: "continuous_on U g" and pg_f: "\x. x \ U \ p(g x) = f x" + shows "g holomorphic_on U" + unfolding holomorphic_on_def +proof (intro strip) + fix z + assume "z \ U" + with fim have "f z \ S" by blast + then obtain T \ where "f z \ T" and opeT: "openin (top_of_set S) T" + and UV: "\\ = C \ p -` T" + and "\W. W \ \ \ openin (top_of_set C) W" + and disV: "pairwise disjnt \" and homeV: "\W. W \ \ \ \q. homeomorphism W T p q" + using cov fim unfolding covering_space_def by meson + then have "\W. W \ \ \ open W \ W \ C" + by (metis \open C\ inf_le1 open_Int openin_open) + then obtain V where "V \ \" "g z \ V" "open V" "V \ C" + by (metis IntI UnionE image_subset_iff vimageI UV \f z \ T\ \z \ U\ gim pg_f) + have holp: "p holomorphic_on V" + using \V \ C\ \p holomorphic_on C\ holomorphic_on_subset by blast + moreover have injp: "inj_on p V" + by (metis \V \ \\ homeV homeomorphism_def inj_on_inverseI) + ultimately + obtain p' where holp': "p' holomorphic_on (p ` V)" and pp': "\z. z \ V \ p'(p z) = z" + using \open V\ holomorphic_has_inverse by metis + have "z \ U \ g -` V" + using \g z \ V\ \z \ U\ by blast + moreover have "openin (top_of_set U) (U \ g -` V)" + using continuous_openin_preimage [OF contg gim] + by (meson \open V\ contg continuous_openin_preimage_eq) + ultimately obtain \ where "\>0" and e: "ball z \ \ U \ g -` V" + by (force simp add: openin_contains_ball) + show "g field_differentiable at z within U" + proof (rule field_differentiable_transform_within) + show "(0::real) < \" + by (simp add: \0 < \\) + show "z \ U" + by (simp add: \z \ U\) + show "(p' o f) x' = g x'" if "x' \ U" "dist x' z < \" for x' + using that + by (metis Int_iff comp_apply dist_commute e mem_ball pg_f pp' subsetD vimage_eq) + have "open (p ` V)" + using \open V\ holp injp open_mapping_thm3 by force + moreover have "f z \ p ` V" + by (metis \z \ U\ image_iff pg_f \g z \ V\) + ultimately have "p' field_differentiable at (f z)" + using holomorphic_on_imp_differentiable_at holp' by blast + moreover have "f field_differentiable at z within U" + by (metis (no_types) \z \ U\ holf holomorphic_on_def) + ultimately show "(p' o f) field_differentiable at z within U" + by (metis (no_types) field_differentiable_at_within field_differentiable_compose_within) + qed +qed + +lemma covering_space_lift_holomorphic: + assumes cov: "covering_space C p S" + and C: "open C" "p holomorphic_on C" + and f: "f holomorphic_on U" "f ` U \ S" + and U: "simply_connected U" "locally path_connected U" + obtains g where "g holomorphic_on U" "g ` U \ C" "\y. y \ U \ p(g y) = f y" +proof - + obtain g where "continuous_on U g" "g ` U \ C" "\y. y \ U \ p(g y) = f y" + using covering_space_lift [OF cov U] + using f holomorphic_on_imp_continuous_on by blast + then show ?thesis + by (metis C cov covering_space_lift_is_holomorphic f that) +qed + subsection\The Schwarz Lemma\ lemma Schwarz1: @@ -1923,8 +1995,7 @@ qed show ?thesis apply (rule Bloch_unit [OF 1 2]) - apply (rule_tac b="(C * of_real r) * b" in that) - using image_mono sb1 sb2 by fastforce + using image_mono sb1 sb2 that by fastforce qed corollary Bloch_general: @@ -1954,10 +2025,7 @@ 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 + using Bloch [OF 1 \t > 0\ rle] * by (metis image_mono order_trans that) qed qed qed