# HG changeset patch # User paulson # Date 1675255413 0 # Node ID 0fb350e7477b0d7d87909bfb4652066fc3894814 # Parent 9a60c17595437d39a706fd295440c611f468f8a7 More new material thanks to Manuel diff -r 9a60c1759543 -r 0fb350e7477b src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Tue Jan 31 14:05:16 2023 +0000 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Wed Feb 01 12:43:33 2023 +0000 @@ -250,39 +250,56 @@ by (auto simp: not_le) qed (insert assms, auto) -lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \ (\z. -(f z)) holomorphic_on s" +lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on A \ (\z. -(f z)) holomorphic_on A" by (metis field_differentiable_minus holomorphic_on_def) lemma holomorphic_on_add [holomorphic_intros]: - "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z + g z) holomorphic_on s" + "\f holomorphic_on A; g holomorphic_on A\ \ (\z. f z + g z) holomorphic_on A" unfolding holomorphic_on_def by (metis field_differentiable_add) lemma holomorphic_on_diff [holomorphic_intros]: - "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z - g z) holomorphic_on s" + "\f holomorphic_on A; g holomorphic_on A\ \ (\z. f z - g z) holomorphic_on A" unfolding holomorphic_on_def by (metis field_differentiable_diff) lemma holomorphic_on_mult [holomorphic_intros]: - "\f holomorphic_on s; g holomorphic_on s\ \ (\z. f z * g z) holomorphic_on s" + "\f holomorphic_on A; g holomorphic_on A\ \ (\z. f z * g z) holomorphic_on A" unfolding holomorphic_on_def by (metis field_differentiable_mult) lemma holomorphic_on_inverse [holomorphic_intros]: - "\f holomorphic_on s; \z. z \ s \ f z \ 0\ \ (\z. inverse (f z)) holomorphic_on s" + "\f holomorphic_on A; \z. z \ A \ f z \ 0\ \ (\z. inverse (f z)) holomorphic_on A" unfolding holomorphic_on_def by (metis field_differentiable_inverse) lemma holomorphic_on_divide [holomorphic_intros]: - "\f holomorphic_on s; g holomorphic_on s; \z. z \ s \ g z \ 0\ \ (\z. f z / g z) holomorphic_on s" + "\f holomorphic_on A; g holomorphic_on A; \z. z \ A \ g z \ 0\ \ (\z. f z / g z) holomorphic_on A" unfolding holomorphic_on_def by (metis field_differentiable_divide) lemma holomorphic_on_power [holomorphic_intros]: - "f holomorphic_on s \ (\z. (f z)^n) holomorphic_on s" + "f holomorphic_on A \ (\z. (f z)^n) holomorphic_on A" unfolding holomorphic_on_def by (metis field_differentiable_power) +lemma holomorphic_on_power_int [holomorphic_intros]: + assumes nz: "n \ 0 \ (\x\A. f x \ 0)" and f: "f holomorphic_on A" + shows "(\x. f x powi n) holomorphic_on A" +proof (cases "n \ 0") + case True + have "(\x. f x ^ nat n) holomorphic_on A" + by (simp add: f holomorphic_on_power) + with True show ?thesis + by (simp add: power_int_def) +next + case False + hence "(\x. inverse (f x ^ nat (-n))) holomorphic_on A" + using nz by (auto intro!: holomorphic_intros f) + with False show ?thesis + by (simp add: power_int_def power_inverse) +qed + lemma holomorphic_on_sum [holomorphic_intros]: - "(\i. i \ I \ (f i) holomorphic_on s) \ (\x. sum (\i. f i x) I) holomorphic_on s" + "(\i. i \ I \ (f i) holomorphic_on A) \ (\x. sum (\i. f i x) I) holomorphic_on A" unfolding holomorphic_on_def by (metis field_differentiable_sum) lemma holomorphic_on_prod [holomorphic_intros]: - "(\i. i \ I \ (f i) holomorphic_on s) \ (\x. prod (\i. f i x) I) holomorphic_on s" + "(\i. i \ I \ (f i) holomorphic_on A) \ (\x. prod (\i. f i x) I) holomorphic_on A" by (induction I rule: infinite_finite_induct) (auto intro: holomorphic_intros) lemma holomorphic_pochhammer [holomorphic_intros]: @@ -327,6 +344,17 @@ by (rule DERIV_imp_deriv) (metis DERIV_deriv_iff_field_differentiable has_field_derivative_transform_within_open at_within_open) +lemma holomorphic_on_compose_cnj_cnj: + assumes "f holomorphic_on cnj ` A" "open A" + shows "cnj \ f \ cnj holomorphic_on A" +proof - + have [simp]: "open (cnj ` A)" + unfolding image_cnj_conv_vimage_cnj using assms by (intro open_vimage) auto + show ?thesis + using assms unfolding holomorphic_on_def + by (auto intro!: field_differentiable_cnj_cnj simp: at_within_open_NO_MATCH) +qed + lemma holomorphic_nonconstant: assumes holf: "f holomorphic_on S" and "open S" "\ \ S" "deriv f \ \ 0" shows "\ f constant_on S" @@ -528,6 +556,23 @@ "f analytic_on S \ (\z. (f z) ^ n) analytic_on S" by (induct n) (auto simp: analytic_on_mult) +lemma analytic_on_power_int [analytic_intros]: + assumes nz: "n \ 0 \ (\x\A. f x \ 0)" and f: "f analytic_on A" + shows "(\x. f x powi n) analytic_on A" +proof (cases "n \ 0") + case True + have "(\x. f x ^ nat n) analytic_on A" + using analytic_on_power f by blast + with True show ?thesis + by (simp add: power_int_def) +next + case False + hence "(\x. inverse (f x ^ nat (-n))) analytic_on A" + using nz by (auto intro!: analytic_intros f) + with False show ?thesis + by (simp add: power_int_def power_inverse) +qed + lemma analytic_on_sum [analytic_intros]: "(\i. i \ I \ (f i) analytic_on S) \ (\x. sum (\i. f i x) I) analytic_on S" by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_add) diff -r 9a60c1759543 -r 0fb350e7477b src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy Tue Jan 31 14:05:16 2023 +0000 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Wed Feb 01 12:43:33 2023 +0000 @@ -2296,7 +2296,17 @@ lemma continuous_within_Arg: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Arg" using continuous_at_Arg continuous_at_imp_continuous_within by blast - +lemma Arg_Re_pos: "\Arg z\ < pi / 2 \ Re z > 0 \ z = 0" + using Arg_def Re_Ln_pos_lt by auto + +lemma Arg_Re_nonneg: "\Arg z\ \ pi / 2 \ Re z \ 0" + using Re_Ln_pos_le[of z] by (cases "z = 0") (auto simp: Arg_eq_Im_Ln Arg_zero) + +lemma Arg_times: + assumes "Arg z + Arg w \ {-pi<..pi}" "z \ 0" "w \ 0" + shows "Arg (z * w) = Arg z + Arg w" + using Arg_eq_Im_Ln Ln_times_simple assms by auto + subsection\The Unwinding Number and the Ln product Formula\ text\Note that in this special case the unwinding number is -1, 0 or 1. But it's always an integer.\ diff -r 9a60c1759543 -r 0fb350e7477b src/HOL/Analysis/Derivative.thy --- a/src/HOL/Analysis/Derivative.thy Tue Jan 31 14:05:16 2023 +0000 +++ b/src/HOL/Analysis/Derivative.thy Wed Feb 01 12:43:33 2023 +0000 @@ -2104,6 +2104,12 @@ using assms unfolding field_differentiable_def by (metis DERIV_power) +lemma field_differentiable_cnj_cnj: + assumes "f field_differentiable (at (cnj z))" + shows "(cnj \ f \ cnj) field_differentiable (at z)" + using has_field_derivative_cnj_cnj assms + by (auto simp: field_differentiable_def) + lemma field_differentiable_transform_within: "0 < d \ x \ S \ diff -r 9a60c1759543 -r 0fb350e7477b src/HOL/Complex.thy --- a/src/HOL/Complex.thy Tue Jan 31 14:05:16 2023 +0000 +++ b/src/HOL/Complex.thy Wed Feb 01 12:43:33 2023 +0000 @@ -671,6 +671,22 @@ shows "((\z. cnj (f z)) has_vector_derivative cnj f') (at z within A)" using assms by (auto simp: has_vector_derivative_complex_iff intro: derivative_intros) +lemma has_field_derivative_cnj_cnj: + assumes "(f has_field_derivative F) (at (cnj z))" + shows "((cnj \ f \ cnj) has_field_derivative cnj F) (at z)" +proof - + have "cnj \0\ cnj 0" + by (subst lim_cnj) auto + also have "cnj 0 = 0" + by simp + finally have *: "filterlim cnj (at 0) (at 0)" + by (auto simp: filterlim_at eventually_at_filter) + have "(\h. (f (cnj z + cnj h) - f (cnj z)) / cnj h) \0\ F" + by (rule filterlim_compose[OF _ *]) (use assms in \auto simp: DERIV_def\) + thus ?thesis + by (subst (asm) lim_cnj [symmetric]) (simp add: DERIV_def) +qed + subsection \Basic Lemmas\ @@ -1388,5 +1404,4 @@ then show ?thesis using that by simp qed - end