--- a/NEWS Mon Feb 20 13:59:16 2023 +0100
+++ b/NEWS Mon Feb 20 13:59:42 2023 +0100
@@ -214,6 +214,8 @@
* Theory "HOL-Library.Multiset_Order":
- Added lemmas.
+ asymp_multpHO
+ asymp_not_liftable_to_multpHO
irreflp_on_multpHO[simp]
multpHO_plus_plus[simp]
totalp_multpDM
--- a/src/HOL/Analysis/Complex_Transcendental.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Analysis/Complex_Transcendental.thy Mon Feb 20 13:59:42 2023 +0100
@@ -102,12 +102,7 @@
have "((\<lambda>w. w) has_field_derivative 1) (at z)"
by (intro derivative_intros)
also have "?this \<longleftrightarrow> ((\<lambda>w. exp (f w)) has_field_derivative 1) (at z)"
- proof (rule DERIV_cong_ev)
- have "eventually (\<lambda>w. w \<in> A) (nhds z)"
- using assms by (intro eventually_nhds_in_open) auto
- thus "eventually (\<lambda>w. w = exp (f w)) (nhds z)"
- by eventually_elim (use assms in auto)
- qed auto
+ by (smt (verit, best) assms has_field_derivative_transform_within_open)
finally have "((\<lambda>w. exp (f w)) has_field_derivative 1) (at z)" .
moreover have "((\<lambda>w. exp (f w)) has_field_derivative exp (f z) * deriv f z) (at z)"
by (rule derivative_eq_intros refl)+
@@ -133,11 +128,7 @@
theorem exp_Euler: "exp(\<i> * z) = cos(z) + \<i> * sin(z)"
proof -
have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) = (\<lambda>n. (\<i> * z) ^ n /\<^sub>R (fact n))"
- proof
- fix n
- show "(cos_coeff n + \<i> * sin_coeff n) * z^n = (\<i> * z) ^ n /\<^sub>R (fact n)"
- by (auto simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE)
- qed
+ by (force simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE)
also have "\<dots> sums (exp (\<i> * z))"
by (rule exp_converges)
finally have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) sums (exp (\<i> * z))" .
@@ -149,8 +140,7 @@
qed
corollary\<^marker>\<open>tag unimportant\<close> exp_minus_Euler: "exp(-(\<i> * z)) = cos(z) - \<i> * sin(z)"
- using exp_Euler [of "-z"]
- by simp
+ using exp_Euler [of "-z"] by simp
lemma sin_exp_eq: "sin z = (exp(\<i> * z) - exp(-(\<i> * z))) / (2*\<i>)"
by (simp add: exp_Euler exp_minus_Euler)
@@ -198,18 +188,7 @@
by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute)
lemma exp_cnj: "cnj (exp z) = exp (cnj z)"
-proof -
- have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) = (\<lambda>n. (cnj z)^n /\<^sub>R (fact n))"
- by auto
- also have "\<dots> sums (exp (cnj z))"
- by (rule exp_converges)
- finally have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) sums (exp (cnj z))" .
- moreover have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) sums (cnj (exp z))"
- by (metis exp_converges sums_cnj)
- ultimately show ?thesis
- using sums_unique2
- by blast
-qed
+ by (simp add: cis_cnj exp_eq_polar)
lemma cnj_sin: "cnj(sin z) = sin(cnj z)"
by (simp add: sin_exp_eq exp_cnj field_simps)
@@ -264,7 +243,7 @@
subsection\<^marker>\<open>tag unimportant\<close>\<open>More on the Polar Representation of Complex Numbers\<close>
lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)"
- by (simp add: Complex_eq exp_add exp_Euler exp_of_real sin_of_real cos_of_real)
+ using Complex_eq Euler complex.sel by presburger
lemma exp_eq_1: "exp z = 1 \<longleftrightarrow> Re(z) = 0 \<and> (\<exists>n::int. Im(z) = of_int (2 * n) * pi)"
(is "?lhs = ?rhs")
@@ -273,7 +252,7 @@
then have "Re z = 0"
by (metis exp_eq_one_iff norm_exp_eq_Re norm_one)
with \<open>?lhs\<close> show ?rhs
- by (metis Re_exp complex_Re_of_int cos_one_2pi_int exp_zero mult.commute mult_numeral_1 numeral_One of_int_mult of_int_numeral)
+ by (metis Re_exp cos_one_2pi_int exp_zero mult.commute mult_1 of_int_mult of_int_numeral one_complex.simps(1))
next
assume ?rhs then show ?lhs
using Im_exp Re_exp complex_eq_iff
@@ -298,13 +277,7 @@
lemma exp_integer_2pi:
assumes "n \<in> \<int>"
shows "exp((2 * n * pi) * \<i>) = 1"
-proof -
- have "exp((2 * n * pi) * \<i>) = exp 0"
- using assms unfolding Ints_def exp_eq by auto
- also have "\<dots> = 1"
- by simp
- finally show ?thesis .
-qed
+ by (metis assms cis_conv_exp cis_multiple_2pi mult.assoc mult.commute)
lemma exp_plus_2pin [simp]: "exp (z + \<i> * (of_int n * (of_real pi * 2))) = exp z"
by (simp add: exp_eq)
@@ -312,15 +285,8 @@
lemma exp_integer_2pi_plus1:
assumes "n \<in> \<int>"
shows "exp(((2 * n + 1) * pi) * \<i>) = - 1"
-proof -
- from assms obtain n' where [simp]: "n = of_int n'"
- by (auto simp: Ints_def)
- have "exp(((2 * n + 1) * pi) * \<i>) = exp (pi * \<i>)"
- using assms by (subst exp_eq) (auto intro!: exI[of _ n'] simp: algebra_simps)
- also have "\<dots> = - 1"
- by simp
- finally show ?thesis .
-qed
+ using exp_integer_2pi [OF assms]
+ by (metis cis_conv_exp cis_mult cis_pi distrib_left mult.commute mult.right_neutral)
lemma inj_on_exp_pi:
fixes z::complex shows "inj_on exp (ball z pi)"
@@ -338,7 +304,6 @@
lemma cmod_add_squared:
fixes r1 r2::real
- assumes "r1 \<ge> 0" "r2 \<ge> 0"
shows "(cmod (r1 * exp (\<i> * \<theta>1) + r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\<theta>1 - \<theta>2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs")
proof -
have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)"
@@ -355,19 +320,8 @@
lemma cmod_diff_squared:
fixes r1 r2::real
- assumes "r1 \<ge> 0" "r2 \<ge> 0"
- shows "(cmod (r1 * exp (\<i> * \<theta>1) - r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\<theta>1 - \<theta>2)" (is "(cmod (?z1 - ?z2))\<^sup>2 = ?rhs")
-proof -
- have "exp (\<i> * (\<theta>2 + pi)) = - exp (\<i> * \<theta>2)"
- by (simp add: exp_Euler cos_plus_pi sin_plus_pi)
- then have "(cmod (?z1 - ?z2))\<^sup>2 = cmod (?z1 + r2 * exp (\<i> * (\<theta>2 + pi))) ^2"
- by simp
- also have "\<dots> = r1\<^sup>2 + r2\<^sup>2 + 2*r1*r2*cos (\<theta>1 - (\<theta>2 + pi))"
- using assms cmod_add_squared by blast
- also have "\<dots> = ?rhs"
- by (simp add: add.commute diff_add_eq_diff_diff_swap)
- finally show ?thesis .
-qed
+ shows "(cmod (r1 * exp (\<i> * \<theta>1) - r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\<theta>1 - \<theta>2)"
+ using cmod_add_squared [of r1 _ "-r2"] by simp
lemma polar_convergence:
fixes R::real
@@ -381,8 +335,7 @@
moreover obtain k where "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>"
proof -
have "cos (\<theta> j - \<Theta>) = ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)" for j
- apply (subst cmod_diff_squared)
- using assms by (auto simp: field_split_simps less_le)
+ using assms by (auto simp: cmod_diff_squared less_le)
moreover have "(\<lambda>j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \<longlonglongrightarrow> ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))"
by (intro L rR tendsto_intros) (use \<open>R > 0\<close> in force)
moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1"
@@ -430,19 +383,7 @@
lemma exp_i_ne_1:
assumes "0 < x" "x < 2*pi"
shows "exp(\<i> * of_real x) \<noteq> 1"
-proof
- assume "exp (\<i> * of_real x) = 1"
- then have "exp (\<i> * of_real x) = exp 0"
- by simp
- then obtain n where "\<i> * of_real x = (of_int (2 * n) * pi) * \<i>"
- by (simp only: Ints_def exp_eq) auto
- then have "of_real x = (of_int (2 * n) * pi)"
- by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real)
- then have "x = (of_int (2 * n) * pi)"
- by simp
- then show False using assms
- by (cases n) (auto simp: zero_less_mult_iff mult_less_0_iff)
-qed
+ by (smt (verit) Im_i_times Re_complex_of_real assms exp_complex_eqI exp_zero zero_complex.sel(2))
lemma sin_eq_0:
fixes z::complex
@@ -458,16 +399,7 @@
lemma cos_eq_1:
fixes z::complex
shows "cos z = 1 \<longleftrightarrow> (\<exists>n::int. z = complex_of_real(2 * n * pi))"
-proof -
- have "cos z = cos (2*(z/2))"
- by simp
- also have "\<dots> = 1 - 2 * sin (z/2) ^ 2"
- by (simp only: cos_double_sin)
- finally have [simp]: "cos z = 1 \<longleftrightarrow> sin (z/2) = 0"
- by simp
- show ?thesis
- by (auto simp: sin_eq_0)
-qed
+ by (metis Re_complex_of_real cos_of_real cos_one_2pi_int cos_one_sin_zero mult.commute of_real_1 sin_eq_0)
lemma csin_eq_1:
fixes z::complex
@@ -482,10 +414,8 @@
proof -
have "sin z = -1 \<longleftrightarrow> sin (-z) = 1"
by (simp add: equation_minus_iff)
- also have "\<dots> \<longleftrightarrow> (\<exists>n::int. -z = of_real(2 * n * pi) + of_real pi/2)"
- by (simp only: csin_eq_1)
also have "\<dots> \<longleftrightarrow> (\<exists>n::int. z = - of_real(2 * n * pi) - of_real pi/2)"
- by (rule iff_exI) (metis add.inverse_inverse add_uminus_conv_diff minus_add_distrib)
+ by (metis (mono_tags, lifting) add_uminus_conv_diff csin_eq_1 equation_minus_iff minus_add_distrib)
also have "\<dots> = ?rhs"
apply safe
apply (rule_tac [2] x="-(x+1)" in exI)
@@ -506,10 +436,8 @@
proof -
have "sin x = 1 \<longleftrightarrow> sin (complex_of_real x) = 1"
by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real)
- also have "\<dots> \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)"
- by (simp only: csin_eq_1)
also have "\<dots> \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + of_real pi/2)"
- by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
+ by (metis csin_eq_1 Re_complex_of_real id_apply of_real_add of_real_divide of_real_eq_id of_real_numeral)
also have "\<dots> = ?rhs"
by (auto simp: algebra_simps)
finally show ?thesis .
@@ -519,10 +447,8 @@
proof -
have "sin x = -1 \<longleftrightarrow> sin (complex_of_real x) = -1"
by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real)
- also have "\<dots> \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)"
- by (simp add: csin_eq_minus1)
also have "\<dots> \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + 3/2*pi)"
- by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
+ by (metis Re_complex_of_real csin_eq_minus1 id_apply of_real_add of_real_eq_id)
also have "\<dots> = ?rhs"
by (auto simp: algebra_simps)
finally show ?thesis .
@@ -533,10 +459,8 @@
proof -
have "cos x = -1 \<longleftrightarrow> cos (complex_of_real x) = -1"
by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real)
- also have "\<dots> \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + pi)"
- by (simp add: ccos_eq_minus1)
also have "\<dots> \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + pi)"
- by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
+ by (metis ccos_eq_minus1 id_apply of_real_add of_real_eq_id of_real_eq_iff)
also have "\<dots> = ?rhs"
by (auto simp: algebra_simps)
finally show ?thesis .
@@ -568,12 +492,8 @@
(is "?lhs = ?rhs")
proof
assume ?lhs
- then have "sin w - sin z = 0"
- by (auto simp: algebra_simps)
- then have "sin ((w - z) / 2)*cos ((w + z) / 2) = 0"
- by (auto simp: sin_diff_sin)
then consider "sin ((w - z) / 2) = 0" | "cos ((w + z) / 2) = 0"
- using mult_eq_0_iff by blast
+ by (metis divide_eq_0_iff nonzero_eq_divide_eq right_minus_eq sin_diff_sin zero_neq_numeral)
then show ?rhs
proof cases
case 1
@@ -608,14 +528,10 @@
fixes w :: complex
shows "cos w = cos z \<longleftrightarrow> (\<exists>n \<in> \<int>. w = z + of_real(2*n*pi) \<or> w = -z + of_real(2*n*pi))"
(is "?lhs = ?rhs")
-proof
+proof
assume ?lhs
- then have "cos w - cos z = 0"
- by (auto simp: algebra_simps)
- then have "sin ((w + z) / 2) * sin ((z - w) / 2) = 0"
- by (auto simp: cos_diff_cos)
then consider "sin ((w + z) / 2) = 0" | "sin ((z - w) / 2) = 0"
- using mult_eq_0_iff by blast
+ by (metis mult_eq_0_iff cos_diff_cos right_minus_eq zero_neq_numeral)
then show ?rhs
proof cases
case 1
@@ -719,22 +635,7 @@
lemma norm_cos_plus1_le:
fixes z::complex
shows "norm(1 + cos z) \<le> 2 * exp(norm z)"
-proof -
- have mono: "\<And>u w z::real. (1 \<le> w | 1 \<le> z) \<Longrightarrow> (w \<le> u & z \<le> u) \<Longrightarrow> 2 + w + z \<le> 4 * u"
- by arith
- have *: "Im z \<le> cmod z"
- using abs_Im_le_cmod abs_le_D1 by auto
- have triangle3: "\<And>x y z. norm(x + y + z) \<le> norm(x) + norm(y) + norm(z)"
- by (simp add: norm_add_rule_thm)
- have "norm(1 + cos z) = cmod (1 + (exp (\<i> * z) + exp (- (\<i> * z))) / 2)"
- by (simp add: cos_exp_eq)
- also have "\<dots> = cmod ((2 + exp (\<i> * z) + exp (- (\<i> * z))) / 2)"
- by (simp add: field_simps)
- also have "\<dots> = cmod (2 + exp (\<i> * z) + exp (- (\<i> * z))) / 2"
- by (simp add: norm_divide)
- finally show ?thesis
- by (metis exp_eq_one_iff exp_le_cancel_iff mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono)
-qed
+ by (smt (verit, best) exp_ge_add_one_self norm_cos_le norm_ge_zero norm_one norm_triangle_ineq)
lemma sinh_conv_sin: "sinh z = -\<i> * sin (\<i>*z)"
by (simp add: sinh_field_def sin_i_times exp_minus)
@@ -793,6 +694,7 @@
by simp
qed auto
+text \<open>For complex @{term z}, a tighter bound than in the previous result\<close>
lemma Taylor_exp:
"norm(exp z - (\<Sum>k\<le>n. z ^ k / (fact k))) \<le> exp\<bar>Re z\<bar> * (norm z) ^ (Suc n) / (fact n)"
proof (rule complex_Taylor [of _ n "\<lambda>k. exp" "exp\<bar>Re z\<bar>" 0 z, simplified])
@@ -857,14 +759,12 @@
\<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact n)"
proof (rule complex_Taylor [of "closed_segment 0 z" n
"\<lambda>k x. (-1)^(k div 2) * (if even k then sin x else cos x)"
- "exp\<bar>Im z\<bar>" 0 z, simplified])
+ "exp\<bar>Im z\<bar>" 0 z, simplified])
fix k x
show "((\<lambda>x. (- 1) ^ (k div 2) * (if even k then sin x else cos x)) has_field_derivative
(- 1) ^ (Suc k div 2) * (if odd k then sin x else cos x))
(at x within closed_segment 0 z)"
- apply (auto simp: power_Suc)
- apply (intro derivative_eq_intros | simp)+
- done
+ by (cases "even k") (intro derivative_eq_intros | simp add: power_Suc)+
next
fix x
assume "x \<in> closed_segment 0 z"
@@ -887,16 +787,15 @@
have *: "cmod (cos z -
(\<Sum>i\<le>n. (-1) ^ (Suc i div 2) * (if even i then cos 0 else sin 0) * z ^ i / (fact i)))
\<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact n)"
- proof (rule complex_Taylor [of "closed_segment 0 z" n "\<lambda>k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" "exp\<bar>Im z\<bar>" 0 z,
-simplified])
+ proof (rule complex_Taylor [of "closed_segment 0 z" n
+ "\<lambda>k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)"
+ "exp\<bar>Im z\<bar>" 0 z, simplified])
fix k x
assume "x \<in> closed_segment 0 z" "k \<le> n"
show "((\<lambda>x. (- 1) ^ (Suc k div 2) * (if even k then cos x else sin x)) has_field_derivative
(- 1) ^ Suc (k div 2) * (if odd k then cos x else sin x))
(at x within closed_segment 0 z)"
- apply (auto simp: power_Suc)
- apply (intro derivative_eq_intros | simp)+
- done
+ by (cases "even k") (intro derivative_eq_intros | simp add: power_Suc)+
next
fix x
assume "x \<in> closed_segment 0 z"
@@ -942,18 +841,10 @@
by (simp add: algebra_simps is_Arg_def)
lemma is_Arg_eqI:
- assumes r: "is_Arg z r" and s: "is_Arg z s" and rs: "abs (r-s) < 2*pi" and "z \<noteq> 0"
+ assumes "is_Arg z r" and "is_Arg z s" and "abs (r-s) < 2*pi" and "z \<noteq> 0"
shows "r = s"
-proof -
- have zr: "z = (cmod z) * exp (\<i> * r)" and zs: "z = (cmod z) * exp (\<i> * s)"
- using r s by (auto simp: is_Arg_def)
- with \<open>z \<noteq> 0\<close> have eq: "exp (\<i> * r) = exp (\<i> * s)"
- by (metis mult_eq_0_iff mult_left_cancel)
- have "\<i> * r = \<i> * s"
- by (rule exp_complex_eqI) (use rs in \<open>auto simp: eq exp_complex_eqI\<close>)
- then show ?thesis
- by simp
-qed
+ using assms unfolding is_Arg_def
+ by (metis Im_i_times Re_complex_of_real exp_complex_eqI mult_cancel_left mult_eq_0_iff)
text\<open>This function returns the angle of a complex number from its representation in polar coordinates.
Due to periodicity, its range is arbitrary. \<^term>\<open>Arg2pi\<close> follows HOL Light in adopting the interval \<open>[0,2\<pi>)\<close>.
@@ -965,28 +856,13 @@
by (simp add: Arg2pi_def)
lemma Arg2pi_unique_lemma:
- assumes z: "is_Arg z t"
- and z': "is_Arg z t'"
- and t: "0 \<le> t" "t < 2*pi"
- and t': "0 \<le> t'" "t' < 2*pi"
- and nz: "z \<noteq> 0"
+ assumes "is_Arg z t"
+ and "is_Arg z t'"
+ and "0 \<le> t" "t < 2*pi"
+ and "0 \<le> t'" "t' < 2*pi"
+ and "z \<noteq> 0"
shows "t' = t"
-proof -
- have [dest]: "\<And>x y z::real. x\<ge>0 \<Longrightarrow> x+y < z \<Longrightarrow> y<z"
- by arith
- have "of_real (cmod z) * exp (\<i> * of_real t') = of_real (cmod z) * exp (\<i> * of_real t)"
- by (metis z z' is_Arg_def)
- then have "exp (\<i> * of_real t') = exp (\<i> * of_real t)"
- by (metis nz mult_left_cancel mult_zero_left z is_Arg_def)
- then have "sin t' = sin t \<and> cos t' = cos t"
- by (metis cis.simps cis_conv_exp)
- then obtain n::int where n: "t' = t + 2 * n * pi"
- by (auto simp: sin_cos_eq_iff)
- then have "n=0"
- by (cases n) (use t t' in \<open>auto simp: mult_less_0_iff algebra_simps\<close>)
- then show "t' = t"
- by (simp add: n)
-qed
+ using is_Arg_eqI assms by force
lemma Arg2pi: "0 \<le> Arg2pi z \<and> Arg2pi z < 2*pi \<and> is_Arg z (Arg2pi z)"
proof (cases "z=0")
@@ -998,7 +874,7 @@
and ReIm: "Re z / cmod z = cos t" "Im z / cmod z = sin t"
using sincos_total_2pi [OF complex_unit_circle [OF False]]
by blast
- have z: "is_Arg z t"
+ then have z: "is_Arg z t"
unfolding is_Arg_def
using t False ReIm
by (intro complex_eqI) (auto simp: exp_Euler sin_of_real cos_of_real field_split_simps)
@@ -1056,36 +932,20 @@
qed auto
lemma Arg2pi_lt_pi: "0 < Arg2pi z \<and> Arg2pi z < pi \<longleftrightarrow> 0 < Im z"
-proof (cases "z=0")
- case False
- have "0 < Im z \<longleftrightarrow> 0 < Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg2pi z)))"
- by (metis Arg2pi_eq)
- also have "\<dots> = (0 < Im (exp (\<i> * complex_of_real (Arg2pi z))))"
- using False by (simp add: zero_less_mult_iff)
- also have "\<dots> \<longleftrightarrow> 0 < Arg2pi z \<and> Arg2pi z < pi" (is "_ = ?rhs")
- proof -
- have "0 < sin (Arg2pi z) \<Longrightarrow> ?rhs"
- by (meson Arg2pi_ge_0 Arg2pi_lt_2pi less_le_trans not_le sin_le_zero sin_x_le_x)
- then show ?thesis
- by (auto simp: Im_exp sin_gt_zero)
- qed
- finally show ?thesis
- by blast
-qed auto
+ using Arg2pi_le_pi [of z]
+ by (smt (verit, del_insts) Arg2pi_0 Arg2pi_le_pi Arg2pi_minus uminus_complex.simps(2) zero_complex.simps(2))
lemma Arg2pi_eq_0: "Arg2pi z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
proof (cases "z=0")
case False
- have "z \<in> \<real> \<and> 0 \<le> Re z \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (of_real (cmod z) * exp (\<i> * complex_of_real (Arg2pi z)))"
- by (metis Arg2pi_eq)
- also have "\<dots> \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg2pi z)))"
- using False by (simp add: zero_le_mult_iff)
+ then have "z \<in> \<real> \<and> 0 \<le> Re z \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg2pi z)))"
+ by (metis cis.sel(1) cis_conv_exp cos_Arg2pi norm_ge_zero norm_le_zero_iff zero_le_mult_iff)
also have "\<dots> \<longleftrightarrow> Arg2pi z = 0"
proof -
have [simp]: "Arg2pi z = 0 \<Longrightarrow> z \<in> \<real>"
using Arg2pi_eq [of z] by (auto simp: Reals_def)
moreover have "\<lbrakk>z \<in> \<real>; 0 \<le> cos (Arg2pi z)\<rbrakk> \<Longrightarrow> Arg2pi z = 0"
- by (metis Arg2pi_lt_pi Arg2pi_ge_0 Arg2pi_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl)
+ by (smt (verit, ccfv_SIG) Arg2pi_ge_0 Arg2pi_le_pi Arg2pi_lt_pi complex_is_Real_iff cos_pi)
ultimately show ?thesis
by (auto simp: Re_exp)
qed
@@ -1123,7 +983,7 @@
lemma Arg2pi_eq_iff:
assumes "w \<noteq> 0" "z \<noteq> 0"
- shows "Arg2pi w = Arg2pi z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)" (is "?lhs = ?rhs")
+ shows "Arg2pi w = Arg2pi z \<longleftrightarrow> (\<exists>x. 0 < x \<and> w = of_real x * z)" (is "?lhs = ?rhs")
proof
assume ?lhs
then have "(cmod w) * (z / cmod z) = w"
@@ -1169,8 +1029,7 @@
assumes "w \<noteq> 0" "z \<noteq> 0"
shows "Arg2pi (w * z) = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi w + Arg2pi z
else (Arg2pi w + Arg2pi z) - 2*pi)"
- using Arg2pi_add [OF assms]
- by auto
+ using Arg2pi_add [OF assms] by auto
lemma Arg2pi_cnj_eq_inverse:
assumes "z \<noteq> 0" shows "Arg2pi (cnj z) = Arg2pi (inverse z)"
@@ -1264,14 +1123,7 @@
lemma Ln_of_real:
assumes "0 < z"
shows "ln(of_real z::complex) = of_real(ln z)"
-proof -
- have "ln(of_real (exp (ln z))::complex) = ln (exp (of_real (ln z)))"
- by (simp add: exp_of_real)
- also have "\<dots> = of_real(ln z)"
- using assms by (subst Ln_exp) auto
- finally show ?thesis
- using assms by simp
-qed
+ by (smt (verit) Im_complex_of_real Ln_exp assms exp_ln of_real_exp pi_ge_two)
corollary\<^marker>\<open>tag unimportant\<close> Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<real>"
by (auto simp: Ln_of_real elim: Reals_cases)
@@ -1286,16 +1138,10 @@
using Ln_of_real by force
lemma Ln_1 [simp]: "ln 1 = (0::complex)"
-proof -
- have "ln (exp 0) = (0::complex)"
- by (simp add: del: exp_zero)
- then show ?thesis
- by simp
-qed
-
+ by (smt (verit, best) Ln_of_real ln_one of_real_0 of_real_1)
lemma Ln_eq_zero_iff [simp]: "x \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln x = 0 \<longleftrightarrow> x = 1" for x::complex
- by auto (metis exp_Ln exp_zero nonpos_Reals_zero_I)
+ by (metis (mono_tags, lifting) Ln_1 exp_Ln exp_zero nonpos_Reals_zero_I)
instance
by intro_classes (rule ln_complex_def Ln_1)
@@ -1317,17 +1163,12 @@
corollary\<^marker>\<open>tag unimportant\<close> ln_cmod_le:
assumes z: "z \<noteq> 0"
shows "ln (cmod z) \<le> cmod (Ln z)"
- using norm_exp [of "Ln z", simplified exp_Ln [OF z]]
by (metis Re_Ln complex_Re_le_cmod z)
proposition\<^marker>\<open>tag unimportant\<close> exists_complex_root:
fixes z :: complex
assumes "n \<noteq> 0" obtains w where "z = w ^ n"
-proof (cases "z=0")
- case False
- then show ?thesis
- by (rule_tac w = "exp(Ln z / n)" in that) (simp add: assms exp_of_nat_mult [symmetric])
-qed (use assms in auto)
+ by (metis assms exp_Ln exp_of_nat_mult nonzero_mult_div_cancel_left of_nat_eq_0_iff power_0_left times_divide_eq_right)
corollary\<^marker>\<open>tag unimportant\<close> exists_complex_root_nonzero:
fixes z::complex
@@ -1337,17 +1178,23 @@
subsection\<^marker>\<open>tag unimportant\<close>\<open>Derivative of Ln away from the branch cut\<close>
-lemma
+lemma Im_Ln_less_pi:
+ assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"shows "Im (Ln z) < pi"
+proof -
+ have znz [simp]: "z \<noteq> 0"
+ using assms by auto
+ with Im_Ln_le_pi [of z] show ?thesis
+ by (smt (verit, best) Arg2pi_eq_0_pi Arg2pi_exp Ln_in_Reals assms complex_is_Real_iff complex_nonpos_Reals_iff exp_Ln pi_ge_two)
+qed
+
+lemma has_field_derivative_Ln:
assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
- shows has_field_derivative_Ln: "(Ln has_field_derivative inverse(z)) (at z)"
- and Im_Ln_less_pi: "Im (Ln z) < pi"
+ shows "(Ln has_field_derivative inverse(z)) (at z)"
proof -
have znz [simp]: "z \<noteq> 0"
using assms by auto
then have "Im (Ln z) \<noteq> pi"
- by (metis (no_types) Im_exp Ln_in_Reals assms complex_nonpos_Reals_iff complex_is_Real_iff exp_Ln mult_zero_right not_less pi_neq_zero sin_pi znz)
- then show *: "Im (Ln z) < pi" using assms Im_Ln_le_pi
- by (simp add: le_neq_trans)
+ by (smt (verit, best) Arg2pi_eq_0_pi Arg2pi_exp Ln_in_Reals assms complex_is_Real_iff complex_nonpos_Reals_iff exp_Ln pi_ge_two)
let ?U = "{w. -pi < Im(w) \<and> Im(w) < pi}"
have 1: "open ?U"
by (simp add: open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt)
@@ -1356,7 +1203,7 @@
have 3: "continuous_on ?U (\<lambda>x. Blinfun ((*) (exp x)))"
unfolding blinfun_mult_right.abs_eq [symmetric] by (intro continuous_intros)
have 4: "Ln z \<in> ?U"
- by (auto simp: mpi_less_Im_Ln *)
+ by (simp add: Im_Ln_less_pi assms mpi_less_Im_Ln)
have 5: "Blinfun ((*) (inverse z)) o\<^sub>L Blinfun ((*) (exp (Ln z))) = id_blinfun"
by (rule blinfun_eqI) (simp add: bounded_linear_mult_right bounded_linear_Blinfun_apply)
obtain U' V g g' where "open U'" and sub: "U' \<subseteq> ?U"
@@ -1371,21 +1218,15 @@
unfolding has_field_derivative_def
proof (rule has_derivative_transform_within_open)
show g_eq_Ln: "g y = Ln y" if "y \<in> V" for y
- proof -
- obtain x where "y = exp x" "x \<in> U'"
- using hom homeomorphism_image1 that \<open>y \<in> V\<close> by blast
- then show ?thesis
- using sub hom homeomorphism_apply1 by fastforce
- qed
+ by (smt (verit, ccfv_threshold) Ln_exp hom homeomorphism_def imageI mem_Collect_eq sub subset_iff that)
have "0 \<notin> V"
by (meson exp_not_eq_zero hom homeomorphism_def)
then have "\<And>y. y \<in> V \<Longrightarrow> g' y = inv ((*) y)"
by (metis exp_Ln g' g_eq_Ln)
then have g': "g' z = (\<lambda>x. x/z)"
- by (metis (no_types, opaque_lifting) bij \<open>z \<in> V\<close> bij_inv_eq_iff exp_Ln g_eq_Ln nonzero_mult_div_cancel_left znz)
+ by (metis \<open>z \<in> V\<close> bij bij_inv_eq_iff exp_Ln g_eq_Ln nonzero_mult_div_cancel_left znz)
show "(g has_derivative (*) (inverse z)) (at z)"
- using g [OF \<open>z \<in> V\<close>] g'
- by (simp add: \<open>z \<in> V\<close> field_class.field_divide_inverse has_derivative_imp_has_field_derivative has_field_derivative_imp_has_derivative)
+ using g [OF \<open>z \<in> V\<close>] g' by (simp add: divide_inverse_commute)
qed (auto simp: \<open>z \<in> V\<close> \<open>open V\<close>)
qed
@@ -1425,29 +1266,19 @@
by (auto simp: o_def)
lemma tendsto_Ln [tendsto_intros]:
- fixes L F
assumes "(f \<longlongrightarrow> L) F" "L \<notin> \<real>\<^sub>\<le>\<^sub>0"
shows "((\<lambda>x. Ln (f x)) \<longlongrightarrow> Ln L) F"
-proof -
- have "nhds L \<ge> filtermap f F"
- using assms(1) by (simp add: filterlim_def)
- moreover have "\<forall>\<^sub>F y in nhds L. y \<in> - \<real>\<^sub>\<le>\<^sub>0"
- using eventually_nhds_in_open[of "- \<real>\<^sub>\<le>\<^sub>0" L] assms by (auto simp: open_Compl)
- ultimately have "\<forall>\<^sub>F y in filtermap f F. y \<in> - \<real>\<^sub>\<le>\<^sub>0" by (rule filter_leD)
- moreover have "continuous_on (-\<real>\<^sub>\<le>\<^sub>0) Ln" by (rule continuous_on_Ln) auto
- ultimately show ?thesis using continuous_on_tendsto_compose[of "- \<real>\<^sub>\<le>\<^sub>0" Ln f L F] assms
- by (simp add: eventually_filtermap)
-qed
+ by (simp add: assms isCont_tendsto_compose)
lemma divide_ln_mono:
fixes x y::real
assumes "3 \<le> x" "x \<le> y"
shows "x / ln x \<le> y / ln y"
-proof (rule exE [OF complex_mvt_line [of x y "\<lambda>z. z / Ln z" "\<lambda>z. 1/(Ln z) - 1/(Ln z)^2"]];
- clarsimp simp add: closed_segment_Reals closed_segment_eq_real_ivl assms)
- show "\<And>u. \<lbrakk>x \<le> u; u \<le> y\<rbrakk> \<Longrightarrow> ((\<lambda>z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)"
+proof -
+ have "\<And>u. \<lbrakk>x \<le> u; u \<le> y\<rbrakk> \<Longrightarrow> ((\<lambda>z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)"
using \<open>3 \<le> x\<close> by (force intro!: derivative_eq_intros simp: field_simps power_eq_if)
- show "x / ln x \<le> y / ln y"
+ moreover
+ have "x / ln x \<le> y / ln y"
if "Re (y / Ln y) - Re (x / Ln x) = (Re (1 / Ln u) - Re (1 / (Ln u)\<^sup>2)) * (y - x)"
and x: "x \<le> u" "u \<le> y" for u
proof -
@@ -1456,6 +1287,9 @@
show ?thesis
using exp_le \<open>3 \<le> x\<close> x by (simp add: eq) (simp add: power_eq_if divide_simps ln_ge_iff)
qed
+ ultimately show ?thesis
+ using complex_mvt_line [of x y "\<lambda>z. z / Ln z" "\<lambda>z. 1/(Ln z) - 1/(Ln z)^2"] assms
+ by (force simp add: closed_segment_Reals closed_segment_eq_real_ivl)
qed
theorem Ln_series:
@@ -1505,7 +1339,8 @@
by (drule Ln_series) (simp add: power_minus')
lemma ln_series':
- assumes "abs (x::real) < 1"
+ fixes x::real
+ assumes "\<bar>x\<bar> < 1"
shows "(\<lambda>n. - ((-x)^n) / of_nat n) sums ln (1 + x)"
proof -
from assms have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)"
@@ -1513,7 +1348,7 @@
also have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) = (\<lambda>n. complex_of_real (- ((-x)^n) / of_nat n))"
by (rule ext) simp
also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))"
- by (subst Ln_of_real [symmetric]) simp_all
+ by (smt (verit) Ln_of_real of_real_1 of_real_add)
finally show ?thesis by (subst (asm) sums_of_real_iff)
qed
@@ -1604,35 +1439,9 @@
using cos_minus_pi cos_gt_zero_pi [of "x-pi"]
by simp
-lemma Re_Ln_pos_lt:
- assumes "z \<noteq> 0"
- shows "\<bar>Im(Ln z)\<bar> < pi/2 \<longleftrightarrow> 0 < Re(z)"
-proof -
- { fix w
- assume "w = Ln z"
- then have w: "Im w \<le> pi" "- pi < Im w"
- using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms
- by auto
- have "\<bar>Im w\<bar> < pi/2 \<longleftrightarrow> 0 < Re(exp w)"
- proof
- assume "\<bar>Im w\<bar> < pi/2" then show "0 < Re(exp w)"
- by (auto simp: Re_exp cos_gt_zero_pi split: if_split_asm)
- next
- assume R: "0 < Re(exp w)" then
- have "\<bar>Im w\<bar> \<noteq> pi/2"
- by (metis cos_minus cos_pi_half mult_eq_0_iff Re_exp abs_if order_less_irrefl)
- then show "\<bar>Im w\<bar> < pi/2"
- using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"] w R
- by (force simp: Re_exp zero_less_mult_iff abs_if not_less_iff_gr_or_eq)
- qed
- }
- then show ?thesis using assms
- by auto
-qed
-
lemma Re_Ln_pos_le:
assumes "z \<noteq> 0"
- shows "\<bar>Im(Ln z)\<bar> \<le> pi/2 \<longleftrightarrow> 0 \<le> Re(z)"
+ shows "\<bar>Im(Ln z)\<bar> \<le> pi/2 \<longleftrightarrow> 0 \<le> Re(z)"
proof -
{ fix w
assume "w = Ln z"
@@ -1647,23 +1456,11 @@
by auto
qed
-lemma Im_Ln_pos_lt:
+lemma Re_Ln_pos_lt:
assumes "z \<noteq> 0"
- shows "0 < Im(Ln z) \<and> Im(Ln z) < pi \<longleftrightarrow> 0 < Im(z)"
-proof -
- { fix w
- assume "w = Ln z"
- then have w: "Im w \<le> pi" "- pi < Im w"
- using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms
- by auto
- then have "0 < Im w \<and> Im w < pi \<longleftrightarrow> 0 < Im(exp w)"
- using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"] less_linear
- by (fastforce simp add: Im_exp zero_less_mult_iff)
- }
- then show ?thesis using assms
- by auto
-qed
-
+ shows "\<bar>Im(Ln z)\<bar> < pi/2 \<longleftrightarrow> 0 < Re(z)"
+ using Re_Ln_pos_le assms
+ by (smt (verit) Re_exp arccos_cos cos_minus cos_pi_half exp_Ln exp_gt_zero field_sum_of_halves mult_eq_0_iff)
lemma Im_Ln_pos_le:
assumes "z \<noteq> 0"
@@ -1681,6 +1478,12 @@
by auto
qed
+lemma Im_Ln_pos_lt:
+ assumes "z \<noteq> 0"
+ shows "0 < Im(Ln z) \<and> Im(Ln z) < pi \<longleftrightarrow> 0 < Im(z)"
+ using Im_Ln_pos_le [OF assms] assms
+ by (smt (verit, best) Arg2pi_exp Arg2pi_lt_pi exp_Ln)
+
lemma Re_Ln_pos_lt_imp: "0 < Re(z) \<Longrightarrow> \<bar>Im(Ln z)\<bar> < pi/2"
by (metis Re_Ln_pos_lt less_irrefl zero_complex.simps(1))
@@ -1701,23 +1504,7 @@
proof (cases "z=0")
case False
show ?thesis
- proof (rule exp_complex_eqI)
- have "\<bar>Im (cnj (Ln z)) - Im (Ln (cnj z))\<bar> \<le> \<bar>Im (cnj (Ln z))\<bar> + \<bar>Im (Ln (cnj z))\<bar>"
- by (rule abs_triangle_ineq4)
- also have "\<dots> < pi + pi"
- proof -
- have "\<bar>Im (cnj (Ln z))\<bar> < pi"
- by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln)
- moreover have "\<bar>Im (Ln (cnj z))\<bar> \<le> pi"
- by (meson abs_le_iff complex_cnj_zero_iff less_eq_real_def minus_less_iff False Im_Ln_le_pi mpi_less_Im_Ln)
- ultimately show ?thesis
- by simp
- qed
- finally show "\<bar>Im (cnj (Ln z)) - Im (Ln (cnj z))\<bar> < 2 * pi"
- by simp
- show "exp (cnj (Ln z)) = exp (Ln (cnj z))"
- by (metis False complex_cnj_zero_iff exp_Ln exp_cnj)
- qed
+ by (smt (verit) False Im_Ln_less_pi Ln_exp assms cnj.sel(2) exp_Ln exp_cnj mpi_less_Im_Ln)
qed (use assms in auto)
@@ -1725,23 +1512,7 @@
proof (cases "z=0")
case False
show ?thesis
- proof (rule exp_complex_eqI)
- have "\<bar>Im (Ln (inverse z)) - Im (- Ln z)\<bar> \<le> \<bar>Im (Ln (inverse z))\<bar> + \<bar>Im (- Ln z)\<bar>"
- by (rule abs_triangle_ineq4)
- also have "\<dots> < pi + pi"
- proof -
- have "\<bar>Im (Ln (inverse z))\<bar> < pi"
- by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln)
- moreover have "\<bar>Im (- Ln z)\<bar> \<le> pi"
- using False Im_Ln_le_pi mpi_less_Im_Ln by fastforce
- ultimately show ?thesis
- by simp
- qed
- finally show "\<bar>Im (Ln (inverse z)) - Im (- Ln z)\<bar> < 2 * pi"
- by simp
- show "exp (Ln (inverse z)) = exp (- Ln z)"
- by (simp add: False exp_minus)
- qed
+ by (smt (verit) False Im_Ln_less_pi Ln_exp assms exp_Ln exp_minus mpi_less_Im_Ln uminus_complex.sel(2))
qed (use assms in auto)
lemma Ln_minus1 [simp]: "Ln(-1) = \<i> * pi"
@@ -1756,12 +1527,7 @@
by simp
lemma Ln_minus_ii [simp]: "Ln(-\<i>) = - (\<i> * pi/2)"
-proof -
- have "Ln(-\<i>) = Ln(inverse \<i>)" by simp
- also have "\<dots> = - (Ln \<i>)" using Ln_inverse by blast
- also have "\<dots> = - (\<i> * pi/2)" by simp
- finally show ?thesis .
-qed
+ using Ln_inverse by fastforce
lemma Ln_times:
assumes "w \<noteq> 0" "z \<noteq> 0"
@@ -1792,10 +1558,9 @@
using Ln_Reals_eq Ln_times_of_real by fastforce
corollary\<^marker>\<open>tag unimportant\<close> Ln_divide_of_real:
- "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(z / of_real r) = Ln(z) - ln r"
-using Ln_times_of_real [of "inverse r" z]
-by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric]
- del: of_real_inverse)
+ "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(z / of_real r) = Ln(z) - ln r"
+ using Ln_times_of_real [of "inverse r" z]
+ by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse flip: of_real_inverse)
corollary\<^marker>\<open>tag unimportant\<close> Ln_prod:
fixes f :: "'a \<Rightarrow> complex"
@@ -1818,10 +1583,10 @@
assumes "z \<noteq> 0"
shows "Ln(-z) = (if Im(z) \<le> 0 \<and> \<not>(Re(z) < 0 \<and> Im(z) = 0)
then Ln(z) + \<i> * pi
- else Ln(z) - \<i> * pi)" (is "_ = ?rhs")
+ else Ln(z) - \<i> * pi)"
using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms
Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z]
- by (fastforce simp: exp_add exp_diff exp_Euler intro!: Ln_unique)
+ by (intro Ln_unique) (auto simp: exp_add exp_diff)
lemma Ln_inverse_if:
assumes "z \<noteq> 0"
@@ -1854,7 +1619,7 @@
by (simp add: Ln_times) auto
lemma Ln_of_nat [simp]: "0 < n \<Longrightarrow> Ln (of_nat n) = of_real (ln (of_nat n))"
- by (subst of_real_of_nat_eq[symmetric], subst Ln_of_real[symmetric]) simp_all
+ by (metis Ln_of_real of_nat_0_less_iff of_real_of_nat_eq)
lemma Ln_of_nat_over_of_nat:
assumes "m > 0" "n > 0"
@@ -1913,13 +1678,11 @@
have "f n x + 1 \<notin> \<real>\<^sub>\<le>\<^sub>0" if "x \<in> A" for n x
proof
assume "f n x + 1 \<in> \<real>\<^sub>\<le>\<^sub>0"
- then obtain t where t: "t \<le> 0" "f n x + 1 = of_real t"
- by (auto elim!: nonpos_Reals_cases)
- hence "f n x = of_real (t - 1)"
- by (simp add: algebra_simps)
- also have "norm \<dots> \<ge> 1"
+ then obtain t where t: "t \<le> 0" "f n x = of_real (t - 1)"
+ by (metis add_diff_cancel nonpos_Reals_cases of_real_1 of_real_diff)
+ moreover have "norm \<dots> \<ge> 1"
using t by (subst norm_of_real) auto
- finally show False
+ ultimately show False
using norm_f[of x n] that by auto
qed
thus "\<forall>\<^sub>F n in sequentially. continuous_on A (\<lambda>x. \<Sum>n<n. Ln (1 + f n x))"
@@ -1961,19 +1724,14 @@
have "exp (\<Sum>n<N. ln (1 + f n x)) = (\<Prod>n<N. exp (ln (1 + f n x)))"
by (simp add: exp_sum)
also have "\<dots> = (\<Prod>n<N. 1 + f n x)"
- proof (rule prod.cong)
- fix n assume "n \<in> {..<N}"
- have "f n x \<noteq> -1"
- using norm_f[of x n] x by auto
- thus "exp (ln (1 + f n x)) = 1 + f n x"
- by (simp add: add_eq_0_iff)
- qed auto
+ using norm_f[of x] x
+ by (smt (verit, best) add.right_neutral add_diff_cancel exp_Ln norm_minus_commute norm_one prod.cong)
finally show "exp (\<Sum>n<N. ln (1 + f n x)) = (\<Prod>n<N. 1 + f n x)" .
qed
finally show ?thesis .
qed
-(* Theorem 17.6 by Bak & Newman, roughly *)
+text \<open>Theorem 17.6 by Bak and Newman, Complex Analysis [roughly]\<close>
lemma uniformly_convergent_on_prod:
fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex"
assumes cont: "\<And>n. continuous_on A (f n)"
@@ -2146,15 +1904,13 @@
lemma Arg_of_real [simp]: "Arg (of_real r) = (if r<0 then pi else 0)"
by (simp add: Im_Ln_eq_pi Arg_def)
-lemma mpi_less_Arg: "-pi < Arg z"
- and Arg_le_pi: "Arg z \<le> pi"
+lemma mpi_less_Arg: "-pi < Arg z" and Arg_le_pi: "Arg z \<le> pi"
by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi)
-lemma
+lemma Arg_eq:
assumes "z \<noteq> 0"
- shows Arg_eq: "z = of_real(norm z) * exp(\<i> * Arg z)"
- using assms exp_Ln exp_eq_polar
- by (auto simp: Arg_def)
+ shows "z = of_real(norm z) * exp(\<i> * Arg z)"
+ using cis_conv_exp rcis_cmod_Arg rcis_def by force
lemma is_Arg_Arg: "z \<noteq> 0 \<Longrightarrow> is_Arg z (Arg z)"
by (simp add: Arg_eq is_Arg_def)
@@ -2196,35 +1952,15 @@
by (auto simp: is_Arg_def norm_divide field_simps exp_diff Arg_of_real)
lemma Arg_unique_lemma:
- assumes z: "is_Arg z t"
- and z': "is_Arg z t'"
- and t: "- pi < t" "t \<le> pi"
- and t': "- pi < t'" "t' \<le> pi"
- and nz: "z \<noteq> 0"
+ assumes "is_Arg z t" "is_Arg z t'"
+ and "- pi < t" "t \<le> pi"
+ and "- pi < t'" "t' \<le> pi"
+ and "z \<noteq> 0"
shows "t' = t"
- using Arg2pi_unique_lemma [of "- (inverse z)"]
-proof -
- have "pi - t' = pi - t"
- proof (rule Arg2pi_unique_lemma [of "- (inverse z)"])
- have "- (inverse z) = - (inverse (of_real(norm z) * exp(\<i> * t)))"
- by (metis is_Arg_def z)
- also have "\<dots> = (cmod (- inverse z)) * exp (\<i> * (pi - t))"
- by (auto simp: field_simps exp_diff norm_divide)
- finally show "is_Arg (- inverse z) (pi - t)"
- unfolding is_Arg_def .
- have "- (inverse z) = - (inverse (of_real(norm z) * exp(\<i> * t')))"
- by (metis is_Arg_def z')
- also have "\<dots> = (cmod (- inverse z)) * exp (\<i> * (pi - t'))"
- by (auto simp: field_simps exp_diff norm_divide)
- finally show "is_Arg (- inverse z) (pi - t')"
- unfolding is_Arg_def .
- qed (use assms in auto)
- then show ?thesis
- by simp
-qed
+ using is_Arg_eqI assms by force
lemma complex_norm_eq_1_exp_eq: "norm z = 1 \<longleftrightarrow> exp(\<i> * (Arg z)) = z"
- by (metis Arg_eq exp_not_eq_zero exp_zero mult.left_neutral norm_zero of_real_1 norm_exp_i_times)
+ by (metis Arg2pi_eq Arg_eq complex_norm_eq_1_exp norm_eq_zero norm_exp_i_times)
lemma Arg_unique: "\<lbrakk>of_real r * exp(\<i> * a) = z; 0 < r; -pi < a; a \<le> pi\<rbrakk> \<Longrightarrow> Arg z = a"
by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq])
@@ -2239,9 +1975,8 @@
have [simp]: "cmod z * sin (Arg z) = Im z"
using assms Arg_eq [of z] by (metis Im_exp exp_Ln norm_exp_eq_Re Arg_def)
show ?thesis
- apply (rule Arg_unique [of "norm z", OF complex_eqI])
using mpi_less_Arg [of z] Arg_le_pi [of z] assms
- by (auto simp: Re_exp Im_exp)
+ by (intro Arg_unique [of "norm z", OF complex_eqI]) (auto simp: Re_exp Im_exp)
qed
lemma Arg_1 [simp]: "Arg 1 = 0"
@@ -2292,14 +2027,10 @@
lemma Arg_inverse: "Arg(inverse z) = (if z \<in> \<real> then Arg z else - Arg z)"
proof (cases "z \<in> \<real>")
- case True
- then show ?thesis
- by (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse)
-next
case False
then show ?thesis
by (simp add: Arg_def Ln_inverse complex_is_Real_iff complex_nonpos_Reals_iff)
-qed
+qed (use Arg_real Re_inverse in auto)
lemma Arg_eq_iff:
assumes "w \<noteq> 0" "z \<noteq> 0"
@@ -2437,28 +2168,7 @@
lemma Im_Ln_eq_pi_half:
"z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = pi/2 \<longleftrightarrow> 0 < Im(z) \<and> Re(z) = 0)"
"z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = -pi/2 \<longleftrightarrow> Im(z) < 0 \<and> Re(z) = 0)"
-proof -
- show "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = pi/2 \<longleftrightarrow> 0 < Im(z) \<and> Re(z) = 0)"
- by (metis Im_Ln_eq_pi Im_Ln_le_pi Im_Ln_pos_lt Re_Ln_pos_le Re_Ln_pos_lt
- abs_of_nonneg less_eq_real_def order_less_irrefl pi_half_gt_zero)
-next
- assume "z\<noteq>0"
- have "Im (Ln z) = - pi / 2 \<Longrightarrow> Im z < 0 \<and> Re z = 0"
- by (metis Im_Ln_pos_le Re_Ln_pos_le Re_Ln_pos_lt_imp \<open>z \<noteq> 0\<close> abs_if
- add.inverse_inverse divide_minus_left less_eq_real_def linorder_not_le minus_pi_half_less_zero)
- moreover have "Im (Ln z) = - pi / 2" when "Im z < 0" "Re z = 0"
- proof -
- obtain r::real where "r>0" "z=r * (-\<i>)"
- by (smt (verit) \<open>Im z < 0\<close> \<open>Re z = 0\<close> add_0 complex_eq mult.commute mult_minus_right of_real_0 of_real_minus)
- then have "Im (Ln z) = Im (Ln (r*(-\<i>)))" by auto
- also have "... = Im (Ln (complex_of_real r) + Ln (- \<i>))"
- by (metis Ln_times_of_real \<open>0 < r\<close> add.inverse_inverse add.inverse_neutral complex_i_not_zero)
- also have "... = - pi/2"
- using \<open>r>0\<close> by simp
- finally show "Im (Ln z) = - pi / 2" .
- qed
- ultimately show "(Im(Ln z) = -pi/2 \<longleftrightarrow> Im(z) < 0 \<and> Re(z) = 0)" by auto
-qed
+ using Im_Ln_pos_lt Im_Ln_pos_le Re_Ln_pos_le Re_Ln_pos_lt pi_ge_two by fastforce+
lemma Im_Ln_eq:
assumes "z\<noteq>0"
@@ -2472,32 +2182,8 @@
else
if Im z>0 then pi/2 else -pi/2)"
proof -
- have eq_arctan_pos:"Im (Ln z) = arctan (Im z/Re z)" when "Re z>0" for z
- proof -
- define wR where "wR \<equiv> Re (Ln z)"
- define \<theta> where "\<theta> \<equiv> arctan (Im z/Re z)"
- have "z\<noteq>0" using that by auto
- have "exp (Complex wR \<theta>) = z"
- proof (rule complex_eqI)
- have "Im (exp (Complex wR \<theta>)) =exp wR * sin \<theta> "
- unfolding Im_exp by simp
- also have "... = Im z"
- unfolding wR_def Re_Ln[OF \<open>z\<noteq>0\<close>] \<theta>_def using \<open>z\<noteq>0\<close> \<open>Re z>0\<close>
- by (auto simp add:sin_arctan divide_simps complex_neq_0 cmod_def real_sqrt_divide)
- finally show "Im (exp (Complex wR \<theta>)) = Im z" .
- next
- have "Re (exp (Complex wR \<theta>)) = exp wR * cos \<theta> "
- unfolding Re_exp by simp
- also have "... = Re z"
- by (metis Arg_eq_Im_Ln Re_exp \<open>z \<noteq> 0\<close> \<theta>_def arg_conv_arctan exp_Ln that wR_def)
- finally show "Re (exp (Complex wR \<theta>)) = Re z" .
- qed
- moreover have "-pi<\<theta>" "\<theta>\<le>pi"
- using arctan_lbound [of \<open>Im z / Re z\<close>] arctan_ubound [of \<open>Im z / Re z\<close>]
- by (simp_all add: \<theta>_def)
- ultimately have "Ln z = Complex wR \<theta>" using Ln_unique by auto
- then show ?thesis using that unfolding \<theta>_def by auto
- qed
+ have eq_arctan_pos: "Im (Ln z) = arctan (Im z/Re z)" when "Re z>0" for z
+ by (metis Arg_eq_Im_Ln arg_conv_arctan order_less_irrefl that zero_complex.simps(1))
have ?thesis when "Re z=0"
using Im_Ln_eq_pi_half[OF \<open>z\<noteq>0\<close>] that
using assms complex_eq_iff by auto
@@ -2532,12 +2218,12 @@
assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
shows "continuous (at z) Arg2pi"
proof -
- have *: "isCont (\<lambda>z. Im (Ln (- z)) + pi) z"
+ have "isCont (\<lambda>z. Im (Ln (- z)) + pi) z"
by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+
- consider "Re z < 0" | "Im z \<noteq> 0" using assms
+ moreover consider "Re z < 0" | "Im z \<noteq> 0" using assms
using complex_nonneg_Reals_iff not_le by blast
- then have "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg2pi z"
- using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within)
+ ultimately have "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg2pi z"
+ by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within)
then show ?thesis
unfolding continuous_at
by (metis (mono_tags, lifting) Arg2pi_Ln Arg2pi_gt_0 Compl_iff Lim_transform_within_open assms
@@ -2658,7 +2344,7 @@
lemma powr_complexpow [simp]:
fixes x::complex shows "x \<noteq> 0 \<Longrightarrow> x powr (of_nat n) = x^n"
- by (induct n) (auto simp: ac_simps powr_add)
+ by (simp add: powr_nat')
lemma powr_complexnumeral [simp]:
fixes x::complex shows "x powr (numeral n) = x ^ (numeral n)"
@@ -2715,18 +2401,19 @@
lemma
fixes w::complex
- shows Reals_powr [simp]: "\<lbrakk>w \<in> \<real>\<^sub>\<ge>\<^sub>0; z \<in> \<real>\<rbrakk> \<Longrightarrow> w powr z \<in> \<real>"
- and nonneg_Reals_powr [simp]: "\<lbrakk>w \<in> \<real>\<^sub>\<ge>\<^sub>0; z \<in> \<real>\<rbrakk> \<Longrightarrow> w powr z \<in> \<real>\<^sub>\<ge>\<^sub>0"
- by (auto simp: nonneg_Reals_def Reals_def powr_of_real)
+ assumes "w \<in> \<real>\<^sub>\<ge>\<^sub>0" "z \<in> \<real>"
+ shows Reals_powr [simp]: "w powr z \<in> \<real>" and nonneg_Reals_powr [simp]: "w powr z \<in> \<real>\<^sub>\<ge>\<^sub>0"
+ using assms by (auto simp: nonneg_Reals_def Reals_def powr_of_real)
lemma powr_neg_real_complex:
- "(- of_real x) powr a = (-1) powr (of_real (sgn x) * a) * of_real x powr (a :: complex)"
+ fixes w::complex
+ shows "(- of_real x) powr w = (-1) powr (of_real (sgn x) * w) * of_real x powr w"
proof (cases "x = 0")
assume x: "x \<noteq> 0"
- hence "(-x) powr a = exp (a * ln (-of_real x))" by (simp add: powr_def)
+ hence "(-x) powr w = exp (w * ln (-of_real x))" by (simp add: powr_def)
also from x have "ln (-of_real x) = Ln (of_real x) + of_real (sgn x) * pi * \<i>"
by (simp add: Ln_minus Ln_of_real)
- also from x have "exp (a * \<dots>) = cis pi powr (of_real (sgn x) * a) * of_real x powr a"
+ also from x have "exp (w * \<dots>) = cis pi powr (of_real (sgn x) * w) * of_real x powr w"
by (simp add: powr_def exp_add algebra_simps Ln_of_real cis_conv_exp)
also note cis_pi
finally show ?thesis by simp
@@ -2818,28 +2505,14 @@
lemma field_differentiable_powr_of_int:
fixes z :: complex
- assumes gderiv: "g field_differentiable (at z within S)" and "g z \<noteq> 0"
+ assumes "g field_differentiable (at z within S)" and "g z \<noteq> 0"
shows "(\<lambda>z. g z powr of_int n) field_differentiable (at z within S)"
-using has_field_derivative_powr_of_int assms(2) field_differentiable_def gderiv by blast
+ using has_field_derivative_powr_of_int assms field_differentiable_def by blast
lemma holomorphic_on_powr_of_int [holomorphic_intros]:
- assumes holf: "f holomorphic_on S" and 0: "\<And>z. z\<in>S \<Longrightarrow> f z \<noteq> 0"
+ assumes "f holomorphic_on S" and "\<And>z. z\<in>S \<Longrightarrow> f z \<noteq> 0"
shows "(\<lambda>z. (f z) powr of_int n) holomorphic_on S"
-proof (cases "n\<ge>0")
- case True
- then have "?thesis \<longleftrightarrow> (\<lambda>z. (f z) ^ nat n) holomorphic_on S"
- by (metis (no_types, lifting) 0 holomorphic_cong powr_of_int)
- moreover have "(\<lambda>z. (f z) ^ nat n) holomorphic_on S"
- using holf by (auto intro: holomorphic_intros)
- ultimately show ?thesis by auto
-next
- case False
- then have "?thesis \<longleftrightarrow> (\<lambda>z. inverse (f z) ^ nat (-n)) holomorphic_on S"
- by (metis (no_types, lifting) "0" holomorphic_cong power_inverse powr_of_int)
- moreover have "(\<lambda>z. inverse (f z) ^ nat (-n)) holomorphic_on S"
- using assms by (auto intro!:holomorphic_intros)
- ultimately show ?thesis by auto
-qed
+ using assms field_differentiable_powr_of_int holomorphic_on_def by auto
lemma has_field_derivative_powr_right [derivative_intros]:
"w \<noteq> 0 \<Longrightarrow> ((\<lambda>z. w powr z) has_field_derivative Ln w * w powr z) (at z)"
@@ -2896,7 +2569,7 @@
define h where
"h = (\<lambda>z. if f z = 0 then 0 else exp (Re (g z) * ln (cmod (f z)) + abs (Im (g z)) * pi))"
{
- fix z :: 'a assume z: "f z \<noteq> 0"
+ fix z :: 'a assume z: "f z \<noteq> 0"
define c where "c = abs (Im (g z)) * pi"
from mpi_less_Im_Ln[OF z] Im_Ln_le_pi[OF z]
have "abs (Im (Ln (f z))) \<le> pi" by simp
@@ -2905,7 +2578,8 @@
hence "-Im (g z) * Im (ln (f z)) \<le> c" by simp
hence "norm (f z powr g z) \<le> h z" by (simp add: powr_def field_simps h_def c_def)
}
- hence le: "norm (f z powr g z) \<le> h z" for z by (cases "f x = 0") (simp_all add: h_def)
+ hence le: "norm (f z powr g z) \<le> h z" for z
+ by (simp add: h_def)
have g': "(g \<longlongrightarrow> b) (inf F (principal {z. f z \<noteq> 0}))"
by (rule tendsto_mono[OF _ g]) simp_all
@@ -2914,8 +2588,7 @@
moreover {
have "filterlim (\<lambda>x. norm (f x)) (principal {0<..}) (principal {z. f z \<noteq> 0})"
by (auto simp: filterlim_def)
- hence "filterlim (\<lambda>x. norm (f x)) (principal {0<..})
- (inf F (principal {z. f z \<noteq> 0}))"
+ hence "filterlim (\<lambda>x. norm (f x)) (principal {0<..}) (inf F (principal {z. f z \<noteq> 0}))"
by (rule filterlim_mono) simp_all
}
ultimately have norm: "filterlim (\<lambda>x. norm (f x)) (at_right 0) (inf F (principal {z. f z \<noteq> 0}))"
@@ -3051,7 +2724,7 @@
lemma lim_ln_over_power:
fixes s :: real
assumes "0 < s"
- shows "((\<lambda>n. ln n / (n powr s)) \<longlongrightarrow> 0) sequentially"
+ shows "(\<lambda>n. ln (real n) / real n powr s) \<longlonglongrightarrow> 0"
proof -
have "(\<lambda>n. ln (Suc n) / (Suc n) powr s) \<longlonglongrightarrow> 0"
using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms
@@ -3095,7 +2768,7 @@
apply (subst filterlim_sequentially_Suc [symmetric])
by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide)
-lemma lim_1_over_Ln: "((\<lambda>n. 1 / Ln(of_nat n)) \<longlongrightarrow> 0) sequentially"
+lemma lim_1_over_Ln: "(\<lambda>n. 1 / Ln (complex_of_nat n)) \<longlonglongrightarrow> 0"
proof (clarsimp simp add: lim_sequentially dist_norm norm_divide field_split_simps)
fix r::real
assume "0 < r"
@@ -3117,7 +2790,7 @@
by (rule_tac x=n in exI) (force simp: field_split_simps intro: less_le_trans)
qed
-lemma lim_1_over_ln: "((\<lambda>n. 1 / ln(real_of_nat n)) \<longlongrightarrow> 0) sequentially"
+lemma lim_1_over_ln: "(\<lambda>n. 1 / ln (real n)) \<longlonglongrightarrow> 0"
using lim_1_over_Ln [THEN filterlim_sequentially_Suc [THEN iffD2]]
apply (subst filterlim_sequentially_Suc [symmetric])
by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide)
@@ -3151,19 +2824,14 @@
qed
lemma lim_ln_over_ln1: "(\<lambda>n. ln n / ln(Suc n)) \<longlonglongrightarrow> 1"
-proof -
- have "(\<lambda>n. inverse (ln(Suc n) / ln n)) \<longlonglongrightarrow> inverse 1"
- by (rule tendsto_inverse [OF lim_ln1_over_ln]) auto
- then show ?thesis
- by simp
-qed
+ using tendsto_inverse [OF lim_ln1_over_ln] by force
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation between Square Root and exp/ln, hence its derivative\<close>
lemma csqrt_exp_Ln:
assumes "z \<noteq> 0"
- shows "csqrt z = exp(Ln(z) / 2)"
+ shows "csqrt z = exp(Ln z / 2)"
proof -
have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))"
by (metis exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral)
@@ -3263,12 +2931,7 @@
by (simp add: continuous_at_imp_continuous_on continuous_within_csqrt)
lemma holomorphic_on_csqrt [holomorphic_intros]: "csqrt holomorphic_on -\<real>\<^sub>\<le>\<^sub>0"
-proof -
- have *: "(\<lambda>z. exp (ln z / 2)) holomorphic_on -\<real>\<^sub>\<le>\<^sub>0"
- by (intro holomorphic_intros) auto
- then show ?thesis
- using field_differentiable_within_csqrt holomorphic_on_def by auto
-qed
+ by (simp add: field_differentiable_within_csqrt holomorphic_on_def)
lemma holomorphic_on_csqrt' [holomorphic_intros]:
"f holomorphic_on A \<Longrightarrow> (\<And>z. z \<in> A \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> (\<lambda>z. csqrt (f z)) holomorphic_on A"
@@ -3283,8 +2946,7 @@
lemma continuous_within_closed_nontrivial:
"closed s \<Longrightarrow> a \<notin> s ==> continuous (at a within s) f"
- using open_Compl
- by (force simp add: continuous_def eventually_at_topological filterlim_iff open_Collect_neg)
+ using Compl_iff continuous_within_topological open_Compl by fastforce
lemma continuous_within_csqrt_posreal:
"continuous (at z within (\<real> \<inter> {w. 0 \<le> Re(w)})) csqrt"
@@ -3344,17 +3006,13 @@
lemma tan_Arctan:
assumes "z\<^sup>2 \<noteq> -1"
- shows [simp]:"tan(Arctan z) = z"
+ shows [simp]: "tan(Arctan z) = z"
proof -
- have "1 + \<i>*z \<noteq> 0"
- by (metis assms complex_i_mult_minus i_squared minus_unique power2_eq_square power2_minus)
- moreover
- have "1 - \<i>*z \<noteq> 0"
- by (metis assms complex_i_mult_minus i_squared power2_eq_square power2_minus right_minus_eq)
- ultimately
- show ?thesis
- by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus csqrt_exp_Ln [symmetric]
- divide_simps power2_eq_square [symmetric])
+ obtain "1 + \<i>*z \<noteq> 0" "1 - \<i>*z \<noteq> 0"
+ by (metis add_diff_cancel_left' assms diff_0 i_times_eq_iff mult_cancel_left2 power2_i power2_minus right_minus_eq)
+ then show ?thesis
+ by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps
+ flip: csqrt_exp_Ln power2_eq_square)
qed
lemma Arctan_tan [simp]:
@@ -3377,8 +3035,7 @@
by (simp add: algebra_simps)
also have "\<dots> \<longleftrightarrow> False"
using assms ge_pi2
- apply (auto simp: algebra_simps)
- by (metis abs_mult_pos not_less of_nat_less_0_iff of_nat_numeral)
+ by (metis eq_divide_eq linorder_not_less mult.commute zero_neq_numeral)
finally have "exp (\<i>*z)*exp (\<i>*z) + 1 \<noteq> 0"
by (auto simp: add.commute minus_unique)
then show "exp (2 * z / \<i>) = (1 - \<i> * tan z) / (1 + \<i> * tan z)"
@@ -3464,7 +3121,7 @@
define G where [abs_def]: "G z = (\<Sum>n. g n * z^n)" for z
have summable: "summable (\<lambda>n. g n * u^n)" if "norm u < 1" for u
proof (cases "u = 0")
- assume u: "u \<noteq> 0"
+ case False
have "(\<lambda>n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\<lambda>n. ereal (inverse (norm u)^2) *
ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))"
proof
@@ -3485,10 +3142,10 @@
by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all
finally have "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) = inverse (norm u)^2"
by (intro lim_imp_Liminf) simp_all
- moreover from power_strict_mono[OF that, of 2] u have "inverse (norm u)^2 > 1"
+ moreover from power_strict_mono[OF that, of 2] False have "inverse (norm u)^2 > 1"
by (simp add: field_split_simps)
ultimately have A: "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp
- from u have "summable (h u)"
+ from False have "summable (h u)"
by (intro summable_norm_cancel[OF ratio_test_convergence[OF _ A]])
(auto simp: h_def norm_divide norm_mult norm_power simp del: of_nat_Suc
intro!: mult_pos_pos divide_pos_pos always_eventually)
@@ -3590,7 +3247,7 @@
by (simp add: Arctan_def)
next
have "tan (Re (Arctan (of_real x))) = Re (tan (Arctan (of_real x)))"
- by (auto simp: tan_def Complex.Re_divide Re_sin Re_cos Im_sin Im_cos field_simps power2_eq_square)
+ by (metis Im_Arctan_of_real Re_complex_of_real complex_is_Real_iff of_real_Re tan_of_real)
also have "\<dots> = x"
proof -
have "(complex_of_real x)\<^sup>2 \<noteq> - 1"
@@ -3632,8 +3289,7 @@
show 12: "- (pi / 2) < arctan x + arctan y" "arctan x + arctan y < pi / 2"
using assms by linarith+
show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)"
- using cos_gt_zero_pi [OF 12]
- by (simp add: arctan tan_add)
+ using cos_gt_zero_pi [OF 12] by (simp add: arctan tan_add)
qed
lemma arctan_inverse:
@@ -3677,8 +3333,7 @@
lemma arctan_bounds:
assumes "0 \<le> x" "x < 1"
shows arctan_lower_bound:
- "(\<Sum>k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \<le> arctan x"
- (is "(\<Sum>k<_. (- 1)^ k * ?a k) \<le> _")
+ "(\<Sum>k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \<le> arctan x" (is "(\<Sum>k<_. _ * ?a k) \<le> _")
and arctan_upper_bound:
"arctan x \<le> (\<Sum>k<2 * n + 1. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))"
proof -
@@ -3724,7 +3379,7 @@
lemma Arcsin_body_lemma: "\<i> * z + csqrt(1 - z\<^sup>2) \<noteq> 0"
using power2_csqrt [of "1 - z\<^sup>2"]
- by (metis add.inverse_inverse complex_i_mult_minus diff_0 diff_add_cancel diff_minus_eq_add mult.assoc mult.commute numeral_One power2_eq_square zero_neq_numeral)
+ by (metis add.inverse_unique diff_0 diff_add_cancel mult.left_commute mult_minus1_right power2_i power2_minus power_mult_distrib zero_neq_one)
lemma Arcsin_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Re(\<i> * z + csqrt(1 - z\<^sup>2))"
using Complex.cmod_power2 [of z, symmetric]
@@ -4038,37 +3693,6 @@
lemma holomorphic_on_Arccos: "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos holomorphic_on s"
by (simp add: field_differentiable_within_Arccos holomorphic_on_def)
-text \<open>This theorem is about REAL cos/arccos but relies on theorems about @{term Arg}\<close>
-lemma cos_eq_arccos_Ex:
- "cos x = y \<longleftrightarrow> -1\<le>y \<and> y\<le>1 \<and> (\<exists>k::int. x = arccos y + 2*k*pi \<or> x = - arccos y + 2*k*pi)" (is "?L=?R")
-proof
- assume R: ?R
- then obtain k::int where "x = arccos y + 2*k*pi \<or> x = - arccos y + 2*k*pi" by auto
- moreover have "cos x = y" when "x = arccos y + 2*k*pi"
- by (metis (no_types) R cos_arccos cos_eq_periodic_intro cos_minus minus_add_cancel)
- moreover have "cos x = y" when "x = -arccos y + 2*k*pi"
- by (metis add_minus_cancel R cos_arccos cos_eq_periodic_intro uminus_add_conv_diff)
- ultimately show "cos x = y" by auto
-next
- assume L: ?L
- let ?goal = "(\<exists>k::int. x = arccos y + 2*k*pi \<or> x = - arccos y + 2*k*pi)"
- obtain k::int where k: "-pi < x-k*2*pi" "x-k*2*pi \<le> pi"
- by (metis Arg_bounded Arg_exp_diff_2pi complex.sel(2) mult.assoc of_int_mult of_int_numeral)
- have *: "cos (x - k * 2*pi) = y"
- using cos.periodic_simps(3)[of x "-k"] L by (auto simp add:field_simps)
- then have \<section>: ?goal when "x-k*2*pi \<ge> 0"
- using arccos_cos k that by force
- moreover have ?goal when "\<not> x-k*2*pi \<ge>0"
- proof -
- have "cos (k * 2*pi - x) = y"
- by (metis * cos_minus minus_diff_eq)
- then show ?goal
- using arccos_cos \<section> k by fastforce
- qed
- ultimately show "-1\<le>y \<and> y\<le>1 \<and> ?goal"
- using L by auto
-qed
-
subsection\<^marker>\<open>tag unimportant\<close>\<open>Upper and Lower Bounds for Inverse Sine and Cosine\<close>
lemma Arcsin_bounds: "\<bar>Re z\<bar> < 1 \<Longrightarrow> \<bar>Re(Arcsin z)\<bar> < pi/2"
--- a/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Analysis/Elementary_Metric_Spaces.thy Mon Feb 20 13:59:42 2023 +0100
@@ -916,6 +916,72 @@
using continuous_attains_sup[of "S \<times> S" "\<lambda>x. dist (fst x) (snd x)"] by auto
qed
+text \<open>
+ If \<open>A\<close> is a compact subset of an open set \<open>B\<close> in a metric space, then there exists an \<open>\<epsilon> > 0\<close>
+ such that the Minkowski sum of \<open>A\<close> with an open ball of radius \<open>\<epsilon>\<close> is also a subset of \<open>B\<close>.
+\<close>
+lemma compact_subset_open_imp_ball_epsilon_subset:
+ assumes "compact A" "open B" "A \<subseteq> B"
+ obtains e where "e > 0" "(\<Union>x\<in>A. ball x e) \<subseteq> B"
+proof -
+ have "\<forall>x\<in>A. \<exists>e. e > 0 \<and> ball x e \<subseteq> B"
+ using assms unfolding open_contains_ball by blast
+ then obtain e where e: "\<And>x. x \<in> A \<Longrightarrow> e x > 0" "\<And>x. x \<in> A \<Longrightarrow> ball x (e x) \<subseteq> B"
+ by metis
+ define C where "C = e ` A"
+ obtain X where X: "X \<subseteq> A" "finite X" "A \<subseteq> (\<Union>c\<in>X. ball c (e c / 2))"
+ using assms(1)
+ proof (rule compactE_image)
+ show "open (ball x (e x / 2))" if "x \<in> A" for x
+ by simp
+ show "A \<subseteq> (\<Union>c\<in>A. ball c (e c / 2))"
+ using e by auto
+ qed auto
+
+ define e' where "e' = Min (insert 1 ((\<lambda>x. e x / 2) ` X))"
+ have "e' > 0"
+ unfolding e'_def using e X by (subst Min_gr_iff) auto
+ have e': "e' \<le> e x / 2" if "x \<in> X" for x
+ using that X unfolding e'_def by (intro Min.coboundedI) auto
+
+ show ?thesis
+ proof
+ show "e' > 0"
+ by fact
+ next
+ show "(\<Union>x\<in>A. ball x e') \<subseteq> B"
+ proof clarify
+ fix x y assume xy: "x \<in> A" "y \<in> ball x e'"
+ from xy(1) X obtain z where z: "z \<in> X" "x \<in> ball z (e z / 2)"
+ by auto
+ have "dist y z \<le> dist x y + dist z x"
+ by (metis dist_commute dist_triangle)
+ also have "dist z x < e z / 2"
+ using xy z by auto
+ also have "dist x y < e'"
+ using xy by auto
+ also have "\<dots> \<le> e z / 2"
+ using z by (intro e') auto
+ finally have "y \<in> ball z (e z)"
+ by (simp add: dist_commute)
+ also have "\<dots> \<subseteq> B"
+ using z X by (intro e) auto
+ finally show "y \<in> B" .
+ qed
+ qed
+qed
+
+lemma compact_subset_open_imp_cball_epsilon_subset:
+ assumes "compact A" "open B" "A \<subseteq> B"
+ obtains e where "e > 0" "(\<Union>x\<in>A. cball x e) \<subseteq> B"
+proof -
+ obtain e where "e > 0" and e: "(\<Union>x\<in>A. ball x e) \<subseteq> B"
+ using compact_subset_open_imp_ball_epsilon_subset [OF assms] by blast
+ then have "(\<Union>x\<in>A. cball x (e / 2)) \<subseteq> (\<Union>x\<in>A. ball x e)"
+ by auto
+ with \<open>0 < e\<close> that show ?thesis
+ by (metis e half_gt_zero_iff order_trans)
+qed
subsubsection\<open>Totally bounded\<close>
--- a/src/HOL/Complex.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Complex.thy Mon Feb 20 13:59:42 2023 +0100
@@ -1001,6 +1001,21 @@
"continuous_on A f \<Longrightarrow> continuous_on A (\<lambda>x. cis (f x))"
by (auto simp: cis_conv_exp intro!: continuous_intros)
+lemma tendsto_exp_0_Re_at_bot: "(exp \<longlongrightarrow> 0) (filtercomap Re at_bot)"
+proof -
+ have "((\<lambda>z. cmod (exp z)) \<longlongrightarrow> 0) (filtercomap Re at_bot)"
+ by (auto intro!: filterlim_filtercomapI exp_at_bot)
+ thus ?thesis
+ using tendsto_norm_zero_iff by blast
+qed
+
+lemma filterlim_exp_at_infinity_Re_at_top: "filterlim exp at_infinity (filtercomap Re at_top)"
+proof -
+ have "filterlim (\<lambda>z. norm (exp z)) at_top (filtercomap Re at_top)"
+ by (auto intro!: filterlim_filtercomapI exp_at_top)
+ thus ?thesis
+ using filterlim_norm_at_top_imp_at_infinity by blast
+qed
subsubsection \<open>Complex argument\<close>
--- a/src/HOL/Complex_Analysis/Complex_Analysis.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Complex_Analysis/Complex_Analysis.thy Mon Feb 20 13:59:42 2023 +0100
@@ -1,7 +1,7 @@
theory Complex_Analysis
imports
Residue_Theorem
- Riemann_Mapping
+ Meromorphic
begin
end
--- a/src/HOL/Complex_Analysis/Complex_Singularities.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Complex_Analysis/Complex_Singularities.thy Mon Feb 20 13:59:42 2023 +0100
@@ -249,6 +249,12 @@
shows "not_essential f z \<longleftrightarrow> not_essential g z'"
unfolding not_essential_def using assms filterlim_cong is_pole_cong by fastforce
+lemma not_essential_compose_iff:
+ assumes "filtermap g (at z) = at z'"
+ shows "not_essential (f \<circ> g) z = not_essential f z'"
+ unfolding not_essential_def filterlim_def filtermap_compose assms is_pole_compose_iff[OF assms]
+ by blast
+
lemma isolated_singularity_at_cong:
assumes "eventually (\<lambda>x. f x = g x) (at z)" "z = z'"
shows "isolated_singularity_at f z \<longleftrightarrow> isolated_singularity_at g z'"
@@ -362,8 +368,8 @@
ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
by (simp add: continuous_within)
moreover have "(g \<longlongrightarrow> g z) F"
- using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \<open>r>0\<close>
- unfolding F_def by auto
+ unfolding F_def
+ using \<open>r>0\<close> centre_in_ball continuous_on_def g_holo holomorphic_on_imp_continuous_on by blast
ultimately show " ((\<lambda>w. f' w * g w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
qed
moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)"
@@ -1058,7 +1064,7 @@
using analytic_at not_is_pole_holomorphic by blast
lemma not_essential_const [singularity_intros]: "not_essential (\<lambda>_. c) z"
- unfolding not_essential_def by (rule exI[of _ c]) auto
+ by blast
lemma not_essential_uminus [singularity_intros]:
assumes f_ness: "not_essential f z"
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Complex_Analysis/Laurent_Convergence.thy Mon Feb 20 13:59:42 2023 +0100
@@ -0,0 +1,2917 @@
+theory Laurent_Convergence
+ imports "HOL-Computational_Algebra.Formal_Laurent_Series" "HOL-Library.Landau_Symbols"
+ Residue_Theorem
+
+begin
+
+(* TODO: Move *)
+text \<open>TODO: Better than @{thm deriv_compose_linear}?\<close>
+lemma deriv_compose_linear':
+ assumes "f field_differentiable at (c * z+a)"
+ shows "deriv (\<lambda>w. f (c * w+a)) z = c * deriv f (c * z+a)"
+ apply (subst deriv_chain[where f="\<lambda>w. c * w+a",unfolded comp_def])
+ using assms by (auto intro:derivative_intros)
+
+text \<open>TODO: Better than @{thm higher_deriv_compose_linear}?\<close>
+lemma higher_deriv_compose_linear':
+ fixes z::complex
+ assumes f: "f holomorphic_on T" and S: "open S" and T: "open T" and z: "z \<in> S"
+ and fg: "\<And>w. w \<in> S \<Longrightarrow> u * w+c \<in> T"
+ shows "(deriv ^^ n) (\<lambda>w. f (u * w+c)) z = u^n * (deriv ^^ n) f (u * z+c)"
+using z
+proof (induction n arbitrary: z)
+ case 0 then show ?case by simp
+next
+ case (Suc n z)
+ have holo0: "f holomorphic_on (\<lambda>w. u * w+c) ` S"
+ by (meson fg f holomorphic_on_subset image_subset_iff)
+ have holo2: "(deriv ^^ n) f holomorphic_on (\<lambda>w. u * w+c) ` S"
+ by (meson f fg holomorphic_higher_deriv holomorphic_on_subset image_subset_iff T)
+ have holo3: "(\<lambda>z. u ^ n * (deriv ^^ n) f (u * z+c)) holomorphic_on S"
+ by (intro holo2 holomorphic_on_compose [where g="(deriv ^^ n) f", unfolded o_def] holomorphic_intros)
+ have "(\<lambda>w. u * w+c) holomorphic_on S" "f holomorphic_on (\<lambda>w. u * w+c) ` S"
+ by (rule holo0 holomorphic_intros)+
+ then have holo1: "(\<lambda>w. f (u * w+c)) holomorphic_on S"
+ by (rule holomorphic_on_compose [where g=f, unfolded o_def])
+ have "deriv ((deriv ^^ n) (\<lambda>w. f (u * w+c))) z = deriv (\<lambda>z. u^n * (deriv ^^ n) f (u*z+c)) z"
+ proof (rule complex_derivative_transform_within_open [OF _ holo3 S Suc.prems])
+ show "(deriv ^^ n) (\<lambda>w. f (u * w+c)) holomorphic_on S"
+ by (rule holomorphic_higher_deriv [OF holo1 S])
+ qed (simp add: Suc.IH)
+ also have "\<dots> = u^n * deriv (\<lambda>z. (deriv ^^ n) f (u * z+c)) z"
+ proof -
+ have "(deriv ^^ n) f analytic_on T"
+ by (simp add: analytic_on_open f holomorphic_higher_deriv T)
+ then have "(\<lambda>w. (deriv ^^ n) f (u * w+c)) analytic_on S"
+ proof -
+ have "(deriv ^^ n) f \<circ> (\<lambda>w. u * w+c) holomorphic_on S"
+ using holomorphic_on_compose[OF _ holo2] \<open>(\<lambda>w. u * w+c) holomorphic_on S\<close>
+ by simp
+ then show ?thesis
+ by (simp add: S analytic_on_open o_def)
+ qed
+ then show ?thesis
+ by (intro deriv_cmult analytic_on_imp_differentiable_at [OF _ Suc.prems])
+ qed
+ also have "\<dots> = u * u ^ n * deriv ((deriv ^^ n) f) (u * z+c)"
+ proof -
+ have "(deriv ^^ n) f field_differentiable at (u * z+c)"
+ using Suc.prems T f fg holomorphic_higher_deriv holomorphic_on_imp_differentiable_at by blast
+ then show ?thesis
+ by (simp add: deriv_compose_linear')
+ qed
+ finally show ?case
+ by simp
+qed
+
+lemma fps_to_fls_numeral [simp]: "fps_to_fls (numeral n) = numeral n"
+ by (metis fps_to_fls_of_nat of_nat_numeral)
+
+lemma fls_const_power: "fls_const (a ^ b) = fls_const a ^ b"
+ by (induction b) (auto simp flip: fls_const_mult_const)
+
+lemma fls_deriv_numeral [simp]: "fls_deriv (numeral n) = 0"
+ by (metis fls_deriv_of_int of_int_numeral)
+
+lemma fls_const_numeral [simp]: "fls_const (numeral n) = numeral n"
+ by (metis fls_of_nat of_nat_numeral)
+
+lemma fls_mult_of_int_nth [simp]:
+ shows "fls_nth (numeral k * f) n = numeral k * fls_nth f n"
+ and "fls_nth (f * numeral k) n = fls_nth f n * numeral k"
+ by (metis fls_const_numeral fls_mult_const_nth)+
+
+lemma fls_nth_numeral' [simp]:
+ "fls_nth (numeral n) 0 = numeral n" "k \<noteq> 0 \<Longrightarrow> fls_nth (numeral n) k = 0"
+ by (subst fls_const_numeral [symmetric], subst fls_const_nth, simp)+
+
+lemma fls_subdegree_prod:
+ fixes F :: "'a \<Rightarrow> 'b :: field_char_0 fls"
+ assumes "\<And>x. x \<in> I \<Longrightarrow> F x \<noteq> 0"
+ shows "fls_subdegree (\<Prod>x\<in>I. F x) = (\<Sum>x\<in>I. fls_subdegree (F x))"
+ using assms by (induction I rule: infinite_finite_induct) auto
+
+lemma fls_subdegree_prod':
+ fixes F :: "'a \<Rightarrow> 'b :: field_char_0 fls"
+ assumes "\<And>x. x \<in> I \<Longrightarrow> fls_subdegree (F x) \<noteq> 0"
+ shows "fls_subdegree (\<Prod>x\<in>I. F x) = (\<Sum>x\<in>I. fls_subdegree (F x))"
+proof (intro fls_subdegree_prod)
+ show "F x \<noteq> 0" if "x \<in> I" for x
+ using assms[OF that] by auto
+qed
+
+instance fps :: (semiring_char_0) semiring_char_0
+proof
+ show "inj (of_nat :: nat \<Rightarrow> 'a fps)"
+ proof
+ fix m n :: nat
+ assume "of_nat m = (of_nat n :: 'a fps)"
+ hence "fps_nth (of_nat m) 0 = (fps_nth (of_nat n) 0 :: 'a)"
+ by (simp only: )
+ thus "m = n"
+ by simp
+ qed
+qed
+
+instance fls :: (semiring_char_0) semiring_char_0
+proof
+ show "inj (of_nat :: nat \<Rightarrow> 'a fls)"
+ proof
+ fix m n :: nat
+ assume "of_nat m = (of_nat n :: 'a fls)"
+ hence "fls_nth (of_nat m) 0 = (fls_nth (of_nat n) 0 :: 'a)"
+ by (simp only: )
+ thus "m = n"
+ by (simp add: fls_of_nat_nth)
+ qed
+qed
+
+lemma fls_const_eq_0_iff [simp]: "fls_const c = 0 \<longleftrightarrow> c = 0"
+ using fls_const_0 fls_const_nonzero by blast
+
+lemma fls_subdegree_add_eq1:
+ assumes "f \<noteq> 0" "fls_subdegree f < fls_subdegree g"
+ shows "fls_subdegree (f + g) = fls_subdegree f"
+proof (intro antisym)
+ from assms have *: "fls_nth (f + g) (fls_subdegree f) \<noteq> 0"
+ by auto
+ from * show "fls_subdegree (f + g) \<le> fls_subdegree f"
+ by (rule fls_subdegree_leI)
+ from * have "f + g \<noteq> 0"
+ using fls_nonzeroI by blast
+ thus "fls_subdegree f \<le> fls_subdegree (f + g)"
+ using assms(2) fls_plus_subdegree by force
+qed
+
+lemma fls_subdegree_add_eq2:
+ assumes "g \<noteq> 0" "fls_subdegree g < fls_subdegree f"
+ shows "fls_subdegree (f + g) = fls_subdegree g"
+proof (intro antisym)
+ from assms have *: "fls_nth (f + g) (fls_subdegree g) \<noteq> 0"
+ by auto
+ from * show "fls_subdegree (f + g) \<le> fls_subdegree g"
+ by (rule fls_subdegree_leI)
+ from * have "f + g \<noteq> 0"
+ using fls_nonzeroI by blast
+ thus "fls_subdegree g \<le> fls_subdegree (f + g)"
+ using assms(2) fls_plus_subdegree by force
+qed
+
+lemma fls_subdegree_diff_eq1:
+ assumes "f \<noteq> 0" "fls_subdegree f < fls_subdegree g"
+ shows "fls_subdegree (f - g) = fls_subdegree f"
+ using fls_subdegree_add_eq1[of f "-g"] assms by simp
+
+lemma fls_subdegree_diff_eq2:
+ assumes "g \<noteq> 0" "fls_subdegree g < fls_subdegree f"
+ shows "fls_subdegree (f - g) = fls_subdegree g"
+ using fls_subdegree_add_eq2[of "-g" f] assms by simp
+
+lemma nat_minus_fls_subdegree_plus_const_eq:
+ "nat (-fls_subdegree (F + fls_const c)) = nat (-fls_subdegree F)"
+proof (cases "fls_subdegree F < 0")
+ case True
+ hence "fls_subdegree (F + fls_const c) = fls_subdegree F"
+ by (intro fls_subdegree_add_eq1) auto
+ thus ?thesis
+ by simp
+next
+ case False
+ thus ?thesis
+ by (auto simp: fls_subdegree_ge0I)
+qed
+
+lemma at_to_0': "NO_MATCH 0 z \<Longrightarrow> at z = filtermap (\<lambda>x. x + z) (at 0)"
+ for z :: "'a::real_normed_vector"
+ by (rule at_to_0)
+
+lemma nhds_to_0: "nhds (x :: 'a :: real_normed_vector) = filtermap ((+) x) (nhds 0)"
+proof -
+ have "(\<lambda>xa. xa - - x) = (+) x"
+ by auto
+ thus ?thesis
+ using filtermap_nhds_shift[of "-x" 0] by simp
+qed
+
+lemma nhds_to_0': "NO_MATCH 0 x \<Longrightarrow> nhds (x :: 'a :: real_normed_vector) = filtermap ((+) x) (nhds 0)"
+ by (rule nhds_to_0)
+
+
+definition%important fls_conv_radius :: "complex fls \<Rightarrow> ereal" where
+ "fls_conv_radius f = fps_conv_radius (fls_regpart f)"
+
+definition%important eval_fls :: "complex fls \<Rightarrow> complex \<Rightarrow> complex" where
+ "eval_fls F z = eval_fps (fls_base_factor_to_fps F) z * z powi fls_subdegree F"
+
+definition\<^marker>\<open>tag important\<close>
+ has_laurent_expansion :: "(complex \<Rightarrow> complex) \<Rightarrow> complex fls \<Rightarrow> bool"
+ (infixl "has'_laurent'_expansion" 60)
+ where "(f has_laurent_expansion F) \<longleftrightarrow>
+ fls_conv_radius F > 0 \<and> eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
+
+lemma has_laurent_expansion_schematicI:
+ "f has_laurent_expansion F \<Longrightarrow> F = G \<Longrightarrow> f has_laurent_expansion G"
+ by simp
+
+lemma has_laurent_expansion_cong:
+ assumes "eventually (\<lambda>x. f x = g x) (at 0)" "F = G"
+ shows "(f has_laurent_expansion F) \<longleftrightarrow> (g has_laurent_expansion G)"
+proof -
+ have "eventually (\<lambda>z. eval_fls F z = g z) (at 0)"
+ if "eventually (\<lambda>z. eval_fls F z = f z) (at 0)" "eventually (\<lambda>x. f x = g x) (at 0)" for f g
+ using that by eventually_elim auto
+ from this[of f g] this[of g f] show ?thesis
+ using assms by (auto simp: eq_commute has_laurent_expansion_def)
+qed
+
+lemma has_laurent_expansion_cong':
+ assumes "eventually (\<lambda>x. f x = g x) (at z)" "F = G" "z = z'"
+ shows "((\<lambda>x. f (z + x)) has_laurent_expansion F) \<longleftrightarrow> ((\<lambda>x. g (z' + x)) has_laurent_expansion G)"
+ by (intro has_laurent_expansion_cong)
+ (use assms in \<open>auto simp: at_to_0' eventually_filtermap add_ac\<close>)
+
+lemma fls_conv_radius_altdef:
+ "fls_conv_radius F = fps_conv_radius (fls_base_factor_to_fps F)"
+proof -
+ have "conv_radius (\<lambda>n. fls_nth F (int n)) = conv_radius (\<lambda>n. fls_nth F (int n + fls_subdegree F))"
+ proof (cases "fls_subdegree F \<ge> 0")
+ case True
+ hence "conv_radius (\<lambda>n. fls_nth F (int n + fls_subdegree F)) =
+ conv_radius (\<lambda>n. fls_nth F (int (n + nat (fls_subdegree F))))"
+ by auto
+ thus ?thesis
+ by (subst (asm) conv_radius_shift) auto
+ next
+ case False
+ hence "conv_radius (\<lambda>n. fls_nth F (int n)) =
+ conv_radius (\<lambda>n. fls_nth F (fls_subdegree F + int (n + nat (-fls_subdegree F))))"
+ by auto
+ thus ?thesis
+ by (subst (asm) conv_radius_shift) (auto simp: add_ac)
+ qed
+ thus ?thesis
+ by (simp add: fls_conv_radius_def fps_conv_radius_def)
+qed
+
+lemma eval_fps_of_nat [simp]: "eval_fps (of_nat n) z = of_nat n"
+ and eval_fps_of_int [simp]: "eval_fps (of_int m) z = of_int m"
+ by (simp_all flip: fps_of_nat fps_of_int)
+
+lemma fls_subdegree_numeral [simp]: "fls_subdegree (numeral n) = 0"
+ by (metis fls_subdegree_of_nat of_nat_numeral)
+
+lemma fls_regpart_numeral [simp]: "fls_regpart (numeral n) = numeral n"
+ by (metis fls_regpart_of_nat of_nat_numeral)
+
+lemma fps_conv_radius_of_nat [simp]: "fps_conv_radius (of_nat n) = \<infinity>"
+ and fps_conv_radius_of_int [simp]: "fps_conv_radius (of_int m) = \<infinity>"
+ by (simp_all flip: fps_of_nat fps_of_int)
+
+lemma fps_conv_radius_fls_regpart: "fps_conv_radius (fls_regpart F) = fls_conv_radius F"
+ by (simp add: fls_conv_radius_def)
+
+lemma fls_conv_radius_0 [simp]: "fls_conv_radius 0 = \<infinity>"
+ and fls_conv_radius_1 [simp]: "fls_conv_radius 1 = \<infinity>"
+ and fls_conv_radius_const [simp]: "fls_conv_radius (fls_const c) = \<infinity>"
+ and fls_conv_radius_numeral [simp]: "fls_conv_radius (numeral num) = \<infinity>"
+ and fls_conv_radius_of_nat [simp]: "fls_conv_radius (of_nat n) = \<infinity>"
+ and fls_conv_radius_of_int [simp]: "fls_conv_radius (of_int m) = \<infinity>"
+ and fls_conv_radius_X [simp]: "fls_conv_radius fls_X = \<infinity>"
+ and fls_conv_radius_X_inv [simp]: "fls_conv_radius fls_X_inv = \<infinity>"
+ and fls_conv_radius_X_intpow [simp]: "fls_conv_radius (fls_X_intpow m) = \<infinity>"
+ by (simp_all add: fls_conv_radius_def fls_X_intpow_regpart)
+
+lemma fls_conv_radius_shift [simp]: "fls_conv_radius (fls_shift n F) = fls_conv_radius F"
+ unfolding fls_conv_radius_altdef by (subst fls_base_factor_to_fps_shift) (rule refl)
+
+lemma fls_conv_radius_fps_to_fls [simp]: "fls_conv_radius (fps_to_fls F) = fps_conv_radius F"
+ by (simp add: fls_conv_radius_def)
+
+lemma fls_conv_radius_deriv [simp]: "fls_conv_radius (fls_deriv F) \<ge> fls_conv_radius F"
+proof -
+ have "fls_conv_radius (fls_deriv F) = fps_conv_radius (fls_regpart (fls_deriv F))"
+ by (simp add: fls_conv_radius_def)
+ also have "fls_regpart (fls_deriv F) = fps_deriv (fls_regpart F)"
+ by (intro fps_ext) (auto simp: add_ac)
+ also have "fps_conv_radius \<dots> \<ge> fls_conv_radius F"
+ by (simp add: fls_conv_radius_def fps_conv_radius_deriv)
+ finally show ?thesis .
+qed
+
+lemma fls_conv_radius_uminus [simp]: "fls_conv_radius (-F) = fls_conv_radius F"
+ by (simp add: fls_conv_radius_def)
+
+lemma fls_conv_radius_add: "fls_conv_radius (F + G) \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
+ by (simp add: fls_conv_radius_def fps_conv_radius_add)
+
+lemma fls_conv_radius_diff: "fls_conv_radius (F - G) \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
+ by (simp add: fls_conv_radius_def fps_conv_radius_diff)
+
+lemma fls_conv_radius_mult: "fls_conv_radius (F * G) \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
+proof (cases "F = 0 \<or> G = 0")
+ case False
+ hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
+ by auto
+ have "fls_conv_radius (F * G) = fps_conv_radius (fls_regpart (fls_shift (fls_subdegree F + fls_subdegree G) (F * G)))"
+ by (simp add: fls_conv_radius_altdef)
+ also have "fls_regpart (fls_shift (fls_subdegree F + fls_subdegree G) (F * G)) =
+ fls_base_factor_to_fps F * fls_base_factor_to_fps G"
+ by (simp add: fls_times_def)
+ also have "fps_conv_radius \<dots> \<ge> min (fls_conv_radius F) (fls_conv_radius G)"
+ unfolding fls_conv_radius_altdef by (rule fps_conv_radius_mult)
+ finally show ?thesis .
+qed auto
+
+lemma fps_conv_radius_add_ge:
+ "fps_conv_radius F \<ge> r \<Longrightarrow> fps_conv_radius G \<ge> r \<Longrightarrow> fps_conv_radius (F + G) \<ge> r"
+ using fps_conv_radius_add[of F G] by (simp add: min_def split: if_splits)
+
+lemma fps_conv_radius_diff_ge:
+ "fps_conv_radius F \<ge> r \<Longrightarrow> fps_conv_radius G \<ge> r \<Longrightarrow> fps_conv_radius (F - G) \<ge> r"
+ using fps_conv_radius_diff[of F G] by (simp add: min_def split: if_splits)
+
+lemma fps_conv_radius_mult_ge:
+ "fps_conv_radius F \<ge> r \<Longrightarrow> fps_conv_radius G \<ge> r \<Longrightarrow> fps_conv_radius (F * G) \<ge> r"
+ using fps_conv_radius_mult[of F G] by (simp add: min_def split: if_splits)
+
+lemma fls_conv_radius_add_ge:
+ "fls_conv_radius F \<ge> r \<Longrightarrow> fls_conv_radius G \<ge> r \<Longrightarrow> fls_conv_radius (F + G) \<ge> r"
+ using fls_conv_radius_add[of F G] by (simp add: min_def split: if_splits)
+
+lemma fls_conv_radius_diff_ge:
+ "fls_conv_radius F \<ge> r \<Longrightarrow> fls_conv_radius G \<ge> r \<Longrightarrow> fls_conv_radius (F - G) \<ge> r"
+ using fls_conv_radius_diff[of F G] by (simp add: min_def split: if_splits)
+
+lemma fls_conv_radius_mult_ge:
+ "fls_conv_radius F \<ge> r \<Longrightarrow> fls_conv_radius G \<ge> r \<Longrightarrow> fls_conv_radius (F * G) \<ge> r"
+ using fls_conv_radius_mult[of F G] by (simp add: min_def split: if_splits)
+
+lemma fls_conv_radius_power: "fls_conv_radius (F ^ n) \<ge> fls_conv_radius F"
+ by (induction n) (auto intro!: fls_conv_radius_mult_ge)
+
+lemma eval_fls_0 [simp]: "eval_fls 0 z = 0"
+ and eval_fls_1 [simp]: "eval_fls 1 z = 1"
+ and eval_fls_const [simp]: "eval_fls (fls_const c) z = c"
+ and eval_fls_numeral [simp]: "eval_fls (numeral num) z = numeral num"
+ and eval_fls_of_nat [simp]: "eval_fls (of_nat n) z = of_nat n"
+ and eval_fls_of_int [simp]: "eval_fls (of_int m) z = of_int m"
+ and eval_fls_X [simp]: "eval_fls fls_X z = z"
+ and eval_fls_X_intpow [simp]: "eval_fls (fls_X_intpow m) z = z powi m"
+ by (simp_all add: eval_fls_def)
+
+lemma eval_fls_at_0: "eval_fls F 0 = (if fls_subdegree F \<ge> 0 then fls_nth F 0 else 0)"
+ by (cases "fls_subdegree F = 0")
+ (simp_all add: eval_fls_def fls_regpart_def eval_fps_at_0)
+
+lemma eval_fps_to_fls:
+ assumes "norm z < fps_conv_radius F"
+ shows "eval_fls (fps_to_fls F) z = eval_fps F z"
+proof (cases "F = 0")
+ case [simp]: False
+ have "eval_fps F z = eval_fps (unit_factor F * normalize F) z"
+ by (metis unit_factor_mult_normalize)
+ also have "\<dots> = eval_fps (unit_factor F * fps_X ^ subdegree F) z"
+ by simp
+ also have "\<dots> = eval_fps (unit_factor F) z * z ^ subdegree F"
+ using assms by (subst eval_fps_mult) auto
+ also have "\<dots> = eval_fls (fps_to_fls F) z"
+ unfolding eval_fls_def fls_base_factor_to_fps_to_fls fls_subdegree_fls_to_fps
+ power_int_of_nat ..
+ finally show ?thesis ..
+qed auto
+
+lemma eval_fls_shift:
+ assumes [simp]: "z \<noteq> 0"
+ shows "eval_fls (fls_shift n F) z = eval_fls F z * z powi -n"
+proof (cases "F = 0")
+ case [simp]: False
+ show ?thesis
+ unfolding eval_fls_def
+ by (subst fls_base_factor_to_fps_shift, subst fls_shift_subdegree[OF \<open>F \<noteq> 0\<close>], subst power_int_diff)
+ (auto simp: power_int_minus divide_simps)
+qed auto
+
+lemma eval_fls_add:
+ assumes "ereal (norm z) < fls_conv_radius F" "ereal (norm z) < fls_conv_radius G" "z \<noteq> 0"
+ shows "eval_fls (F + G) z = eval_fls F z + eval_fls G z"
+ using assms
+proof (induction "fls_subdegree F" "fls_subdegree G" arbitrary: F G rule: linorder_wlog)
+ case (sym F G)
+ show ?case
+ using sym(1)[of G F] sym(2-) by (simp add: add_ac)
+next
+ case (le F G)
+ show ?case
+ proof (cases "F = 0 \<or> G = 0")
+ case False
+ hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
+ by auto
+ note [simp] = \<open>z \<noteq> 0\<close>
+ define F' G' where "F' = fls_base_factor_to_fps F" "G' = fls_base_factor_to_fps G"
+ define m n where "m = fls_subdegree F" "n = fls_subdegree G"
+ have "m \<le> n"
+ using le by (auto simp: m_n_def)
+ have conv1: "ereal (cmod z) < fps_conv_radius F'" "ereal (cmod z) < fps_conv_radius G'"
+ using assms le by (simp_all add: F'_G'_def fls_conv_radius_altdef)
+ have conv2: "ereal (cmod z) < fps_conv_radius (G' * fps_X ^ nat (n - m))"
+ using conv1 by (intro less_le_trans[OF _ fps_conv_radius_mult]) auto
+ have conv3: "ereal (cmod z) < fps_conv_radius (F' + G' * fps_X ^ nat (n - m))"
+ using conv1 conv2 by (intro less_le_trans[OF _ fps_conv_radius_add]) auto
+
+ have "eval_fls F z + eval_fls G z = eval_fps F' z * z powi m + eval_fps G' z * z powi n"
+ unfolding eval_fls_def m_n_def[symmetric] F'_G'_def[symmetric]
+ by (simp add: power_int_add algebra_simps)
+ also have "\<dots> = (eval_fps F' z + eval_fps G' z * z powi (n - m)) * z powi m"
+ by (simp add: algebra_simps power_int_diff)
+ also have "eval_fps G' z * z powi (n - m) = eval_fps (G' * fps_X ^ nat (n - m)) z"
+ using assms \<open>m \<le> n\<close> conv1 by (subst eval_fps_mult) (auto simp: power_int_def)
+ also have "eval_fps F' z + \<dots> = eval_fps (F' + G' * fps_X ^ nat (n - m)) z"
+ using conv1 conv2 by (subst eval_fps_add) auto
+ also have "\<dots> = eval_fls (fps_to_fls (F' + G' * fps_X ^ nat (n - m))) z"
+ using conv3 by (subst eval_fps_to_fls) auto
+ also have "\<dots> * z powi m = eval_fls (fls_shift (-m) (fps_to_fls (F' + G' * fps_X ^ nat (n - m)))) z"
+ by (subst eval_fls_shift) auto
+ also have "fls_shift (-m) (fps_to_fls (F' + G' * fps_X ^ nat (n - m))) = F + G"
+ using \<open>m \<le> n\<close>
+ by (simp add: fls_times_fps_to_fls fps_to_fls_power fls_X_power_conv_shift_1
+ fls_shifted_times_simps F'_G'_def m_n_def)
+ finally show ?thesis ..
+ qed auto
+qed
+
+lemma eval_fls_minus:
+ assumes "ereal (norm z) < fls_conv_radius F"
+ shows "eval_fls (-F) z = -eval_fls F z"
+ using assms by (simp add: eval_fls_def eval_fps_minus fls_conv_radius_altdef)
+
+lemma eval_fls_diff:
+ assumes "ereal (norm z) < fls_conv_radius F" "ereal (norm z) < fls_conv_radius G"
+ and [simp]: "z \<noteq> 0"
+ shows "eval_fls (F - G) z = eval_fls F z - eval_fls G z"
+proof -
+ have "eval_fls (F + (-G)) z = eval_fls F z - eval_fls G z"
+ using assms by (subst eval_fls_add) (auto simp: eval_fls_minus)
+ thus ?thesis
+ by simp
+qed
+
+lemma eval_fls_mult:
+ assumes "ereal (norm z) < fls_conv_radius F" "ereal (norm z) < fls_conv_radius G" "z \<noteq> 0"
+ shows "eval_fls (F * G) z = eval_fls F z * eval_fls G z"
+proof (cases "F = 0 \<or> G = 0")
+ case False
+ hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
+ by auto
+ note [simp] = \<open>z \<noteq> 0\<close>
+ define F' G' where "F' = fls_base_factor_to_fps F" "G' = fls_base_factor_to_fps G"
+ define m n where "m = fls_subdegree F" "n = fls_subdegree G"
+ have "eval_fls F z * eval_fls G z = (eval_fps F' z * eval_fps G' z) * z powi (m + n)"
+ unfolding eval_fls_def m_n_def[symmetric] F'_G'_def[symmetric]
+ by (simp add: power_int_add algebra_simps)
+ also have "\<dots> = eval_fps (F' * G') z * z powi (m + n)"
+ using assms by (subst eval_fps_mult) (auto simp: F'_G'_def fls_conv_radius_altdef)
+ also have "\<dots> = eval_fls (F * G) z"
+ by (simp add: eval_fls_def F'_G'_def m_n_def) (simp add: fls_times_def)
+ finally show ?thesis ..
+qed auto
+
+lemma eval_fls_power:
+ assumes "ereal (norm z) < fls_conv_radius F" "z \<noteq> 0"
+ shows "eval_fls (F ^ n) z = eval_fls F z ^ n"
+proof (induction n)
+ case (Suc n)
+ have "eval_fls (F ^ Suc n) z = eval_fls (F * F ^ n) z"
+ by simp
+ also have "\<dots> = eval_fls F z * eval_fls (F ^ n) z"
+ using assms by (subst eval_fls_mult) (auto intro!: less_le_trans[OF _ fls_conv_radius_power])
+ finally show ?case
+ using Suc by simp
+qed auto
+
+lemma norm_summable_fls:
+ "norm z < fls_conv_radius f \<Longrightarrow> summable (\<lambda>n. norm (fls_nth f n * z ^ n))"
+ using norm_summable_fps[of z "fls_regpart f"] by (simp add: fls_conv_radius_def)
+
+lemma norm_summable_fls':
+ "norm z < fls_conv_radius f \<Longrightarrow> summable (\<lambda>n. norm (fls_nth f (n + fls_subdegree f) * z ^ n))"
+ using norm_summable_fps[of z "fls_base_factor_to_fps f"] by (simp add: fls_conv_radius_altdef)
+
+lemma summable_fls:
+ "norm z < fls_conv_radius f \<Longrightarrow> summable (\<lambda>n. fls_nth f n * z ^ n)"
+ by (rule summable_norm_cancel[OF norm_summable_fls])
+
+theorem sums_eval_fls:
+ fixes f
+ defines "n \<equiv> fls_subdegree f"
+ assumes "norm z < fls_conv_radius f" and "z \<noteq> 0 \<or> n \<ge> 0"
+ shows "(\<lambda>k. fls_nth f (int k + n) * z powi (int k + n)) sums eval_fls f z"
+proof (cases "z = 0")
+ case [simp]: False
+ have "(\<lambda>k. fps_nth (fls_base_factor_to_fps f) k * z ^ k * z powi n) sums
+ (eval_fps (fls_base_factor_to_fps f) z * z powi n)"
+ using assms(2) by (intro sums_eval_fps sums_mult2) (auto simp: fls_conv_radius_altdef)
+ thus ?thesis
+ by (simp add: power_int_add n_def eval_fls_def mult_ac)
+next
+ case [simp]: True
+ with assms have "n \<ge> 0"
+ by auto
+ have "(\<lambda>k. fls_nth f (int k + n) * z powi (int k + n)) sums
+ (\<Sum>k\<in>(if n \<le> 0 then {nat (-n)} else {}). fls_nth f (int k + n) * z powi (int k + n))"
+ by (intro sums_finite) (auto split: if_splits)
+ also have "\<dots> = eval_fls f z"
+ using \<open>n \<ge> 0\<close> by (auto simp: eval_fls_at_0 n_def not_le)
+ finally show ?thesis .
+qed
+
+lemma holomorphic_on_eval_fls:
+ fixes f
+ defines "n \<equiv> fls_subdegree f"
+ assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
+ shows "eval_fls f holomorphic_on A"
+proof (cases "n \<ge> 0")
+ case True
+ have "eval_fls f = (\<lambda>z. eval_fps (fls_base_factor_to_fps f) z * z ^ nat n)"
+ using True by (simp add: fun_eq_iff eval_fls_def power_int_def n_def)
+ moreover have "\<dots> holomorphic_on A"
+ using True assms(2) by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef)
+ ultimately show ?thesis
+ by simp
+next
+ case False
+ show ?thesis using assms
+ unfolding eval_fls_def by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef)
+qed
+
+lemma holomorphic_on_eval_fls' [holomorphic_intros]:
+ assumes "g holomorphic_on A"
+ assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
+ shows "(\<lambda>x. eval_fls f (g x)) holomorphic_on A"
+proof -
+ have "eval_fls f \<circ> g holomorphic_on A"
+ by (intro holomorphic_on_compose[OF assms(1) holomorphic_on_eval_fls]) (use assms in auto)
+ thus ?thesis
+ by (simp add: o_def)
+qed
+
+lemma continuous_on_eval_fls:
+ fixes f
+ defines "n \<equiv> fls_subdegree f"
+ assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
+ shows "continuous_on A (eval_fls f)"
+ by (intro holomorphic_on_imp_continuous_on holomorphic_on_eval_fls)
+ (use assms in auto)
+
+lemma continuous_on_eval_fls' [continuous_intros]:
+ fixes f
+ defines "n \<equiv> fls_subdegree f"
+ assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if n \<ge> 0 then {} else {0})"
+ assumes "continuous_on A g"
+ shows "continuous_on A (\<lambda>x. eval_fls f (g x))"
+ using assms(3)
+ by (intro continuous_on_compose2[OF continuous_on_eval_fls _ assms(2)])
+ (auto simp: n_def)
+
+lemmas has_field_derivative_eval_fps' [derivative_intros] =
+ DERIV_chain2[OF has_field_derivative_eval_fps]
+
+lemma fps_deriv_fls_regpart: "fps_deriv (fls_regpart F) = fls_regpart (fls_deriv F)"
+ by (intro fps_ext) (auto simp: add_ac)
+
+(* TODO: generalise for nonneg subdegree *)
+lemma has_field_derivative_eval_fls:
+ assumes "z \<in> eball 0 (fls_conv_radius f) - {0}"
+ shows "(eval_fls f has_field_derivative eval_fls (fls_deriv f) z) (at z within A)"
+proof -
+ define g where "g = fls_base_factor_to_fps f"
+ define n where "n = fls_subdegree f"
+ have [simp]: "fps_conv_radius g = fls_conv_radius f"
+ by (simp add: fls_conv_radius_altdef g_def)
+ have conv1: "fps_conv_radius (fps_deriv g * fps_X) \<ge> fls_conv_radius f"
+ by (intro fps_conv_radius_mult_ge order.trans[OF _ fps_conv_radius_deriv]) auto
+ have conv2: "fps_conv_radius (of_int n * g) \<ge> fls_conv_radius f"
+ by (intro fps_conv_radius_mult_ge) auto
+ have conv3: "fps_conv_radius (fps_deriv g * fps_X + of_int n * g) \<ge> fls_conv_radius f"
+ by (intro fps_conv_radius_add_ge conv1 conv2)
+
+ have [simp]: "fps_conv_radius g = fls_conv_radius f"
+ by (simp add: g_def fls_conv_radius_altdef)
+ have "((\<lambda>z. eval_fps g z * z powi fls_subdegree f) has_field_derivative
+ (eval_fps (fps_deriv g) z * z powi n + of_int n * z powi (n - 1) * eval_fps g z))
+ (at z within A)"
+ using assms by (auto intro!: derivative_eq_intros simp: n_def)
+ also have "(\<lambda>z. eval_fps g z * z powi fls_subdegree f) = eval_fls f"
+ by (simp add: eval_fls_def g_def fun_eq_iff)
+ also have "eval_fps (fps_deriv g) z * z powi n + of_int n * z powi (n - 1) * eval_fps g z =
+ (z * eval_fps (fps_deriv g) z + of_int n * eval_fps g z) * z powi (n - 1)"
+ using assms by (auto simp: power_int_diff field_simps)
+ also have "(z * eval_fps (fps_deriv g) z + of_int n * eval_fps g z) =
+ eval_fps (fps_deriv g * fps_X + of_int n * g) z"
+ using conv1 conv2 assms fps_conv_radius_deriv[of g]
+ by (subst eval_fps_add) (auto simp: eval_fps_mult)
+ also have "\<dots> = eval_fls (fps_to_fls (fps_deriv g * fps_X + of_int n * g)) z"
+ using conv3 assms by (subst eval_fps_to_fls) auto
+ also have "\<dots> * z powi (n - 1) = eval_fls (fls_shift (1 - n) (fps_to_fls (fps_deriv g * fps_X + of_int n * g))) z"
+ using assms by (subst eval_fls_shift) auto
+ also have "fls_shift (1 - n) (fps_to_fls (fps_deriv g * fps_X + of_int n * g)) = fls_deriv f"
+ by (intro fls_eqI) (auto simp: g_def n_def algebra_simps eq_commute[of _ "fls_subdegree f"])
+ finally show ?thesis .
+qed
+
+lemma eval_fls_deriv:
+ assumes "z \<in> eball 0 (fls_conv_radius F) - {0}"
+ shows "eval_fls (fls_deriv F) z = deriv (eval_fls F) z"
+ by (rule sym, rule DERIV_imp_deriv, rule has_field_derivative_eval_fls, rule assms)
+
+lemma analytic_on_eval_fls:
+ assumes "A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
+ shows "eval_fls f analytic_on A"
+proof (rule analytic_on_subset [OF _ assms])
+ show "eval_fls f analytic_on eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
+ using holomorphic_on_eval_fls[OF order.refl]
+ by (subst analytic_on_open) auto
+qed
+
+lemma analytic_on_eval_fls' [analytic_intros]:
+ assumes "g analytic_on A"
+ assumes "g ` A \<subseteq> eball 0 (fls_conv_radius f) - (if fls_subdegree f \<ge> 0 then {} else {0})"
+ shows "(\<lambda>x. eval_fls f (g x)) analytic_on A"
+proof -
+ have "eval_fls f \<circ> g analytic_on A"
+ by (intro analytic_on_compose[OF assms(1) analytic_on_eval_fls]) (use assms in auto)
+ thus ?thesis
+ by (simp add: o_def)
+qed
+
+lemma continuous_eval_fls [continuous_intros]:
+ assumes "z \<in> eball 0 (fls_conv_radius F) - (if fls_subdegree F \<ge> 0 then {} else {0})"
+ shows "continuous (at z within A) (eval_fls F)"
+proof -
+ have "isCont (eval_fls F) z"
+ using continuous_on_eval_fls[OF order.refl] assms
+ by (subst (asm) continuous_on_eq_continuous_at) auto
+ thus ?thesis
+ using continuous_at_imp_continuous_at_within by blast
+qed
+
+
+
+
+named_theorems laurent_expansion_intros
+
+lemma has_laurent_expansion_imp_asymp_equiv_0:
+ assumes F: "f has_laurent_expansion F"
+ defines "n \<equiv> fls_subdegree F"
+ shows "f \<sim>[at 0] (\<lambda>z. fls_nth F n * z powi n)"
+proof (cases "F = 0")
+ case True
+ thus ?thesis using assms
+ by (auto simp: has_laurent_expansion_def)
+next
+ case [simp]: False
+ define G where "G = fls_base_factor_to_fps F"
+ have "fls_conv_radius F > 0"
+ using F by (auto simp: has_laurent_expansion_def)
+ hence "isCont (eval_fps G) 0"
+ by (intro continuous_intros) (auto simp: G_def fps_conv_radius_fls_regpart zero_ereal_def)
+ hence lim: "eval_fps G \<midarrow>0\<rightarrow> eval_fps G 0"
+ by (meson isContD)
+ have [simp]: "fps_nth G 0 \<noteq> 0"
+ by (auto simp: G_def)
+
+ have "f \<sim>[at 0] eval_fls F"
+ using F by (intro asymp_equiv_refl_ev) (auto simp: has_laurent_expansion_def eq_commute)
+ also have "\<dots> = (\<lambda>z. eval_fps G z * z powi n)"
+ by (intro ext) (simp_all add: eval_fls_def G_def n_def)
+ also have "\<dots> \<sim>[at 0] (\<lambda>z. fps_nth G 0 * z powi n)" using lim
+ by (intro asymp_equiv_intros tendsto_imp_asymp_equiv_const) (auto simp: eval_fps_at_0)
+ also have "fps_nth G 0 = fls_nth F n"
+ by (simp add: G_def n_def)
+ finally show ?thesis
+ by simp
+qed
+
+lemma has_laurent_expansion_imp_asymp_equiv:
+ assumes F: "(\<lambda>w. f (z + w)) has_laurent_expansion F"
+ defines "n \<equiv> fls_subdegree F"
+ shows "f \<sim>[at z] (\<lambda>w. fls_nth F n * (w - z) powi n)"
+ using has_laurent_expansion_imp_asymp_equiv_0[OF assms(1)] unfolding n_def
+ by (simp add: at_to_0[of z] asymp_equiv_filtermap_iff add_ac)
+
+
+lemmas [tendsto_intros del] = tendsto_power_int
+
+lemma has_laurent_expansion_imp_tendsto_0:
+ assumes F: "f has_laurent_expansion F" and "fls_subdegree F \<ge> 0"
+ shows "f \<midarrow>0\<rightarrow> fls_nth F 0"
+proof (rule asymp_equiv_tendsto_transfer)
+ show "(\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) \<sim>[at 0] f"
+ by (rule asymp_equiv_symI, rule has_laurent_expansion_imp_asymp_equiv_0) fact
+ show "(\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) \<midarrow>0\<rightarrow> fls_nth F 0"
+ by (rule tendsto_eq_intros refl | use assms(2) in simp)+
+ (use assms(2) in \<open>auto simp: power_int_0_left_If\<close>)
+qed
+
+lemma has_laurent_expansion_imp_tendsto:
+ assumes F: "(\<lambda>w. f (z + w)) has_laurent_expansion F" and "fls_subdegree F \<ge> 0"
+ shows "f \<midarrow>z\<rightarrow> fls_nth F 0"
+ using has_laurent_expansion_imp_tendsto_0[OF assms]
+ by (simp add: at_to_0[of z] filterlim_filtermap add_ac)
+
+lemma has_laurent_expansion_imp_filterlim_infinity_0:
+ assumes F: "f has_laurent_expansion F" and "fls_subdegree F < 0"
+ shows "filterlim f at_infinity (at 0)"
+proof (rule asymp_equiv_at_infinity_transfer)
+ have [simp]: "F \<noteq> 0"
+ using assms(2) by auto
+ show "(\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) \<sim>[at 0] f"
+ by (rule asymp_equiv_symI, rule has_laurent_expansion_imp_asymp_equiv_0) fact
+ show "filterlim (\<lambda>z. fls_nth F (fls_subdegree F) * z powi fls_subdegree F) at_infinity (at 0)"
+ by (rule tendsto_mult_filterlim_at_infinity tendsto_const
+ filterlim_power_int_neg_at_infinity | use assms(2) in simp)+
+ (auto simp: eventually_at_filter)
+qed
+
+lemma has_laurent_expansion_imp_neg_fls_subdegree:
+ assumes F: "f has_laurent_expansion F"
+ and infy:"filterlim f at_infinity (at 0)"
+ shows "fls_subdegree F < 0"
+proof (rule ccontr)
+ assume asm:"\<not> fls_subdegree F < 0"
+ define ff where "ff=(\<lambda>z. fls_nth F (fls_subdegree F)
+ * z powi fls_subdegree F)"
+
+ have "(ff \<longlongrightarrow> (if fls_subdegree F =0 then fls_nth F 0 else 0)) (at 0)"
+ using asm unfolding ff_def
+ by (auto intro!: tendsto_eq_intros)
+ moreover have "filterlim ff at_infinity (at 0)"
+ proof (rule asymp_equiv_at_infinity_transfer)
+ show "f \<sim>[at 0] ff" unfolding ff_def
+ using has_laurent_expansion_imp_asymp_equiv_0[OF F] unfolding ff_def .
+ show "filterlim f at_infinity (at 0)" by fact
+ qed
+ ultimately show False
+ using not_tendsto_and_filterlim_at_infinity[of "at (0::complex)"] by auto
+qed
+
+lemma has_laurent_expansion_imp_filterlim_infinity:
+ assumes F: "(\<lambda>w. f (z + w)) has_laurent_expansion F" and "fls_subdegree F < 0"
+ shows "filterlim f at_infinity (at z)"
+ using has_laurent_expansion_imp_filterlim_infinity_0[OF assms]
+ by (simp add: at_to_0[of z] filterlim_filtermap add_ac)
+
+lemma has_laurent_expansion_imp_is_pole_0:
+ assumes F: "f has_laurent_expansion F" and "fls_subdegree F < 0"
+ shows "is_pole f 0"
+ using has_laurent_expansion_imp_filterlim_infinity_0[OF assms]
+ by (simp add: is_pole_def)
+
+lemma is_pole_0_imp_neg_fls_subdegree:
+ assumes F: "f has_laurent_expansion F" and "is_pole f 0"
+ shows "fls_subdegree F < 0"
+ using F assms(2) has_laurent_expansion_imp_neg_fls_subdegree is_pole_def
+ by blast
+
+lemma has_laurent_expansion_imp_is_pole:
+ assumes F: "(\<lambda>x. f (z + x)) has_laurent_expansion F" and "fls_subdegree F < 0"
+ shows "is_pole f z"
+ using has_laurent_expansion_imp_is_pole_0[OF assms]
+ by (simp add: is_pole_shift_0')
+
+lemma is_pole_imp_neg_fls_subdegree:
+ assumes F: "(\<lambda>x. f (z + x)) has_laurent_expansion F" and "is_pole f z"
+ shows "fls_subdegree F < 0"
+ apply (rule is_pole_0_imp_neg_fls_subdegree[OF F])
+ using assms(2) is_pole_shift_0 by blast
+
+lemma is_pole_fls_subdegree_iff:
+ assumes "(\<lambda>x. f (z + x)) has_laurent_expansion F"
+ shows "is_pole f z \<longleftrightarrow> fls_subdegree F < 0"
+ using assms is_pole_imp_neg_fls_subdegree has_laurent_expansion_imp_is_pole
+ by auto
+
+lemma
+ assumes "f has_laurent_expansion F"
+ shows has_laurent_expansion_isolated_0: "isolated_singularity_at f 0"
+ and has_laurent_expansion_not_essential_0: "not_essential f 0"
+proof -
+ from assms have "eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
+ by (auto simp: has_laurent_expansion_def)
+ then obtain r where r: "r > 0" "\<And>z. z \<in> ball 0 r - {0} \<Longrightarrow> eval_fls F z = f z"
+ by (auto simp: eventually_at_filter ball_def eventually_nhds_metric)
+
+ have "fls_conv_radius F > 0"
+ using assms by (auto simp: has_laurent_expansion_def)
+ then obtain R :: real where R: "R > 0" "R \<le> min r (fls_conv_radius F)"
+ using \<open>r > 0\<close> by (metis dual_order.strict_implies_order ereal_dense2 ereal_less(2) min_def)
+
+ have "eval_fls F holomorphic_on ball 0 R - {0}"
+ using r R by (intro holomorphic_intros ball_eball_mono Diff_mono) (auto simp: ereal_le_less)
+ also have "?this \<longleftrightarrow> f holomorphic_on ball 0 R - {0}"
+ using r R by (intro holomorphic_cong) auto
+ also have "\<dots> \<longleftrightarrow> f analytic_on ball 0 R - {0}"
+ by (subst analytic_on_open) auto
+ finally show "isolated_singularity_at f 0"
+ unfolding isolated_singularity_at_def using \<open>R > 0\<close> by blast
+
+ show "not_essential f 0"
+ proof (cases "fls_subdegree F \<ge> 0")
+ case True
+ hence "f \<midarrow>0\<rightarrow> fls_nth F 0"
+ by (intro has_laurent_expansion_imp_tendsto_0[OF assms])
+ thus ?thesis
+ by (auto simp: not_essential_def)
+ next
+ case False
+ hence "is_pole f 0"
+ by (intro has_laurent_expansion_imp_is_pole_0[OF assms]) auto
+ thus ?thesis
+ by (auto simp: not_essential_def)
+ qed
+qed
+
+lemma
+ assumes "(\<lambda>w. f (z + w)) has_laurent_expansion F"
+ shows has_laurent_expansion_isolated: "isolated_singularity_at f z"
+ and has_laurent_expansion_not_essential: "not_essential f z"
+ using has_laurent_expansion_isolated_0[OF assms] has_laurent_expansion_not_essential_0[OF assms]
+ by (simp_all add: isolated_singularity_at_shift_0 not_essential_shift_0)
+
+lemma has_laurent_expansion_fps:
+ assumes "f has_fps_expansion F"
+ shows "f has_laurent_expansion fps_to_fls F"
+proof -
+ from assms have radius: "0 < fps_conv_radius F" and eval: "\<forall>\<^sub>F z in nhds 0. eval_fps F z = f z"
+ by (auto simp: has_fps_expansion_def)
+ from eval have eval': "\<forall>\<^sub>F z in at 0. eval_fps F z = f z"
+ using eventually_at_filter eventually_mono by fastforce
+ moreover have "eventually (\<lambda>z. z \<in> eball 0 (fps_conv_radius F) - {0}) (at 0)"
+ using radius by (intro eventually_at_in_open) (auto simp: zero_ereal_def)
+ ultimately have "eventually (\<lambda>z. eval_fls (fps_to_fls F) z = f z) (at 0)"
+ by eventually_elim (auto simp: eval_fps_to_fls)
+ thus ?thesis using radius
+ by (auto simp: has_laurent_expansion_def)
+qed
+
+lemma has_laurent_expansion_const [simp, intro, laurent_expansion_intros]:
+ "(\<lambda>_. c) has_laurent_expansion fls_const c"
+ by (auto simp: has_laurent_expansion_def)
+
+lemma has_laurent_expansion_0 [simp, intro, laurent_expansion_intros]:
+ "(\<lambda>_. 0) has_laurent_expansion 0"
+ by (auto simp: has_laurent_expansion_def)
+
+lemma has_fps_expansion_0_iff: "f has_fps_expansion 0 \<longleftrightarrow> eventually (\<lambda>z. f z = 0) (nhds 0)"
+ by (auto simp: has_fps_expansion_def)
+
+lemma has_laurent_expansion_1 [simp, intro, laurent_expansion_intros]:
+ "(\<lambda>_. 1) has_laurent_expansion 1"
+ by (auto simp: has_laurent_expansion_def)
+
+lemma has_laurent_expansion_numeral [simp, intro, laurent_expansion_intros]:
+ "(\<lambda>_. numeral n) has_laurent_expansion numeral n"
+ by (auto simp: has_laurent_expansion_def)
+
+lemma has_laurent_expansion_fps_X_power [laurent_expansion_intros]:
+ "(\<lambda>x. x ^ n) has_laurent_expansion (fls_X_intpow n)"
+ by (auto simp: has_laurent_expansion_def)
+
+lemma has_laurent_expansion_fps_X_power_int [laurent_expansion_intros]:
+ "(\<lambda>x. x powi n) has_laurent_expansion (fls_X_intpow n)"
+ by (auto simp: has_laurent_expansion_def)
+
+lemma has_laurent_expansion_fps_X [laurent_expansion_intros]:
+ "(\<lambda>x. x) has_laurent_expansion fls_X"
+ by (auto simp: has_laurent_expansion_def)
+
+lemma has_laurent_expansion_cmult_left [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "(\<lambda>x. c * f x) has_laurent_expansion fls_const c * F"
+proof -
+ from assms have "eventually (\<lambda>z. z \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
+ by (intro eventually_at_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
+ moreover from assms have "eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
+ by (auto simp: has_laurent_expansion_def)
+ ultimately have "eventually (\<lambda>z. eval_fls (fls_const c * F) z = c * f z) (at 0)"
+ by eventually_elim (simp_all add: eval_fls_mult)
+ with assms show ?thesis
+ by (auto simp: has_laurent_expansion_def intro!: less_le_trans[OF _ fls_conv_radius_mult])
+qed
+
+lemma has_laurent_expansion_cmult_right [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "(\<lambda>x. f x * c) has_laurent_expansion F * fls_const c"
+proof -
+ have "F * fls_const c = fls_const c * F"
+ by (intro fls_eqI) (auto simp: mult.commute)
+ with has_laurent_expansion_cmult_left [OF assms] show ?thesis
+ by (simp add: mult.commute)
+qed
+
+lemma has_fps_expansion_scaleR [fps_expansion_intros]:
+ fixes F :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps"
+ shows "f has_fps_expansion F \<Longrightarrow> (\<lambda>x. c *\<^sub>R f x) has_fps_expansion fps_const (of_real c) * F"
+ unfolding scaleR_conv_of_real by (intro fps_expansion_intros)
+
+lemma has_laurent_expansion_scaleR [laurent_expansion_intros]:
+ "f has_laurent_expansion F \<Longrightarrow> (\<lambda>x. c *\<^sub>R f x) has_laurent_expansion fls_const (of_real c) * F"
+ unfolding scaleR_conv_of_real by (intro laurent_expansion_intros)
+
+lemma has_laurent_expansion_minus [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "(\<lambda>x. - f x) has_laurent_expansion -F"
+proof -
+ from assms have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
+ by (intro eventually_at_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
+ moreover from assms have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
+ by (auto simp: has_laurent_expansion_def)
+ ultimately have "eventually (\<lambda>x. eval_fls (-F) x = -f x) (at 0)"
+ by eventually_elim (auto simp: eval_fls_minus)
+ thus ?thesis using assms by (auto simp: has_laurent_expansion_def)
+qed
+
+lemma has_laurent_expansion_add [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F" "g has_laurent_expansion G"
+ shows "(\<lambda>x. f x + g x) has_laurent_expansion F + G"
+proof -
+ from assms have "0 < min (fls_conv_radius F) (fls_conv_radius G)"
+ by (auto simp: has_laurent_expansion_def)
+ also have "\<dots> \<le> fls_conv_radius (F + G)"
+ by (rule fls_conv_radius_add)
+ finally have radius: "\<dots> > 0" .
+
+ from assms have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
+ "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius G) - {0}) (at 0)"
+ by (intro eventually_at_in_open; force simp: has_laurent_expansion_def zero_ereal_def)+
+ moreover have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
+ and "eventually (\<lambda>x. eval_fls G x = g x) (at 0)"
+ using assms by (auto simp: has_laurent_expansion_def)
+ ultimately have "eventually (\<lambda>x. eval_fls (F + G) x = f x + g x) (at 0)"
+ by eventually_elim (auto simp: eval_fls_add)
+ with radius show ?thesis by (auto simp: has_laurent_expansion_def)
+qed
+
+lemma has_laurent_expansion_diff [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F" "g has_laurent_expansion G"
+ shows "(\<lambda>x. f x - g x) has_laurent_expansion F - G"
+ using has_laurent_expansion_add[of f F "\<lambda>x. - g x" "-G"] assms
+ by (simp add: has_laurent_expansion_minus)
+
+lemma has_laurent_expansion_mult [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F" "g has_laurent_expansion G"
+ shows "(\<lambda>x. f x * g x) has_laurent_expansion F * G"
+proof -
+ from assms have "0 < min (fls_conv_radius F) (fls_conv_radius G)"
+ by (auto simp: has_laurent_expansion_def)
+ also have "\<dots> \<le> fls_conv_radius (F * G)"
+ by (rule fls_conv_radius_mult)
+ finally have radius: "\<dots> > 0" .
+
+ from assms have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
+ "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius G) - {0}) (at 0)"
+ by (intro eventually_at_in_open; force simp: has_laurent_expansion_def zero_ereal_def)+
+ moreover have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
+ and "eventually (\<lambda>x. eval_fls G x = g x) (at 0)"
+ using assms by (auto simp: has_laurent_expansion_def)
+ ultimately have "eventually (\<lambda>x. eval_fls (F * G) x = f x * g x) (at 0)"
+ by eventually_elim (auto simp: eval_fls_mult)
+ with radius show ?thesis by (auto simp: has_laurent_expansion_def)
+qed
+
+lemma has_fps_expansion_power [fps_expansion_intros]:
+ fixes F :: "'a :: {banach, real_normed_div_algebra, comm_ring_1} fps"
+ shows "f has_fps_expansion F \<Longrightarrow> (\<lambda>x. f x ^ m) has_fps_expansion F ^ m"
+ by (induction m) (auto intro!: fps_expansion_intros)
+
+lemma has_laurent_expansion_power [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "(\<lambda>x. f x ^ n) has_laurent_expansion F ^ n"
+ by (induction n) (auto intro!: laurent_expansion_intros assms)
+
+lemma has_laurent_expansion_sum [laurent_expansion_intros]:
+ assumes "\<And>x. x \<in> I \<Longrightarrow> f x has_laurent_expansion F x"
+ shows "(\<lambda>y. \<Sum>x\<in>I. f x y) has_laurent_expansion (\<Sum>x\<in>I. F x)"
+ using assms by (induction I rule: infinite_finite_induct) (auto intro!: laurent_expansion_intros)
+
+lemma has_laurent_expansion_prod [laurent_expansion_intros]:
+ assumes "\<And>x. x \<in> I \<Longrightarrow> f x has_laurent_expansion F x"
+ shows "(\<lambda>y. \<Prod>x\<in>I. f x y) has_laurent_expansion (\<Prod>x\<in>I. F x)"
+ using assms by (induction I rule: infinite_finite_induct) (auto intro!: laurent_expansion_intros)
+
+lemma has_laurent_expansion_deriv [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "deriv f has_laurent_expansion fls_deriv F"
+proof -
+ have "eventually (\<lambda>z. z \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
+ using assms by (intro eventually_at_in_open)
+ (auto simp: has_laurent_expansion_def zero_ereal_def)
+ moreover from assms have "eventually (\<lambda>z. eval_fls F z = f z) (at 0)"
+ by (auto simp: has_laurent_expansion_def)
+ then obtain s where "open s" "0 \<in> s" and s: "\<And>w. w \<in> s - {0} \<Longrightarrow> eval_fls F w = f w"
+ by (auto simp: eventually_nhds eventually_at_filter)
+ hence "eventually (\<lambda>w. w \<in> s - {0}) (at 0)"
+ by (intro eventually_at_in_open) auto
+ ultimately have "eventually (\<lambda>z. eval_fls (fls_deriv F) z = deriv f z) (at 0)"
+ proof eventually_elim
+ case (elim z)
+ hence "eval_fls (fls_deriv F) z = deriv (eval_fls F) z"
+ by (simp add: eval_fls_deriv)
+ also have "eventually (\<lambda>w. w \<in> s - {0}) (nhds z)"
+ using elim and \<open>open s\<close> by (intro eventually_nhds_in_open) auto
+ hence "eventually (\<lambda>w. eval_fls F w = f w) (nhds z)"
+ by eventually_elim (use s in auto)
+ hence "deriv (eval_fls F) z = deriv f z"
+ by (intro deriv_cong_ev refl)
+ finally show ?case .
+ qed
+ with assms show ?thesis
+ by (auto simp: has_laurent_expansion_def intro!: less_le_trans[OF _ fls_conv_radius_deriv])
+qed
+
+lemma has_laurent_expansion_shift [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "(\<lambda>x. f x * x powi n) has_laurent_expansion (fls_shift (-n) F)"
+proof -
+ have "eventually (\<lambda>x. x \<in> eball 0 (fls_conv_radius F) - {0}) (at 0)"
+ using assms by (intro eventually_at_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
+ moreover have "eventually (\<lambda>x. eval_fls F x = f x) (at 0)"
+ using assms by (auto simp: has_laurent_expansion_def)
+ ultimately have "eventually (\<lambda>x. eval_fls (fls_shift (-n) F) x = f x * x powi n) (at 0)"
+ by eventually_elim (auto simp: eval_fls_shift assms)
+ with assms show ?thesis by (auto simp: has_laurent_expansion_def)
+qed
+
+lemma has_laurent_expansion_shift' [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "(\<lambda>x. f x * x powi (-n)) has_laurent_expansion (fls_shift n F)"
+ using has_laurent_expansion_shift[OF assms, of "-n"] by simp
+
+
+lemma has_laurent_expansion_deriv':
+ assumes "f has_laurent_expansion F"
+ assumes "open A" "0 \<in> A" "\<And>x. x \<in> A - {0} \<Longrightarrow> (f has_field_derivative f' x) (at x)"
+ shows "f' has_laurent_expansion fls_deriv F"
+proof -
+ have "deriv f has_laurent_expansion fls_deriv F"
+ by (intro laurent_expansion_intros assms)
+ also have "?this \<longleftrightarrow> ?thesis"
+ proof (intro has_laurent_expansion_cong refl)
+ have "eventually (\<lambda>z. z \<in> A - {0}) (at 0)"
+ by (intro eventually_at_in_open assms)
+ thus "eventually (\<lambda>z. deriv f z = f' z) (at 0)"
+ by eventually_elim (auto intro!: DERIV_imp_deriv assms)
+ qed
+ finally show ?thesis .
+qed
+
+definition laurent_expansion :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex fls" where
+ "laurent_expansion f z =
+ (if eventually (\<lambda>z. f z = 0) (at z) then 0
+ else fls_shift (-zorder f z) (fps_to_fls (fps_expansion (zor_poly f z) z)))"
+
+lemma laurent_expansion_cong:
+ assumes "eventually (\<lambda>w. f w = g w) (at z)" "z = z'"
+ shows "laurent_expansion f z = laurent_expansion g z'"
+ unfolding laurent_expansion_def
+ using zor_poly_cong[OF assms(1,2)] zorder_cong[OF assms] assms
+ by (intro if_cong refl) (auto elim: eventually_elim2)
+
+theorem not_essential_has_laurent_expansion_0:
+ assumes "isolated_singularity_at f 0" "not_essential f 0"
+ shows "f has_laurent_expansion laurent_expansion f 0"
+proof (cases "\<exists>\<^sub>F w in at 0. f w \<noteq> 0")
+ case False
+ have "(\<lambda>_. 0) has_laurent_expansion 0"
+ by simp
+ also have "?this \<longleftrightarrow> f has_laurent_expansion 0"
+ using False by (intro has_laurent_expansion_cong) (auto simp: frequently_def)
+ finally show ?thesis
+ using False by (simp add: laurent_expansion_def frequently_def)
+next
+ case True
+ define n where "n = zorder f 0"
+ obtain r where r: "zor_poly f 0 0 \<noteq> 0" "zor_poly f 0 holomorphic_on cball 0 r" "r > 0"
+ "\<forall>w\<in>cball 0 r - {0}. f w = zor_poly f 0 w * w powr of_int n \<and>
+ zor_poly f 0 w \<noteq> 0"
+ using zorder_exist[OF assms True] unfolding n_def by auto
+ have holo: "zor_poly f 0 holomorphic_on ball 0 r"
+ by (rule holomorphic_on_subset[OF r(2)]) auto
+
+ define F where "F = fps_expansion (zor_poly f 0) 0"
+ have F: "zor_poly f 0 has_fps_expansion F"
+ unfolding F_def by (rule has_fps_expansion_fps_expansion[OF _ _ holo]) (use \<open>r > 0\<close> in auto)
+ have "(\<lambda>z. zor_poly f 0 z * z powi n) has_laurent_expansion fls_shift (-n) (fps_to_fls F)"
+ by (intro laurent_expansion_intros has_laurent_expansion_fps[OF F])
+ also have "?this \<longleftrightarrow> f has_laurent_expansion fls_shift (-n) (fps_to_fls F)"
+ by (intro has_laurent_expansion_cong refl eventually_mono[OF eventually_at_in_open[of "ball 0 r"]])
+ (use r in \<open>auto simp: complex_powr_of_int\<close>)
+ finally show ?thesis using True
+ by (simp add: laurent_expansion_def F_def n_def frequently_def)
+qed
+
+lemma not_essential_has_laurent_expansion:
+ assumes "isolated_singularity_at f z" "not_essential f z"
+ shows "(\<lambda>x. f (z + x)) has_laurent_expansion laurent_expansion f z"
+proof -
+ from assms(1) have iso:"isolated_singularity_at (\<lambda>x. f (z + x)) 0"
+ by (simp add: isolated_singularity_at_shift_0)
+ moreover from assms(2) have ness:"not_essential (\<lambda>x. f (z + x)) 0"
+ by (simp add: not_essential_shift_0)
+ ultimately have "(\<lambda>x. f (z + x)) has_laurent_expansion laurent_expansion (\<lambda>x. f (z + x)) 0"
+ by (rule not_essential_has_laurent_expansion_0)
+
+ also have "\<dots> = laurent_expansion f z"
+ proof (cases "\<exists>\<^sub>F w in at z. f w \<noteq> 0")
+ case False
+ then have "\<forall>\<^sub>F w in at z. f w = 0" using not_frequently by force
+ then have "laurent_expansion (\<lambda>x. f (z + x)) 0 = 0"
+ by (smt (verit, best) add.commute eventually_at_to_0 eventually_mono
+ laurent_expansion_def)
+ moreover have "laurent_expansion f z = 0"
+ using \<open>\<forall>\<^sub>F w in at z. f w = 0\<close> unfolding laurent_expansion_def by auto
+ ultimately show ?thesis by auto
+ next
+ case True
+ define df where "df=zor_poly (\<lambda>x. f (z + x)) 0"
+ define g where "g=(\<lambda>u. u-z)"
+
+ have "fps_expansion df 0
+ = fps_expansion (df o g) z"
+ proof -
+ have "\<exists>\<^sub>F w in at 0. f (z + w) \<noteq> 0" using True
+ by (smt (verit, best) add.commute eventually_at_to_0
+ eventually_mono not_frequently)
+ from zorder_exist[OF iso ness this,folded df_def]
+ obtain r where "r>0" and df_holo:"df holomorphic_on cball 0 r" and "df 0 \<noteq> 0"
+ "\<forall>w\<in>cball 0 r - {0}.
+ f (z + w) = df w * w powr of_int (zorder (\<lambda>w. f (z + w)) 0) \<and>
+ df w \<noteq> 0"
+ by auto
+ then have df_nz:"\<forall>w\<in>ball 0 r. df w\<noteq>0" by auto
+
+ have "(deriv ^^ n) df 0 = (deriv ^^ n) (df \<circ> g) z" for n
+ unfolding comp_def g_def
+ proof (subst higher_deriv_compose_linear'[where u=1 and c="-z",simplified])
+ show "df holomorphic_on ball 0 r"
+ using df_holo by auto
+ show "open (ball z r)" "open (ball 0 r)" "z \<in> ball z r"
+ using \<open>r>0\<close> by auto
+ show " \<And>w. w \<in> ball z r \<Longrightarrow> w - z \<in> ball 0 r"
+ by (simp add: dist_norm)
+ qed auto
+ then show ?thesis
+ unfolding fps_expansion_def by auto
+ qed
+ also have "... = fps_expansion (zor_poly f z) z"
+ proof (rule fps_expansion_cong)
+ have "\<forall>\<^sub>F w in nhds z. zor_poly f z w
+ = zor_poly (\<lambda>u. f (z + u)) 0 (w - z)"
+ apply (rule zor_poly_shift)
+ using True assms by auto
+ then show "\<forall>\<^sub>F w in nhds z. (df \<circ> g) w = zor_poly f z w"
+ unfolding df_def g_def comp_def
+ by (auto elim:eventually_mono)
+ qed
+ finally show ?thesis unfolding df_def
+ by (auto simp: laurent_expansion_def at_to_0[of z]
+ eventually_filtermap add_ac zorder_shift')
+ qed
+ finally show ?thesis .
+qed
+
+lemma has_fps_expansion_to_laurent:
+ "f has_fps_expansion F \<longleftrightarrow> f has_laurent_expansion fps_to_fls F \<and> f 0 = fps_nth F 0"
+proof safe
+ assume *: "f has_laurent_expansion fps_to_fls F" "f 0 = fps_nth F 0"
+ have "eventually (\<lambda>z. z \<in> eball 0 (fps_conv_radius F)) (nhds 0)"
+ using * by (intro eventually_nhds_in_open) (auto simp: has_laurent_expansion_def zero_ereal_def)
+ moreover have "eventually (\<lambda>z. z \<noteq> 0 \<longrightarrow> eval_fls (fps_to_fls F) z = f z) (nhds 0)"
+ using * by (auto simp: has_laurent_expansion_def eventually_at_filter)
+ ultimately have "eventually (\<lambda>z. f z = eval_fps F z) (nhds 0)"
+ by eventually_elim
+ (auto simp: has_laurent_expansion_def eventually_at_filter eval_fps_at_0 eval_fps_to_fls *(2))
+ thus "f has_fps_expansion F"
+ using * by (auto simp: has_fps_expansion_def has_laurent_expansion_def eq_commute)
+next
+ assume "f has_fps_expansion F"
+ thus "f 0 = fps_nth F 0"
+ by (metis eval_fps_at_0 has_fps_expansion_imp_holomorphic)
+qed (auto intro: has_laurent_expansion_fps)
+
+lemma eval_fps_fls_base_factor [simp]:
+ assumes "z \<noteq> 0"
+ shows "eval_fps (fls_base_factor_to_fps F) z = eval_fls F z * z powi -fls_subdegree F"
+ using assms unfolding eval_fls_def by (simp add: power_int_minus field_simps)
+
+lemma has_fps_expansion_imp_analytic_0:
+ assumes "f has_fps_expansion F"
+ shows "f analytic_on {0}"
+ by (meson analytic_at_two assms has_fps_expansion_imp_holomorphic)
+
+lemma has_fps_expansion_imp_analytic:
+ assumes "(\<lambda>x. f (z + x)) has_fps_expansion F"
+ shows "f analytic_on {z}"
+proof -
+ have "(\<lambda>x. f (z + x)) analytic_on {0}"
+ by (rule has_fps_expansion_imp_analytic_0) fact
+ hence "(\<lambda>x. f (z + x)) \<circ> (\<lambda>x. x - z) analytic_on {z}"
+ by (intro analytic_on_compose_gen analytic_intros) auto
+ thus ?thesis
+ by (simp add: o_def)
+qed
+
+lemma is_pole_cong_asymp_equiv:
+ assumes "f \<sim>[at z] g" "z = z'"
+ shows "is_pole f z = is_pole g z'"
+ using asymp_equiv_at_infinity_transfer[OF assms(1)]
+ asymp_equiv_at_infinity_transfer[OF asymp_equiv_symI[OF assms(1)]] assms(2)
+ unfolding is_pole_def by auto
+
+lemma not_is_pole_const [simp]: "\<not>is_pole (\<lambda>_::'a::perfect_space. c :: complex) z"
+ using not_tendsto_and_filterlim_at_infinity[of "at z" "\<lambda>_::'a. c" c] by (auto simp: is_pole_def)
+
+lemma has_laurent_expansion_imp_is_pole_iff:
+ assumes F: "(\<lambda>x. f (z + x)) has_laurent_expansion F"
+ shows "is_pole f z \<longleftrightarrow> fls_subdegree F < 0"
+proof
+ assume pole: "is_pole f z"
+ have [simp]: "F \<noteq> 0"
+ proof
+ assume "F = 0"
+ hence "is_pole f z \<longleftrightarrow> is_pole (\<lambda>_. 0 :: complex) z" using assms
+ by (intro is_pole_cong)
+ (auto simp: has_laurent_expansion_def at_to_0[of z] eventually_filtermap add_ac)
+ with pole show False
+ by simp
+ qed
+
+ note pole
+ also have "is_pole f z \<longleftrightarrow>
+ is_pole (\<lambda>w. fls_nth F (fls_subdegree F) * (w - z) powi fls_subdegree F) z"
+ using has_laurent_expansion_imp_asymp_equiv[OF F] by (intro is_pole_cong_asymp_equiv refl)
+ also have "\<dots> \<longleftrightarrow> is_pole (\<lambda>w. (w - z) powi fls_subdegree F) z"
+ by simp
+ finally have pole': \<dots> .
+
+ have False if "fls_subdegree F \<ge> 0"
+ proof -
+ have "(\<lambda>w. (w - z) powi fls_subdegree F) holomorphic_on UNIV"
+ using that by (intro holomorphic_intros) auto
+ hence "\<not>is_pole (\<lambda>w. (w - z) powi fls_subdegree F) z"
+ by (meson UNIV_I not_is_pole_holomorphic open_UNIV)
+ with pole' show False
+ by simp
+ qed
+ thus "fls_subdegree F < 0"
+ by force
+qed (use has_laurent_expansion_imp_is_pole[OF assms] in auto)
+
+lemma analytic_at_imp_has_fps_expansion_0:
+ assumes "f analytic_on {0}"
+ shows "f has_fps_expansion fps_expansion f 0"
+ using assms has_fps_expansion_fps_expansion analytic_at by fast
+
+lemma deriv_shift_0: "deriv f z = deriv (f \<circ> (\<lambda>x. z + x)) 0"
+proof -
+ have *: "(f \<circ> (+) z has_field_derivative D) (at z')"
+ if "(f has_field_derivative D) (at (z + z'))" for D z z' and f :: "'a \<Rightarrow> 'a"
+ proof -
+ have "(f \<circ> (+) z has_field_derivative D * 1) (at z')"
+ by (rule DERIV_chain that derivative_eq_intros refl)+ auto
+ thus ?thesis by simp
+ qed
+ have "(\<lambda>D. (f has_field_derivative D) (at z)) = (\<lambda> D. (f \<circ> (+) z has_field_derivative D) (at 0))"
+ using *[of f _ z 0] *[of "f \<circ> (+) z" _ "-z" z] by (intro ext iffI) (auto simp: o_def)
+ thus ?thesis
+ by (simp add: deriv_def)
+qed
+
+lemma deriv_shift_0': "NO_MATCH 0 z \<Longrightarrow> deriv f z = deriv (f \<circ> (\<lambda>x. z + x)) 0"
+ by (rule deriv_shift_0)
+
+lemma higher_deriv_shift_0: "(deriv ^^ n) f z = (deriv ^^ n) (f \<circ> (\<lambda>x. z + x)) 0"
+proof (induction n arbitrary: f)
+ case (Suc n)
+ have "(deriv ^^ Suc n) f z = (deriv ^^ n) (deriv f) z"
+ by (subst funpow_Suc_right) auto
+ also have "\<dots> = (deriv ^^ n) (\<lambda>x. deriv f (z + x)) 0"
+ by (subst Suc) (auto simp: o_def)
+ also have "\<dots> = (deriv ^^ n) (\<lambda>x. deriv (\<lambda>xa. f (z + x + xa)) 0) 0"
+ by (subst deriv_shift_0) (auto simp: o_def)
+ also have "(\<lambda>x. deriv (\<lambda>xa. f (z + x + xa)) 0) = deriv (\<lambda>x. f (z + x))"
+ by (rule ext) (simp add: deriv_shift_0' o_def add_ac)
+ also have "(deriv ^^ n) \<dots> 0 = (deriv ^^ Suc n) (f \<circ> (\<lambda>x. z + x)) 0"
+ by (subst funpow_Suc_right) (auto simp: o_def)
+ finally show ?case .
+qed auto
+
+lemma higher_deriv_shift_0': "NO_MATCH 0 z \<Longrightarrow> (deriv ^^ n) f z = (deriv ^^ n) (f \<circ> (\<lambda>x. z + x)) 0"
+ by (rule higher_deriv_shift_0)
+
+lemma analytic_at_imp_has_fps_expansion:
+ assumes "f analytic_on {z}"
+ shows "(\<lambda>x. f (z + x)) has_fps_expansion fps_expansion f z"
+proof -
+ have "f \<circ> (\<lambda>x. z + x) analytic_on {0}"
+ by (intro analytic_on_compose_gen[OF _ assms] analytic_intros) auto
+ hence "(f \<circ> (\<lambda>x. z + x)) has_fps_expansion fps_expansion (f \<circ> (\<lambda>x. z + x)) 0"
+ unfolding o_def by (intro analytic_at_imp_has_fps_expansion_0) auto
+ also have "\<dots> = fps_expansion f z"
+ by (simp add: fps_expansion_def higher_deriv_shift_0')
+ finally show ?thesis by (simp add: add_ac)
+qed
+
+lemma has_laurent_expansion_zorder_0:
+ assumes "f has_laurent_expansion F" "F \<noteq> 0"
+ shows "zorder f 0 = fls_subdegree F"
+proof -
+ define G where "G = fls_base_factor_to_fps F"
+ from assms obtain A where A: "0 \<in> A" "open A" "\<And>x. x \<in> A - {0} \<Longrightarrow> eval_fls F x = f x"
+ unfolding has_laurent_expansion_def eventually_at_filter eventually_nhds
+ by blast
+
+ show ?thesis
+ proof (rule zorder_eqI)
+ show "open (A \<inter> eball 0 (fls_conv_radius F))" "0 \<in> A \<inter> eball 0 (fls_conv_radius F)"
+ using assms A by (auto simp: has_laurent_expansion_def zero_ereal_def)
+ show "eval_fps G holomorphic_on A \<inter> eball 0 (fls_conv_radius F)"
+ by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef G_def)
+ show "eval_fps G 0 \<noteq> 0" using \<open>F \<noteq> 0\<close>
+ by (auto simp: eval_fps_at_0 G_def)
+ next
+ fix w :: complex assume "w \<in> A \<inter> eball 0 (fls_conv_radius F)" "w \<noteq> 0"
+ thus "f w = eval_fps G w * (w - 0) powr of_int (fls_subdegree F)"
+ using A unfolding G_def
+ by (subst eval_fps_fls_base_factor)
+ (auto simp: complex_powr_of_int power_int_minus field_simps)
+ qed
+qed
+
+lemma has_laurent_expansion_zorder:
+ assumes "(\<lambda>w. f (z + w)) has_laurent_expansion F" "F \<noteq> 0"
+ shows "zorder f z = fls_subdegree F"
+ using has_laurent_expansion_zorder_0[OF assms] by (simp add: zorder_shift' add_ac)
+
+lemma has_fps_expansion_zorder_0:
+ assumes "f has_fps_expansion F" "F \<noteq> 0"
+ shows "zorder f 0 = int (subdegree F)"
+ using assms has_laurent_expansion_zorder_0[of f "fps_to_fls F"]
+ by (auto simp: has_fps_expansion_to_laurent fls_subdegree_fls_to_fps)
+
+lemma has_fps_expansion_zorder:
+ assumes "(\<lambda>w. f (z + w)) has_fps_expansion F" "F \<noteq> 0"
+ shows "zorder f z = int (subdegree F)"
+ using has_fps_expansion_zorder_0[OF assms]
+ by (simp add: zorder_shift' add_ac)
+
+lemma has_fps_expansion_fls_base_factor_to_fps:
+ assumes "f has_laurent_expansion F"
+ defines "n \<equiv> fls_subdegree F"
+ defines "c \<equiv> fps_nth (fls_base_factor_to_fps F) 0"
+ shows "(\<lambda>z. if z = 0 then c else f z * z powi -n) has_fps_expansion fls_base_factor_to_fps F"
+proof -
+ have "(\<lambda>z. f z * z powi -n) has_laurent_expansion fls_shift (-(-n)) F"
+ by (intro laurent_expansion_intros assms)
+ also have "fls_shift (-(-n)) F = fps_to_fls (fls_base_factor_to_fps F)"
+ by (simp add: n_def fls_shift_nonneg_subdegree)
+ also have "(\<lambda>z. f z * z powi - n) has_laurent_expansion fps_to_fls (fls_base_factor_to_fps F) \<longleftrightarrow>
+ (\<lambda>z. if z = 0 then c else f z * z powi -n) has_laurent_expansion fps_to_fls (fls_base_factor_to_fps F)"
+ by (intro has_laurent_expansion_cong) (auto simp: eventually_at_filter)
+ also have "\<dots> \<longleftrightarrow> (\<lambda>z. if z = 0 then c else f z * z powi -n) has_fps_expansion fls_base_factor_to_fps F"
+ by (subst has_fps_expansion_to_laurent) (auto simp: c_def)
+ finally show ?thesis .
+qed
+
+lemma zero_has_laurent_expansion_imp_eq_0:
+ assumes "(\<lambda>_. 0) has_laurent_expansion F"
+ shows "F = 0"
+proof -
+ have "at (0 :: complex) \<noteq> bot"
+ by auto
+ moreover have "(\<lambda>z. if z = 0 then fls_nth F (fls_subdegree F) else 0) has_fps_expansion
+ fls_base_factor_to_fps F" (is "?f has_fps_expansion _")
+ using has_fps_expansion_fls_base_factor_to_fps[OF assms] by (simp cong: if_cong)
+ hence "isCont ?f 0"
+ using has_fps_expansion_imp_continuous by blast
+ hence "?f \<midarrow>0\<rightarrow> fls_nth F (fls_subdegree F)"
+ by (auto simp: isCont_def)
+ moreover have "?f \<midarrow>0\<rightarrow> 0 \<longleftrightarrow> (\<lambda>_::complex. 0 :: complex) \<midarrow>0\<rightarrow> 0"
+ by (intro filterlim_cong) (auto simp: eventually_at_filter)
+ hence "?f \<midarrow>0\<rightarrow> 0"
+ by simp
+ ultimately have "fls_nth F (fls_subdegree F) = 0"
+ by (rule tendsto_unique)
+ thus ?thesis
+ by (meson nth_fls_subdegree_nonzero)
+qed
+
+lemma has_laurent_expansion_unique:
+ assumes "f has_laurent_expansion F" "f has_laurent_expansion G"
+ shows "F = G"
+proof -
+ from assms have "(\<lambda>x. f x - f x) has_laurent_expansion F - G"
+ by (intro laurent_expansion_intros)
+ hence "(\<lambda>_. 0) has_laurent_expansion F - G"
+ by simp
+ hence "F - G = 0"
+ by (rule zero_has_laurent_expansion_imp_eq_0)
+ thus ?thesis
+ by simp
+qed
+
+lemma laurent_expansion_eqI:
+ assumes "(\<lambda>x. f (z + x)) has_laurent_expansion F"
+ shows "laurent_expansion f z = F"
+ using assms has_laurent_expansion_isolated has_laurent_expansion_not_essential
+ has_laurent_expansion_unique not_essential_has_laurent_expansion by blast
+
+lemma laurent_expansion_0_eqI:
+ assumes "f has_laurent_expansion F"
+ shows "laurent_expansion f 0 = F"
+ using assms laurent_expansion_eqI[of f 0] by simp
+
+lemma has_laurent_expansion_nonzero_imp_eventually_nonzero:
+ assumes "f has_laurent_expansion F" "F \<noteq> 0"
+ shows "eventually (\<lambda>x. f x \<noteq> 0) (at 0)"
+proof (rule ccontr)
+ assume "\<not>eventually (\<lambda>x. f x \<noteq> 0) (at 0)"
+ with assms have "eventually (\<lambda>x. f x = 0) (at 0)"
+ by (intro not_essential_frequently_0_imp_eventually_0 has_laurent_expansion_isolated
+ has_laurent_expansion_not_essential)
+ (auto simp: frequently_def)
+ hence "(f has_laurent_expansion 0) \<longleftrightarrow> ((\<lambda>_. 0) has_laurent_expansion 0)"
+ by (intro has_laurent_expansion_cong) auto
+ hence "f has_laurent_expansion 0"
+ by simp
+ with assms(1) have "F = 0"
+ using has_laurent_expansion_unique by blast
+ with \<open>F \<noteq> 0\<close> show False
+ by contradiction
+qed
+
+lemma has_laurent_expansion_eventually_nonzero_iff':
+ assumes "f has_laurent_expansion F"
+ shows "eventually (\<lambda>x. f x \<noteq> 0) (at 0) \<longleftrightarrow> F \<noteq> 0 "
+proof
+ assume "\<forall>\<^sub>F x in at 0. f x \<noteq> 0"
+ moreover have "\<not> (\<forall>\<^sub>F x in at 0. f x \<noteq> 0)" if "F=0"
+ proof -
+ have "\<forall>\<^sub>F x in at 0. f x = 0"
+ using assms that unfolding has_laurent_expansion_def by simp
+ then show ?thesis unfolding not_eventually
+ by (auto elim:eventually_frequentlyE)
+ qed
+ ultimately show "F \<noteq> 0" by auto
+qed (simp add:has_laurent_expansion_nonzero_imp_eventually_nonzero[OF assms])
+
+lemma has_laurent_expansion_eventually_nonzero_iff:
+ assumes "(\<lambda>w. f (z+w)) has_laurent_expansion F"
+ shows "eventually (\<lambda>x. f x \<noteq> 0) (at z) \<longleftrightarrow> F \<noteq> 0"
+ apply (subst eventually_at_to_0)
+ apply (rule has_laurent_expansion_eventually_nonzero_iff')
+ using assms by (simp add:add.commute)
+
+lemma has_laurent_expansion_inverse [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F"
+ shows "(\<lambda>x. inverse (f x)) has_laurent_expansion inverse F"
+proof (cases "F = 0")
+ case True
+ thus ?thesis using assms
+ by (auto simp: has_laurent_expansion_def)
+next
+ case False
+ define G where "G = laurent_expansion (\<lambda>x. inverse (f x)) 0"
+ from False have ev: "eventually (\<lambda>z. f z \<noteq> 0) (at 0)"
+ by (intro has_laurent_expansion_nonzero_imp_eventually_nonzero[OF assms])
+
+ have *: "(\<lambda>x. inverse (f x)) has_laurent_expansion G" unfolding G_def
+ by (intro not_essential_has_laurent_expansion_0 isolated_singularity_at_inverse not_essential_inverse
+ has_laurent_expansion_isolated_0[OF assms] has_laurent_expansion_not_essential_0[OF assms])
+ have "(\<lambda>x. f x * inverse (f x)) has_laurent_expansion F * G"
+ by (intro laurent_expansion_intros assms *)
+ also have "?this \<longleftrightarrow> (\<lambda>x. 1) has_laurent_expansion F * G"
+ by (intro has_laurent_expansion_cong refl eventually_mono[OF ev]) auto
+ finally have "(\<lambda>_. 1) has_laurent_expansion F * G" .
+ moreover have "(\<lambda>_. 1) has_laurent_expansion 1"
+ by simp
+ ultimately have "F * G = 1"
+ using has_laurent_expansion_unique by blast
+ hence "G = inverse F"
+ using inverse_unique by blast
+ with * show ?thesis
+ by simp
+qed
+
+lemma has_laurent_expansion_power_int [laurent_expansion_intros]:
+ "f has_laurent_expansion F \<Longrightarrow> (\<lambda>x. f x powi n) has_laurent_expansion (F powi n)"
+ by (auto simp: power_int_def intro!: laurent_expansion_intros)
+
+
+lemma has_fps_expansion_0_analytic_continuation:
+ assumes "f has_fps_expansion 0" "f holomorphic_on A"
+ assumes "open A" "connected A" "0 \<in> A" "x \<in> A"
+ shows "f x = 0"
+proof -
+ have "eventually (\<lambda>z. z \<in> A \<and> f z = 0) (nhds 0)" using assms
+ by (intro eventually_conj eventually_nhds_in_open) (auto simp: has_fps_expansion_def)
+ then obtain B where B: "open B" "0 \<in> B" "\<forall>z\<in>B. z \<in> A \<and> f z = 0"
+ unfolding eventually_nhds by blast
+ show ?thesis
+ proof (rule analytic_continuation_open[where f = f and g = "\<lambda>_. 0"])
+ show "B \<noteq> {}"
+ using \<open>open B\<close> B by auto
+ show "connected A"
+ using assms by auto
+ qed (use assms B in auto)
+qed
+
+lemma has_laurent_expansion_0_analytic_continuation:
+ assumes "f has_laurent_expansion 0" "f holomorphic_on A - {0}"
+ assumes "open A" "connected A" "0 \<in> A" "x \<in> A - {0}"
+ shows "f x = 0"
+proof -
+ have "eventually (\<lambda>z. z \<in> A - {0} \<and> f z = 0) (at 0)" using assms
+ by (intro eventually_conj eventually_at_in_open) (auto simp: has_laurent_expansion_def)
+ then obtain B where B: "open B" "0 \<in> B" "\<forall>z\<in>B - {0}. z \<in> A - {0} \<and> f z = 0"
+ unfolding eventually_at_filter eventually_nhds by blast
+ show ?thesis
+ proof (rule analytic_continuation_open[where f = f and g = "\<lambda>_. 0"])
+ show "B - {0} \<noteq> {}"
+ using \<open>open B\<close> \<open>0 \<in> B\<close> by (metis insert_Diff not_open_singleton)
+ show "connected (A - {0})"
+ using assms by (intro connected_open_delete) auto
+ qed (use assms B in auto)
+qed
+
+lemma has_fps_expansion_cong:
+ assumes "eventually (\<lambda>x. f x = g x) (nhds 0)" "F = G"
+ shows "f has_fps_expansion F \<longleftrightarrow> g has_fps_expansion G"
+ using assms(2) by (auto simp: has_fps_expansion_def elim!: eventually_elim2[OF assms(1)])
+
+lemma zor_poly_has_fps_expansion:
+ assumes "f has_laurent_expansion F" "F \<noteq> 0"
+ shows "zor_poly f 0 has_fps_expansion fls_base_factor_to_fps F"
+proof -
+ note [simp] = \<open>F \<noteq> 0\<close>
+ have "eventually (\<lambda>z. f z \<noteq> 0) (at 0)"
+ by (rule has_laurent_expansion_nonzero_imp_eventually_nonzero[OF assms])
+ hence freq: "frequently (\<lambda>z. f z \<noteq> 0) (at 0)"
+ by (rule eventually_frequently[rotated]) auto
+
+ have *: "isolated_singularity_at f 0" "not_essential f 0"
+ using has_laurent_expansion_isolated_0[OF assms(1)] has_laurent_expansion_not_essential_0[OF assms(1)]
+ by auto
+
+ define G where "G = fls_base_factor_to_fps F"
+ define n where "n = zorder f 0"
+ have n_altdef: "n = fls_subdegree F"
+ using has_laurent_expansion_zorder_0 [OF assms(1)] by (simp add: n_def)
+ obtain r where r: "zor_poly f 0 0 \<noteq> 0" "zor_poly f 0 holomorphic_on cball 0 r" "r > 0"
+ "\<forall>w\<in>cball 0 r - {0}. f w = zor_poly f 0 w * w powr of_int n \<and>
+ zor_poly f 0 w \<noteq> 0"
+ using zorder_exist[OF * freq] unfolding n_def by auto
+ obtain r' where r': "r' > 0" "\<forall>x\<in>ball 0 r'-{0}. eval_fls F x = f x"
+ using assms(1) unfolding has_laurent_expansion_def eventually_at_filter eventually_nhds_metric ball_def
+ by (auto simp: dist_commute)
+ have holo: "zor_poly f 0 holomorphic_on ball 0 r"
+ by (rule holomorphic_on_subset[OF r(2)]) auto
+
+ have "(\<lambda>z. if z = 0 then fps_nth G 0 else f z * z powi -n) has_fps_expansion G"
+ unfolding G_def n_altdef by (intro has_fps_expansion_fls_base_factor_to_fps assms)
+ also have "?this \<longleftrightarrow> zor_poly f 0 has_fps_expansion G"
+ proof (intro has_fps_expansion_cong)
+ have "eventually (\<lambda>z. z \<in> ball 0 (min r r')) (nhds 0)"
+ using \<open>r > 0\<close> \<open>r' > 0\<close> by (intro eventually_nhds_in_open) auto
+ thus "\<forall>\<^sub>F x in nhds 0. (if x = 0 then G $ 0 else f x * x powi - n) = zor_poly f 0 x"
+ proof eventually_elim
+ case (elim w)
+ have w: "w \<in> ball 0 r" "w \<in> ball 0 r'"
+ using elim by auto
+ show ?case
+ proof (cases "w = 0")
+ case False
+ hence "f w = zor_poly f 0 w * w powr of_int n"
+ using r w by auto
+ thus ?thesis using False
+ by (simp add: powr_minus complex_powr_of_int power_int_minus)
+ next
+ case [simp]: True
+ obtain R where R: "R > 0" "R \<le> r" "R \<le> r'" "R \<le> fls_conv_radius F"
+ using \<open>r > 0\<close> \<open>r' > 0\<close> assms(1) unfolding has_laurent_expansion_def
+ by (smt (verit, ccfv_SIG) ereal_dense2 ereal_less(2) less_ereal.simps(1) order.strict_implies_order order_trans)
+ have "eval_fps G 0 = zor_poly f 0 0"
+ proof (rule analytic_continuation_open[where f = "eval_fps G" and g = "zor_poly f 0"])
+ show "connected (ball 0 R :: complex set)"
+ by auto
+ have "of_real R / 2 \<in> ball 0 R - {0 :: complex}"
+ using R by auto
+ thus "ball 0 R - {0 :: complex} \<noteq> {}"
+ by blast
+ show "eval_fps G holomorphic_on ball 0 R"
+ using R less_le_trans[OF _ R(4)] unfolding G_def
+ by (intro holomorphic_intros) (auto simp: fls_conv_radius_altdef)
+ show "zor_poly f 0 holomorphic_on ball 0 R"
+ by (rule holomorphic_on_subset[OF holo]) (use R in auto)
+ show "eval_fps G z = zor_poly f 0 z" if "z \<in> ball 0 R - {0}" for z
+ using that r r' R n_altdef unfolding G_def
+ by (subst eval_fps_fls_base_factor)
+ (auto simp: complex_powr_of_int field_simps power_int_minus n_def)
+ qed (use R in auto)
+ hence "zor_poly f 0 0 = fps_nth G 0"
+ by (simp add: eval_fps_at_0)
+ thus ?thesis by simp
+ qed
+ qed
+ qed (use r' in auto)
+ finally show ?thesis
+ by (simp add: G_def)
+qed
+
+lemma zorder_geI_0:
+ assumes "f analytic_on {0}" "f holomorphic_on A" "open A" "connected A" "0 \<in> A" "z \<in> A" "f z \<noteq> 0"
+ assumes "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f 0 = 0"
+ shows "zorder f 0 \<ge> n"
+proof -
+ define F where "F = fps_expansion f 0"
+ from assms have "f has_fps_expansion F"
+ unfolding F_def using analytic_at_imp_has_fps_expansion_0 by blast
+ hence laurent: "f has_laurent_expansion fps_to_fls F" and [simp]: "f 0 = fps_nth F 0"
+ by (simp_all add: has_fps_expansion_to_laurent)
+
+ have [simp]: "F \<noteq> 0"
+ proof
+ assume [simp]: "F = 0"
+ hence "f z = 0"
+ proof (cases "z = 0")
+ case False
+ have "f has_laurent_expansion 0"
+ using laurent by simp
+ thus ?thesis
+ proof (rule has_laurent_expansion_0_analytic_continuation)
+ show "f holomorphic_on A - {0}"
+ using assms(2) by (rule holomorphic_on_subset) auto
+ qed (use assms False in auto)
+ qed auto
+ with \<open>f z \<noteq> 0\<close> show False by contradiction
+ qed
+
+ have "zorder f 0 = int (subdegree F)"
+ using has_laurent_expansion_zorder_0[OF laurent] by (simp add: fls_subdegree_fls_to_fps)
+ also have "subdegree F \<ge> n"
+ using assms by (intro subdegree_geI \<open>F \<noteq> 0\<close>) (auto simp: F_def fps_expansion_def)
+ hence "int (subdegree F) \<ge> int n"
+ by simp
+ finally show ?thesis .
+qed
+
+lemma zorder_geI:
+ assumes "f analytic_on {x}" "f holomorphic_on A" "open A" "connected A" "x \<in> A" "z \<in> A" "f z \<noteq> 0"
+ assumes "\<And>k. k < n \<Longrightarrow> (deriv ^^ k) f x = 0"
+ shows "zorder f x \<ge> n"
+proof -
+ have "zorder f x = zorder (f \<circ> (\<lambda>u. u + x)) 0"
+ by (subst zorder_shift) (auto simp: o_def)
+ also have "\<dots> \<ge> n"
+ proof (rule zorder_geI_0)
+ show "(f \<circ> (\<lambda>u. u + x)) analytic_on {0}"
+ by (intro analytic_on_compose_gen[OF _ assms(1)] analytic_intros) auto
+ show "f \<circ> (\<lambda>u. u + x) holomorphic_on ((+) (-x)) ` A"
+ by (intro holomorphic_on_compose_gen[OF _ assms(2)] holomorphic_intros) auto
+ show "connected ((+) (- x) ` A)"
+ by (intro connected_continuous_image continuous_intros assms)
+ show "open ((+) (- x) ` A)"
+ by (intro open_translation assms)
+ show "z - x \<in> (+) (- x) ` A"
+ using \<open>z \<in> A\<close> by auto
+ show "0 \<in> (+) (- x) ` A"
+ using \<open>x \<in> A\<close> by auto
+ show "(f \<circ> (\<lambda>u. u + x)) (z - x) \<noteq> 0"
+ using \<open>f z \<noteq> 0\<close> by auto
+ next
+ fix k :: nat assume "k < n"
+ hence "(deriv ^^ k) f x = 0"
+ using assms by simp
+ also have "(deriv ^^ k) f x = (deriv ^^ k) (f \<circ> (+) x) 0"
+ by (subst higher_deriv_shift_0) auto
+ finally show "(deriv ^^ k) (f \<circ> (\<lambda>u. u + x)) 0 = 0"
+ by (subst add.commute) auto
+ qed
+ finally show ?thesis .
+qed
+
+lemma has_laurent_expansion_divide [laurent_expansion_intros]:
+ assumes "f has_laurent_expansion F" and "g has_laurent_expansion G"
+ shows "(\<lambda>x. f x / g x) has_laurent_expansion (F / G)"
+proof -
+ have "(\<lambda>x. f x * inverse (g x)) has_laurent_expansion (F * inverse G)"
+ by (intro laurent_expansion_intros assms)
+ thus ?thesis
+ by (simp add: field_simps)
+qed
+
+lemma vector_derivative_translate [simp]:
+ "vector_derivative ((+) z \<circ> g) (at x within A) = vector_derivative g (at x within A)"
+proof -
+ have "(((+) z \<circ> g) has_vector_derivative g') (at x within A)"
+ if "(g has_vector_derivative g') (at x within A)" for g :: "real \<Rightarrow> 'a" and z g'
+ unfolding o_def using that by (auto intro!: derivative_eq_intros)
+ from this[of g _ z] this[of "\<lambda>x. z + g x" _ "-z"] show ?thesis
+ unfolding vector_derivative_def
+ by (intro arg_cong[where f = Eps] ext) (auto simp: o_def algebra_simps)
+qed
+
+lemma has_contour_integral_translate:
+ "(f has_contour_integral I) ((+) z \<circ> g) \<longleftrightarrow> ((\<lambda>x. f (x + z)) has_contour_integral I) g"
+ by (simp add: has_contour_integral_def add_ac)
+
+lemma contour_integrable_translate:
+ "f contour_integrable_on ((+) z \<circ> g) \<longleftrightarrow> (\<lambda>x. f (x + z)) contour_integrable_on g"
+ by (simp add: contour_integrable_on_def has_contour_integral_translate)
+
+lemma contour_integral_translate:
+ "contour_integral ((+) z \<circ> g) f = contour_integral g (\<lambda>x. f (x + z))"
+ by (simp add: contour_integral_def contour_integrable_translate has_contour_integral_translate)
+
+lemma residue_shift_0: "residue f z = residue (\<lambda>x. f (z + x)) 0"
+proof -
+ define Q where
+ "Q = (\<lambda>r f z \<epsilon>. (f has_contour_integral complex_of_real (2 * pi) * \<i> * r) (circlepath z \<epsilon>))"
+ define P where
+ "P = (\<lambda>r f z. \<exists>e>0. \<forall>\<epsilon>>0. \<epsilon> < e \<longrightarrow> Q r f z \<epsilon>)"
+ have path_eq: "circlepath (z - w) \<epsilon> = (+) (-w) \<circ> circlepath z \<epsilon>" for z w \<epsilon>
+ by (simp add: circlepath_def o_def part_circlepath_def algebra_simps)
+ have *: "P r f z" if "P r (\<lambda>x. f (x + w)) (z - w)" for r w f z
+ using that by (auto simp: P_def Q_def path_eq has_contour_integral_translate)
+ have "(SOME r. P r f z) = (SOME r. P r (\<lambda>x. f (z + x)) 0)"
+ using *[of _ f z z] *[of _ "\<lambda>x. f (z + x)" "-z"]
+ by (intro arg_cong[where f = Eps] ext iffI) (simp_all add: add_ac)
+ thus ?thesis
+ by (simp add: residue_def P_def Q_def)
+qed
+
+lemma residue_shift_0': "NO_MATCH 0 z \<Longrightarrow> residue f z = residue (\<lambda>x. f (z + x)) 0"
+ by (rule residue_shift_0)
+
+lemma has_laurent_expansion_residue_0:
+ assumes "f has_laurent_expansion F"
+ shows "residue f 0 = fls_residue F"
+proof (cases "fls_subdegree F \<ge> 0")
+ case True
+ have "residue f 0 = residue (eval_fls F) 0"
+ using assms by (intro residue_cong) (auto simp: has_laurent_expansion_def eq_commute)
+ also have "\<dots> = 0"
+ by (rule residue_holo[OF _ _ holomorphic_on_eval_fls[OF order.refl]])
+ (use True assms in \<open>auto simp: has_laurent_expansion_def zero_ereal_def\<close>)
+ also have "\<dots> = fls_residue F"
+ using True by simp
+ finally show ?thesis .
+next
+ case False
+ hence "F \<noteq> 0"
+ by auto
+ have *: "zor_poly f 0 has_fps_expansion fls_base_factor_to_fps F"
+ by (intro zor_poly_has_fps_expansion False assms \<open>F \<noteq> 0\<close>)
+
+ have "residue f 0 = (deriv ^^ (nat (-zorder f 0) - 1)) (zor_poly f 0) 0 / fact (nat (- zorder f 0) - 1)"
+ by (intro residue_pole_order has_laurent_expansion_isolated_0[OF assms]
+ has_laurent_expansion_imp_is_pole_0[OF assms]) (use False in auto)
+ also have "\<dots> = fls_residue F"
+ using has_laurent_expansion_zorder_0[OF assms \<open>F \<noteq> 0\<close>] False
+ by (subst fps_nth_fps_expansion [OF *, symmetric]) (auto simp: of_nat_diff)
+ finally show ?thesis .
+qed
+
+lemma has_laurent_expansion_residue:
+ assumes "(\<lambda>x. f (z + x)) has_laurent_expansion F"
+ shows "residue f z = fls_residue F"
+ using has_laurent_expansion_residue_0[OF assms] by (simp add: residue_shift_0')
+
+lemma eval_fls_has_laurent_expansion [laurent_expansion_intros]:
+ assumes "fls_conv_radius F > 0"
+ shows "eval_fls F has_laurent_expansion F"
+ using assms by (auto simp: has_laurent_expansion_def)
+
+lemma fps_expansion_unique_complex:
+ fixes F G :: "complex fps"
+ assumes "f has_fps_expansion F" "f has_fps_expansion G"
+ shows "F = G"
+ using assms unfolding fps_eq_iff by (auto simp: fps_eq_iff fps_nth_fps_expansion)
+
+lemma fps_expansion_eqI:
+ assumes "f has_fps_expansion F"
+ shows "fps_expansion f 0 = F"
+ using assms unfolding fps_eq_iff
+ by (auto simp: fps_eq_iff fps_nth_fps_expansion fps_expansion_def)
+
+lemma has_fps_expansion_imp_eval_fps_eq:
+ assumes "f has_fps_expansion F" "norm z < r"
+ assumes "f holomorphic_on ball 0 r"
+ shows "eval_fps F z = f z"
+proof -
+ have [simp]: "fps_expansion f 0 = F"
+ by (rule fps_expansion_eqI) fact
+ have *: "f holomorphic_on eball 0 (ereal r)"
+ using assms by simp
+ from conv_radius_fps_expansion[OF *] have "fps_conv_radius F \<ge> ereal r"
+ by simp
+ have "eval_fps (fps_expansion f 0) z = f (0 + z)"
+ by (rule eval_fps_expansion'[OF *]) (use assms in auto)
+ thus ?thesis
+ by simp
+qed
+
+lemma fls_conv_radius_ge:
+ assumes "f has_laurent_expansion F"
+ assumes "f holomorphic_on eball 0 r - {0}"
+ shows "fls_conv_radius F \<ge> r"
+proof -
+ define n where "n = fls_subdegree F"
+ define G where "G = fls_base_factor_to_fps F"
+ define g where "g = (\<lambda>z. if z = 0 then fps_nth G 0 else f z * z powi -n)"
+ have G: "g has_fps_expansion G"
+ unfolding G_def g_def n_def
+ by (intro has_fps_expansion_fls_base_factor_to_fps assms)
+ have "(\<lambda>z. f z * z powi -n) holomorphic_on eball 0 r - {0}"
+ by (intro holomorphic_intros assms) auto
+ also have "?this \<longleftrightarrow> g holomorphic_on eball 0 r - {0}"
+ by (intro holomorphic_cong) (auto simp: g_def)
+ finally have "g analytic_on eball 0 r - {0}"
+ by (subst analytic_on_open) auto
+ moreover have "g analytic_on {0}"
+ using G has_fps_expansion_imp_analytic_0 by auto
+ ultimately have "g analytic_on (eball 0 r - {0} \<union> {0})"
+ by (subst analytic_on_Un) auto
+ hence "g analytic_on eball 0 r"
+ by (rule analytic_on_subset) auto
+ hence "g holomorphic_on eball 0 r"
+ by (subst (asm) analytic_on_open) auto
+ hence "fps_conv_radius (fps_expansion g 0) \<ge> r"
+ by (intro conv_radius_fps_expansion)
+ also have "fps_expansion g 0 = G"
+ using G by (intro fps_expansion_eqI)
+ finally show ?thesis
+ by (simp add: fls_conv_radius_altdef G_def)
+qed
+
+lemma connected_eball [intro]: "connected (eball (z :: 'a :: real_normed_vector) r)"
+ by (cases r) auto
+
+lemma eval_fls_eqI:
+ assumes "f has_laurent_expansion F" "f holomorphic_on eball 0 r - {0}"
+ assumes "z \<in> eball 0 r - {0}"
+ shows "eval_fls F z = f z"
+proof -
+ have conv: "fls_conv_radius F \<ge> r"
+ by (intro fls_conv_radius_ge[OF assms(1,2)])
+ have "(\<lambda>z. eval_fls F z - f z) has_laurent_expansion F - F"
+ using assms by (intro laurent_expansion_intros assms) (auto simp: has_laurent_expansion_def)
+ hence "(\<lambda>z. eval_fls F z - f z) has_laurent_expansion 0"
+ by simp
+ hence "eval_fls F z - f z = 0"
+ proof (rule has_laurent_expansion_0_analytic_continuation)
+ have "ereal 0 \<le> ereal (norm z)"
+ by simp
+ also have "norm z < r"
+ using assms by auto
+ finally have "r > 0"
+ by (simp add: zero_ereal_def)
+ thus "open (eball 0 r :: complex set)" "connected (eball 0 r :: complex set)"
+ "0 \<in> eball 0 r" "z \<in> eball 0 r - {0}"
+ using assms by (auto simp: zero_ereal_def)
+ qed (auto intro!: holomorphic_intros assms less_le_trans[OF _ conv] split: if_splits)
+ thus ?thesis by simp
+qed
+
+lemma fls_nth_as_contour_integral:
+ assumes F: "f has_laurent_expansion F"
+ assumes holo: "f holomorphic_on ball 0 r - {0}"
+ assumes R: "0 < R" "R < r"
+ shows "((\<lambda>z. f z * z powi (-(n+1))) has_contour_integral
+ complex_of_real (2 * pi) * \<i> * fls_nth F n) (circlepath 0 R)"
+proof -
+ define I where "I = (\<lambda>z. f z * z powi (-(n+1)))"
+ have "(I has_contour_integral complex_of_real (2 * pi) * \<i> * residue I 0) (circlepath 0 R)"
+ proof (rule base_residue)
+ show "open (ball (0::complex) r)" "0 \<in> ball (0::complex) r"
+ using R F by (auto simp: has_laurent_expansion_def zero_ereal_def)
+ qed (use R in \<open>auto intro!: holomorphic_intros holomorphic_on_subset[OF holo]
+ simp: I_def split: if_splits\<close>)
+ also have "residue I 0 = fls_residue (fls_shift (n + 1) F)"
+ unfolding I_def by (intro has_laurent_expansion_residue_0 laurent_expansion_intros F)
+ also have "\<dots> = fls_nth F n"
+ by simp
+ finally show ?thesis
+ by (simp add: I_def)
+qed
+
+lemma tendsto_0_subdegree_iff_0:
+ assumes F:"f has_laurent_expansion F" and "F\<noteq>0"
+ shows "(f \<midarrow>0\<rightarrow>0) \<longleftrightarrow> fls_subdegree F > 0"
+proof -
+ have ?thesis if "is_pole f 0"
+ proof -
+ have "fls_subdegree F <0"
+ using is_pole_0_imp_neg_fls_subdegree[OF F that] .
+ moreover then have "\<not> f \<midarrow>0\<rightarrow>0"
+ using \<open>is_pole f 0\<close> F at_neq_bot
+ has_laurent_expansion_imp_filterlim_infinity_0
+ not_tendsto_and_filterlim_at_infinity that
+ by blast
+ ultimately show ?thesis by auto
+ qed
+ moreover have ?thesis if "\<not>is_pole f 0" "\<exists>x. f \<midarrow>0\<rightarrow>x"
+ proof -
+ have "fls_subdegree F \<ge>0"
+ using has_laurent_expansion_imp_is_pole_0[OF F] that(1)
+ by linarith
+ have "f \<midarrow>0\<rightarrow>0" if "fls_subdegree F > 0"
+ using fls_eq0_below_subdegree[OF that]
+ by (metis F \<open>0 \<le> fls_subdegree F\<close> has_laurent_expansion_imp_tendsto_0)
+ moreover have "fls_subdegree F > 0" if "f \<midarrow>0\<rightarrow>0"
+ proof -
+ have False if "fls_subdegree F = 0"
+ proof -
+ have "f \<midarrow>0\<rightarrow> fls_nth F 0"
+ using has_laurent_expansion_imp_tendsto_0
+ [OF F \<open>fls_subdegree F \<ge>0\<close>] .
+ then have "fls_nth F 0 = 0" using \<open>f \<midarrow>0\<rightarrow>0\<close>
+ using LIM_unique by blast
+ then have "F = 0"
+ using nth_fls_subdegree_zero_iff \<open>fls_subdegree F = 0\<close>
+ by metis
+ with \<open>F\<noteq>0\<close> show False by auto
+ qed
+ with \<open>fls_subdegree F \<ge>0\<close>
+ show ?thesis by fastforce
+ qed
+ ultimately show ?thesis by auto
+ qed
+ moreover have "is_pole f 0 \<or> (\<exists>x. f \<midarrow>0\<rightarrow>x)"
+ proof -
+ have "not_essential f 0"
+ using F has_laurent_expansion_not_essential_0 by auto
+ then show ?thesis unfolding not_essential_def
+ by auto
+ qed
+ ultimately show ?thesis by auto
+qed
+
+lemma tendsto_0_subdegree_iff:
+ assumes F:"(\<lambda>w. f (z+w)) has_laurent_expansion F" and "F\<noteq>0"
+ shows "(f \<midarrow>z\<rightarrow>0) \<longleftrightarrow> fls_subdegree F > 0"
+ apply (subst Lim_at_zero)
+ apply (rule tendsto_0_subdegree_iff_0)
+ using assms by auto
+
+lemma is_pole_0_deriv_divide_iff:
+ assumes F:"f has_laurent_expansion F" and "F\<noteq>0"
+ shows "is_pole (\<lambda>x. deriv f x / f x) 0 \<longleftrightarrow> is_pole f 0 \<or> (f \<midarrow>0\<rightarrow>0)"
+proof -
+ have "(\<lambda>x. deriv f x / f x) has_laurent_expansion fls_deriv F / F"
+ using F by (auto intro:laurent_expansion_intros)
+
+ have "is_pole (\<lambda>x. deriv f x / f x) 0 \<longleftrightarrow>
+ fls_subdegree (fls_deriv F / F) < 0"
+ apply (rule is_pole_fls_subdegree_iff)
+ using F by (auto intro:laurent_expansion_intros)
+ also have "... \<longleftrightarrow> is_pole f 0 \<or> (f \<midarrow>0\<rightarrow>0)"
+ proof (cases "fls_subdegree F = 0")
+ case True
+ then have "fls_subdegree (fls_deriv F / F) \<ge> 0"
+ by (metis diff_zero div_0 \<open>F\<noteq>0\<close> fls_deriv_subdegree0
+ fls_divide_subdegree)
+ moreover then have "\<not> is_pole f 0"
+ by (metis F True is_pole_0_imp_neg_fls_subdegree less_le)
+ moreover have "\<not> (f \<midarrow>0\<rightarrow>0)"
+ using tendsto_0_subdegree_iff_0[OF F \<open>F\<noteq>0\<close>] True by auto
+ ultimately show ?thesis by auto
+ next
+ case False
+ then have "fls_deriv F \<noteq> 0"
+ by (metis fls_const_subdegree fls_deriv_eq_0_iff)
+ then have "fls_subdegree (fls_deriv F / F) =
+ fls_subdegree (fls_deriv F) - fls_subdegree F"
+ by (rule fls_divide_subdegree[OF _ \<open>F\<noteq>0\<close>])
+ moreover have "fls_subdegree (fls_deriv F) = fls_subdegree F - 1"
+ using fls_subdegree_deriv[OF False] .
+ ultimately have "fls_subdegree (fls_deriv F / F) < 0" by auto
+ moreover have "f \<midarrow>0\<rightarrow> 0 = (0 < fls_subdegree F)"
+ using tendsto_0_subdegree_iff_0[OF F \<open>F \<noteq> 0\<close>] .
+ moreover have "is_pole f 0 = (fls_subdegree F < 0)"
+ using is_pole_fls_subdegree_iff F by auto
+ ultimately show ?thesis using False by auto
+ qed
+ finally show ?thesis .
+qed
+
+lemma is_pole_deriv_divide_iff:
+ assumes F:"(\<lambda>w. f (z+w)) has_laurent_expansion F" and "F\<noteq>0"
+ shows "is_pole (\<lambda>x. deriv f x / f x) z \<longleftrightarrow> is_pole f z \<or> (f \<midarrow>z\<rightarrow>0)"
+proof -
+ define ff df where "ff=(\<lambda>w. f (z+w))" and "df=(\<lambda>w. deriv f (z + w))"
+ have "is_pole (\<lambda>x. deriv f x / f x) z
+ \<longleftrightarrow> is_pole (\<lambda>x. deriv ff x / ff x) 0"
+ unfolding ff_def df_def
+ by (simp add:deriv_shift_0' is_pole_shift_0' comp_def algebra_simps)
+ moreover have "is_pole f z \<longleftrightarrow> is_pole ff 0"
+ unfolding ff_def by (auto simp:is_pole_shift_0')
+ moreover have "(f \<midarrow>z\<rightarrow>0) \<longleftrightarrow> (ff \<midarrow>0\<rightarrow>0)"
+ unfolding ff_def by (simp add: LIM_offset_zero_iff)
+ moreover have "is_pole (\<lambda>x. deriv ff x / ff x) 0 = (is_pole ff 0 \<or> ff \<midarrow>0\<rightarrow> 0)"
+ apply (rule is_pole_0_deriv_divide_iff)
+ using F ff_def \<open>F\<noteq>0\<close> by blast+
+ ultimately show ?thesis by auto
+qed
+
+lemma subdegree_imp_eventually_deriv_nzero_0:
+ assumes F:"f has_laurent_expansion F" and "fls_subdegree F\<noteq>0"
+ shows "eventually (\<lambda>z. deriv f z \<noteq> 0) (at 0)"
+proof -
+ have "deriv f has_laurent_expansion fls_deriv F"
+ using has_laurent_expansion_deriv[OF F] .
+ moreover have "fls_deriv F\<noteq>0"
+ using \<open>fls_subdegree F\<noteq>0\<close>
+ by (metis fls_const_subdegree fls_deriv_eq_0_iff)
+ ultimately show ?thesis
+ using has_laurent_expansion_eventually_nonzero_iff' by blast
+qed
+
+lemma subdegree_imp_eventually_deriv_nzero:
+ assumes F:"(\<lambda>w. f (z+w)) has_laurent_expansion F"
+ and "fls_subdegree F\<noteq>0"
+ shows "eventually (\<lambda>w. deriv f w \<noteq> 0) (at z)"
+proof -
+ have "\<forall>\<^sub>F x in at 0. deriv (\<lambda>w. f (z + w)) x \<noteq> 0"
+ using subdegree_imp_eventually_deriv_nzero_0 assms by auto
+ then show ?thesis
+ apply (subst eventually_at_to_0)
+ by (simp add:deriv_shift_0' comp_def algebra_simps)
+qed
+
+lemma has_fps_expansion_imp_asymp_equiv_0:
+ fixes f :: "complex \<Rightarrow> complex"
+ assumes F: "f has_fps_expansion F"
+ defines "n \<equiv> subdegree F"
+ shows "f \<sim>[nhds 0] (\<lambda>z. fps_nth F n * z ^ n)"
+proof -
+ have F': "f has_laurent_expansion fps_to_fls F"
+ using F has_laurent_expansion_fps by blast
+
+ have "f \<sim>[at 0] (\<lambda>z. fps_nth F n * z ^ n)"
+ using has_laurent_expansion_imp_asymp_equiv_0[OF F']
+ by (simp add: fls_subdegree_fls_to_fps n_def)
+ moreover have "f 0 = fps_nth F n * 0 ^ n"
+ using F by (auto simp: n_def has_fps_expansion_to_laurent power_0_left)
+ ultimately show ?thesis
+ by (auto simp: asymp_equiv_nhds_iff)
+qed
+
+lemma has_fps_expansion_imp_tendsto_0:
+ fixes f :: "complex \<Rightarrow> complex"
+ assumes "f has_fps_expansion F"
+ shows "(f \<longlongrightarrow> fps_nth F 0) (nhds 0)"
+proof (rule asymp_equiv_tendsto_transfer)
+ show "(\<lambda>z. fps_nth F (subdegree F) * z ^ subdegree F) \<sim>[nhds 0] f"
+ by (rule asymp_equiv_symI, rule has_fps_expansion_imp_asymp_equiv_0) fact
+ have "((\<lambda>z. F $ subdegree F * z ^ subdegree F) \<longlongrightarrow> F $ 0) (at 0)"
+ by (rule tendsto_eq_intros refl | simp)+ (auto simp: power_0_left)
+ thus "((\<lambda>z. F $ subdegree F * z ^ subdegree F) \<longlongrightarrow> F $ 0) (nhds 0)"
+ by (auto simp: tendsto_nhds_iff power_0_left)
+qed
+
+lemma has_fps_expansion_imp_0_eq_fps_nth_0:
+ assumes "f has_fps_expansion F"
+ shows "f 0 = fps_nth F 0"
+proof -
+ have "eventually (\<lambda>x. f x = eval_fps F x) (nhds 0)"
+ using assms by (auto simp: has_fps_expansion_def eq_commute)
+ then obtain A where "open A" "0 \<in> A" "\<forall>x\<in>A. f x = eval_fps F x"
+ unfolding eventually_nhds by blast
+ hence "f 0 = eval_fps F 0"
+ by blast
+ thus ?thesis
+ by (simp add: eval_fps_at_0)
+qed
+
+lemma fls_nth_compose_aux:
+ assumes "f has_fps_expansion F"
+ assumes G: "g has_fps_expansion G" "fps_nth G 0 = 0" "fps_deriv G \<noteq> 0"
+ assumes "(f \<circ> g) has_laurent_expansion H"
+ shows "fls_nth H (int n) = fps_nth (fps_compose F G) n"
+ using assms(1,5)
+proof (induction n arbitrary: f F H rule: less_induct)
+ case (less n f F H)
+ have [simp]: "g 0 = 0"
+ using has_fps_expansion_imp_0_eq_fps_nth_0[OF G(1)] G(2) by simp
+ have ana_f: "f analytic_on {0}"
+ using less.prems by (meson has_fps_expansion_imp_analytic_0)
+ have ana_g: "g analytic_on {0}"
+ using G by (meson has_fps_expansion_imp_analytic_0)
+ have "(f \<circ> g) has_laurent_expansion fps_to_fls (fps_expansion (f \<circ> g) 0)"
+ by (rule analytic_at_imp_has_fps_expansion_0 analytic_intros has_laurent_expansion_fps
+ analytic_on_compose_gen ana_f ana_g)+ auto
+ with less.prems have "H = fps_to_fls (fps_expansion (f \<circ> g) 0)"
+ using has_laurent_expansion_unique by blast
+ also have "fls_subdegree \<dots> \<ge> 0"
+ by (simp add: fls_subdegree_fls_to_fps)
+ finally have subdeg: "fls_subdegree H \<ge> 0" .
+
+ show ?case
+ proof (cases "n = 0")
+ case [simp]: True
+ have lim_g: "g \<midarrow>0\<rightarrow> 0"
+ using has_laurent_expansion_imp_tendsto_0[of g "fps_to_fls G"] G
+ by (auto simp: fls_subdegree_fls_to_fps_gt0 has_fps_expansion_to_laurent)
+ have lim_f: "(f \<longlongrightarrow> fps_nth F 0) (nhds 0)"
+ by (intro has_fps_expansion_imp_tendsto_0 less.prems)
+ have "(\<lambda>x. f (g x)) \<midarrow>0\<rightarrow> fps_nth F 0"
+ by (rule filterlim_compose[OF lim_f lim_g])
+ moreover have "(f \<circ> g) \<midarrow>0\<rightarrow> fls_nth H 0"
+ by (intro has_laurent_expansion_imp_tendsto_0 less.prems subdeg)
+ ultimately have "fps_nth F 0 = fls_nth H 0"
+ using tendsto_unique by (force simp: o_def)
+ thus ?thesis
+ by simp
+ next
+ case n: False
+ define GH where "GH = (fls_deriv H / fls_deriv (fps_to_fls G))"
+ define GH' where "GH' = fls_regpart GH"
+
+ have "(\<lambda>x. deriv (f \<circ> g) x / deriv g x) has_laurent_expansion
+ fls_deriv H / fls_deriv (fps_to_fls G)"
+ by (intro laurent_expansion_intros less.prems has_laurent_expansion_fps[of _ G] G)
+ also have "?this \<longleftrightarrow> (deriv f \<circ> g) has_laurent_expansion fls_deriv H / fls_deriv (fps_to_fls G)"
+ proof (rule has_laurent_expansion_cong)
+ from ana_f obtain r1 where r1: "r1 > 0" "f holomorphic_on ball 0 r1"
+ unfolding analytic_on_def by blast
+ from ana_g obtain r2 where r2: "r2 > 0" "g holomorphic_on ball 0 r2"
+ unfolding analytic_on_def by blast
+ have lim_g: "g \<midarrow>0\<rightarrow> 0"
+ using has_laurent_expansion_imp_tendsto_0[of g "fps_to_fls G"] G
+ by (auto simp: fls_subdegree_fls_to_fps_gt0 has_fps_expansion_to_laurent)
+ moreover have "open (ball 0 r1)" "0 \<in> ball 0 r1"
+ using r1 by auto
+ ultimately have "eventually (\<lambda>x. g x \<in> ball 0 r1) (at 0)"
+ unfolding tendsto_def by blast
+ moreover have "eventually (\<lambda>x. deriv g x \<noteq> 0) (at 0)"
+ using G fps_to_fls_eq_0_iff has_fps_expansion_deriv has_fps_expansion_to_laurent
+ has_laurent_expansion_nonzero_imp_eventually_nonzero by blast
+ moreover have "eventually (\<lambda>x. x \<in> ball 0 (min r1 r2) - {0}) (at 0)"
+ by (intro eventually_at_in_open) (use r1 r2 in auto)
+ ultimately show "eventually (\<lambda>x. deriv (f \<circ> g) x / deriv g x = (deriv f \<circ> g) x) (at 0)"
+ proof eventually_elim
+ case (elim x)
+ thus ?case using r1 r2
+ by (subst deriv_chain)
+ (auto simp: field_simps holomorphic_on_def at_within_open[of _ "ball _ _"])
+ qed
+ qed auto
+ finally have GH: "(deriv f \<circ> g) has_laurent_expansion GH"
+ unfolding GH_def .
+
+ have "(deriv f \<circ> g) has_laurent_expansion fps_to_fls (fps_expansion (deriv f \<circ> g) 0)"
+ by (rule analytic_at_imp_has_fps_expansion_0 analytic_intros has_laurent_expansion_fps
+ analytic_on_compose_gen ana_f ana_g)+ auto
+ with GH have "GH = fps_to_fls (fps_expansion (deriv f \<circ> g) 0)"
+ using has_laurent_expansion_unique by blast
+ also have "fls_subdegree \<dots> \<ge> 0"
+ by (simp add: fls_subdegree_fls_to_fps)
+ finally have subdeg': "fls_subdegree GH \<ge> 0" .
+
+ have "deriv f has_fps_expansion fps_deriv F"
+ by (intro fps_expansion_intros less.prems)
+ from this and GH have IH: "fls_nth GH (int k) = fps_nth (fps_compose (fps_deriv F) G) k"
+ if "k < n" for k
+ by (intro less.IH that)
+
+ have "fps_nth (fps_compose (fps_deriv F) G) n = (\<Sum>i=0..n. of_nat (Suc i) * F $ Suc i * G ^ i $ n)"
+ by (simp add: fps_compose_nth)
+
+ have "fps_nth (fps_compose F G) n =
+ fps_nth (fps_deriv (fps_compose F G)) (n - 1) / of_nat n"
+ using n by (cases n) (auto simp del: of_nat_Suc)
+ also have "fps_deriv (fps_compose F G) = fps_compose (fps_deriv F) G * fps_deriv G "
+ using G by (subst fps_compose_deriv) auto
+ also have "fps_nth \<dots> (n - 1) = (\<Sum>i=0..n-1. (fps_deriv F oo G) $ i * fps_deriv G $ (n - 1 - i))"
+ unfolding fps_mult_nth ..
+ also have "\<dots> = (\<Sum>i=0..n-1. fps_nth GH' i * of_nat (n - i) * G $ (n - i))"
+ using n by (intro sum.cong) (auto simp: IH Suc_diff_Suc GH'_def)
+ also have "\<dots> = (\<Sum>i=0..n. fps_nth GH' i * of_nat (n - i) * G $ (n - i))"
+ by (intro sum.mono_neutral_left) auto
+ also have "\<dots> = fps_nth (GH' * Abs_fps (\<lambda>i. of_nat i * fps_nth G i)) n"
+ by (simp add: fps_mult_nth mult_ac)
+ also have "Abs_fps (\<lambda>i. of_nat i * fps_nth G i) = fps_X * fps_deriv G"
+ by (simp add: fps_mult_fps_X_deriv_shift)
+ also have "fps_nth (GH' * (fps_X * fps_deriv G)) n =
+ fls_nth (fps_to_fls (GH' * (fps_X * fps_deriv G))) (int n)"
+ by simp
+ also have "fps_to_fls (GH' * (fps_X * fps_deriv G)) =
+ GH * fps_to_fls (fps_deriv G) * fls_X"
+ using subdeg' by (simp add: mult_ac fls_times_fps_to_fls GH'_def)
+ also have "GH * fps_to_fls (fps_deriv G) = fls_deriv H"
+ unfolding GH_def using G by (simp add: fls_deriv_fps_to_fls)
+ also have "fls_deriv H * fls_X = fls_shift (-1) (fls_deriv H)"
+ using fls_X_times_conv_shift(2) by blast
+ finally show ?thesis
+ using n by simp
+ qed
+qed
+
+lemma has_fps_expansion_compose [fps_expansion_intros]:
+ fixes f g :: "complex \<Rightarrow> complex"
+ assumes F: "f has_fps_expansion F"
+ assumes G: "g has_fps_expansion G" "fps_nth G 0 = 0"
+ shows "(f \<circ> g) has_fps_expansion fps_compose F G"
+proof (cases "fps_deriv G = 0")
+ case False
+ have [simp]: "g 0 = 0"
+ using has_fps_expansion_imp_0_eq_fps_nth_0[OF G(1)] G(2) False by simp
+ have ana_f: "f analytic_on {0}"
+ using F by (meson has_fps_expansion_imp_analytic_0)
+ have ana_g: "g analytic_on {0}"
+ using G by (meson has_fps_expansion_imp_analytic_0)
+ have fg: "(f \<circ> g) has_fps_expansion fps_expansion (f \<circ> g) 0"
+ by (rule analytic_at_imp_has_fps_expansion_0 analytic_intros
+ analytic_on_compose_gen ana_f ana_g)+ auto
+
+ have "fls_nth (fps_to_fls (fps_expansion (f \<circ> g) 0)) (int n) = fps_nth (fps_compose F G) n" for n
+ by (rule fls_nth_compose_aux has_laurent_expansion_fps F G False fg)+
+ hence "fps_expansion (f \<circ> g) 0 = fps_compose F G"
+ by (simp add: fps_eq_iff)
+ thus ?thesis using fg
+ by simp
+next
+ case True
+ have [simp]: "f 0 = fps_nth F 0"
+ using F by (auto dest: has_fps_expansion_imp_0_eq_fps_nth_0)
+ from True have "fps_nth G n = 0" for n
+ using G(2) by (cases n) (auto simp del: of_nat_Suc)
+ hence [simp]: "G = 0"
+ by (auto simp: fps_eq_iff)
+ have "(\<lambda>_. f 0) has_fps_expansion fps_const (f 0)"
+ by (intro fps_expansion_intros)
+ also have "eventually (\<lambda>x. g x = 0) (nhds 0)"
+ using G by (auto simp: has_fps_expansion_def)
+ hence "(\<lambda>_. f 0) has_fps_expansion fps_const (f 0) \<longleftrightarrow> (f \<circ> g) has_fps_expansion fps_const (f 0)"
+ by (intro has_fps_expansion_cong) (auto elim!: eventually_mono)
+ thus ?thesis
+ by simp
+qed
+
+hide_const (open) fls_compose_fps
+
+definition fls_compose_fps :: "'a :: field fls \<Rightarrow> 'a fps \<Rightarrow> 'a fls" where
+ "fls_compose_fps F G =
+ fps_to_fls (fps_compose (fls_base_factor_to_fps F) G) * fps_to_fls G powi fls_subdegree F"
+
+lemma fps_compose_of_nat [simp]: "fps_compose (of_nat n :: 'a :: comm_ring_1 fps) H = of_nat n"
+ and fps_compose_of_int [simp]: "fps_compose (of_int i) H = of_int i"
+ unfolding fps_of_nat [symmetric] fps_of_int [symmetric] numeral_fps_const
+ by (rule fps_const_compose)+
+
+lemmas [simp] = fps_to_fls_of_nat fps_to_fls_of_int
+
+lemma fls_compose_fps_0 [simp]: "fls_compose_fps 0 H = 0"
+ and fls_compose_fps_1 [simp]: "fls_compose_fps 1 H = 1"
+ and fls_compose_fps_const [simp]: "fls_compose_fps (fls_const c) H = fls_const c"
+ and fls_compose_fps_of_nat [simp]: "fls_compose_fps (of_nat n) H = of_nat n"
+ and fls_compose_fps_of_int [simp]: "fls_compose_fps (of_int i) H = of_int i"
+ and fls_compose_fps_X [simp]: "fls_compose_fps fls_X F = fps_to_fls F"
+ by (simp_all add: fls_compose_fps_def)
+
+lemma fls_compose_fps_0_right:
+ "fls_compose_fps F 0 = (if fls_subdegree F \<ge> 0 then fls_const (fls_nth F 0) else 0)"
+ by (cases "fls_subdegree F = 0") (simp_all add: fls_compose_fps_def)
+
+lemma fls_compose_fps_shift:
+ assumes "H \<noteq> 0"
+ shows "fls_compose_fps (fls_shift n F) H = fls_compose_fps F H * fps_to_fls H powi (-n)"
+proof (cases "F = 0")
+ case False
+ thus ?thesis
+ using assms by (simp add: fls_compose_fps_def power_int_diff power_int_minus field_simps)
+qed auto
+
+lemma fls_compose_fps_to_fls [simp]:
+ assumes [simp]: "G \<noteq> 0" "fps_nth G 0 = 0"
+ shows "fls_compose_fps (fps_to_fls F) G = fps_to_fls (fps_compose F G)"
+proof (cases "F = 0")
+ case False
+ define n where "n = subdegree F"
+ define F' where "F' = fps_shift n F"
+ have [simp]: "F' \<noteq> 0" "subdegree F' = 0"
+ using False by (auto simp: F'_def n_def)
+ have F_eq: "F = F' * fps_X ^ n"
+ unfolding F'_def n_def using subdegree_decompose by blast
+ have "fls_compose_fps (fps_to_fls F) G =
+ fps_to_fls (fps_shift n (fls_regpart (fps_to_fls F' * fls_X_intpow (int n))) oo G) * fps_to_fls (G ^ n)"
+ unfolding F_eq fls_compose_fps_def
+ by (simp add: fls_times_fps_to_fls fls_X_power_conv_shift_1 power_int_add
+ fls_subdegree_fls_to_fps fps_to_fls_power fls_regpart_shift_conv_fps_shift
+ flip: fls_times_both_shifted_simp)
+ also have "fps_to_fls F' * fls_X_intpow (int n) = fps_to_fls F"
+ by (simp add: F_eq fls_times_fps_to_fls fps_to_fls_power fls_X_power_conv_shift_1)
+ also have "fps_to_fls (fps_shift n (fls_regpart (fps_to_fls F)) oo G) * fps_to_fls (G ^ n) =
+ fps_to_fls ((fps_shift n (fls_regpart (fps_to_fls F)) * fps_X ^ n) oo G)"
+ by (simp add: fls_times_fps_to_fls flip: fps_compose_power add: fps_compose_mult_distrib)
+ also have "fps_shift n (fls_regpart (fps_to_fls F)) * fps_X ^ n = F"
+ by (simp add: F_eq)
+ finally show ?thesis .
+qed (auto simp: fls_compose_fps_def)
+
+lemma fls_compose_fps_mult:
+ assumes [simp]: "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps (F * G) H = fls_compose_fps F H * fls_compose_fps G H"
+ using assms
+proof (cases "F * G = 0")
+ case False
+ hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
+ by auto
+ define n m where "n = fls_subdegree F" "m = fls_subdegree G"
+ define F' where "F' = fls_regpart (fls_shift n F)"
+ define G' where "G' = fls_regpart (fls_shift m G)"
+ have F_eq: "F = fls_shift (-n) (fps_to_fls F')" and G_eq: "G = fls_shift (-m) (fps_to_fls G')"
+ by (simp_all add: F'_def G'_def n_m_def)
+ have "fls_compose_fps (F * G) H = fls_compose_fps (fls_shift (-(n + m)) (fps_to_fls (F' * G'))) H"
+ by (simp add: fls_times_fps_to_fls F_eq G_eq fls_shifted_times_simps)
+ also have "\<dots> = fps_to_fls ((F' oo H) * (G' oo H)) * fps_to_fls H powi (m + n)"
+ by (simp add: fls_compose_fps_shift fps_compose_mult_distrib)
+ also have "\<dots> = fls_compose_fps F H * fls_compose_fps G H"
+ by (simp add: F_eq G_eq fls_compose_fps_shift fls_times_fps_to_fls power_int_add)
+ finally show ?thesis .
+qed auto
+
+lemma fls_compose_fps_power:
+ assumes [simp]: "G \<noteq> 0" "fps_nth G 0 = 0"
+ shows "fls_compose_fps (F ^ n) G = fls_compose_fps F G ^ n"
+ by (induction n) (auto simp: fls_compose_fps_mult)
+
+lemma fls_compose_fps_add:
+ assumes [simp]: "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps (F + G) H = fls_compose_fps F H + fls_compose_fps G H"
+proof (cases "F = 0 \<or> G = 0")
+ case False
+ hence [simp]: "F \<noteq> 0" "G \<noteq> 0"
+ by auto
+ define n where "n = min (fls_subdegree F) (fls_subdegree G)"
+ define F' where "F' = fls_regpart (fls_shift n F)"
+ define G' where "G' = fls_regpart (fls_shift n G)"
+ have F_eq: "F = fls_shift (-n) (fps_to_fls F')" and G_eq: "G = fls_shift (-n) (fps_to_fls G')"
+ unfolding n_def by (simp_all add: F'_def G'_def n_def)
+ have "F + G = fls_shift (-n) (fps_to_fls (F' + G'))"
+ by (simp add: F_eq G_eq)
+ also have "fls_compose_fps \<dots> H = fls_compose_fps (fps_to_fls (F' + G')) H * fps_to_fls H powi n"
+ by (subst fls_compose_fps_shift) auto
+ also have "\<dots> = fps_to_fls (fps_compose (F' + G') H) * fps_to_fls H powi n"
+ by (subst fls_compose_fps_to_fls) auto
+ also have "\<dots> = fls_compose_fps F H + fls_compose_fps G H"
+ by (simp add: F_eq G_eq fls_compose_fps_shift fps_compose_add_distrib algebra_simps)
+ finally show ?thesis .
+qed auto
+
+lemma fls_compose_fps_uminus [simp]: "fls_compose_fps (-F) H = -fls_compose_fps F H"
+ by (simp add: fls_compose_fps_def fps_compose_uminus)
+
+lemma fls_compose_fps_diff:
+ assumes [simp]: "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps (F - G) H = fls_compose_fps F H - fls_compose_fps G H"
+ using fls_compose_fps_add[of H F "-G"] by simp
+
+lemma fps_compose_eq_0_iff:
+ fixes F G :: "'a :: idom fps"
+ assumes "fps_nth G 0 = 0"
+ shows "fps_compose F G = 0 \<longleftrightarrow> F = 0 \<or> (G = 0 \<and> fps_nth F 0 = 0)"
+proof safe
+ assume *: "fps_compose F G = 0" "F \<noteq> 0"
+ have "fps_nth (fps_compose F G) 0 = fps_nth F 0"
+ by simp
+ also have "fps_compose F G = 0"
+ by (simp add: *)
+ finally show "fps_nth F 0 = 0"
+ by simp
+ show "G = 0"
+ proof (rule ccontr)
+ assume "G \<noteq> 0"
+ hence "subdegree G > 0" using assms
+ using subdegree_eq_0_iff by blast
+ define N where "N = subdegree F * subdegree G"
+ have "fps_nth (fps_compose F G) N = (\<Sum>i = 0..N. fps_nth F i * fps_nth (G ^ i) N)"
+ unfolding fps_compose_def by (simp add: N_def)
+ also have "\<dots> = (\<Sum>i\<in>{subdegree F}. fps_nth F i * fps_nth (G ^ i) N)"
+ proof (intro sum.mono_neutral_right ballI)
+ fix i assume i: "i \<in> {0..N} - {subdegree F}"
+ show "fps_nth F i * fps_nth (G ^ i) N = 0"
+ proof (cases i "subdegree F" rule: linorder_cases)
+ assume "i > subdegree F"
+ hence "fps_nth (G ^ i) N = 0"
+ using i \<open>subdegree G > 0\<close> by (intro fps_pow_nth_below_subdegree) (auto simp: N_def)
+ thus ?thesis by simp
+ qed (use i in \<open>auto simp: N_def\<close>)
+ qed (use \<open>subdegree G > 0\<close> in \<open>auto simp: N_def\<close>)
+ also have "\<dots> = fps_nth F (subdegree F) * fps_nth (G ^ subdegree F) N"
+ by simp
+ also have "\<dots> \<noteq> 0"
+ using \<open>G \<noteq> 0\<close> \<open>F \<noteq> 0\<close> by (auto simp: N_def)
+ finally show False using * by auto
+ qed
+qed auto
+
+lemma fls_compose_fps_eq_0_iff:
+ assumes "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps F H = 0 \<longleftrightarrow> F = 0"
+ using assms fls_base_factor_to_fps_nonzero[of F]
+ by (cases "F = 0") (auto simp: fls_compose_fps_def fps_compose_eq_0_iff)
+
+lemma fls_compose_fps_inverse:
+ assumes [simp]: "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps (inverse F) H = inverse (fls_compose_fps F H)"
+proof (cases "F = 0")
+ case False
+ have "fls_compose_fps (inverse F) H * fls_compose_fps F H =
+ fls_compose_fps (inverse F * F) H"
+ by (subst fls_compose_fps_mult) auto
+ also have "inverse F * F = 1"
+ using False by simp
+ finally show ?thesis
+ using False by (simp add: field_simps fls_compose_fps_eq_0_iff)
+qed auto
+
+lemma fls_compose_fps_divide:
+ assumes [simp]: "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps (F / G) H = fls_compose_fps F H / fls_compose_fps G H"
+ using fls_compose_fps_mult[of H F "inverse G"] fls_compose_fps_inverse[of H G]
+ by (simp add: field_simps)
+
+lemma fls_compose_fps_powi:
+ assumes [simp]: "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps (F powi n) H = fls_compose_fps F H powi n"
+ by (simp add: power_int_def fls_compose_fps_power fls_compose_fps_inverse)
+
+lemma fls_compose_fps_assoc:
+ assumes [simp]: "G \<noteq> 0" "fps_nth G 0 = 0" "H \<noteq> 0" "fps_nth H 0 = 0"
+ shows "fls_compose_fps (fls_compose_fps F G) H = fls_compose_fps F (fps_compose G H)"
+proof (cases "F = 0")
+ case [simp]: False
+ define n where "n = fls_subdegree F"
+ define F' where "F' = fls_regpart (fls_shift n F)"
+ have F_eq: "F = fls_shift (-n) (fps_to_fls F')"
+ by (simp add: F'_def n_def)
+ show ?thesis
+ by (simp add: F_eq fls_compose_fps_shift fls_compose_fps_mult fls_compose_fps_powi
+ fps_compose_eq_0_iff fps_compose_assoc)
+qed auto
+
+lemma subdegree_pos_iff: "subdegree F > 0 \<longleftrightarrow> F \<noteq> 0 \<and> fps_nth F 0 = 0"
+ using subdegree_eq_0_iff[of F] by auto
+
+lemma has_fps_expansion_fps_to_fls:
+ assumes "f has_laurent_expansion fps_to_fls F"
+ shows "(\<lambda>z. if z = 0 then fps_nth F 0 else f z) has_fps_expansion F"
+ (is "?f' has_fps_expansion _")
+proof -
+ have "f has_laurent_expansion fps_to_fls F \<longleftrightarrow> ?f' has_laurent_expansion fps_to_fls F"
+ by (intro has_laurent_expansion_cong) (auto simp: eventually_at_filter)
+ with assms show ?thesis
+ by (auto simp: has_fps_expansion_to_laurent)
+qed
+
+
+lemma has_laurent_expansion_compose [laurent_expansion_intros]:
+ fixes f g :: "complex \<Rightarrow> complex"
+ assumes F: "f has_laurent_expansion F"
+ assumes G: "g has_laurent_expansion fps_to_fls G" "fps_nth G 0 = 0" "G \<noteq> 0"
+ shows "(f \<circ> g) has_laurent_expansion fls_compose_fps F G"
+proof -
+ from assms have lim_g: "g \<midarrow>0\<rightarrow> 0"
+ by (subst tendsto_0_subdegree_iff_0[OF G(1)])
+ (auto simp: fls_subdegree_fls_to_fps subdegree_pos_iff)
+ have ev1: "eventually (\<lambda>z. g z \<noteq> 0) (at 0)"
+ using \<open>G \<noteq> 0\<close> G(1) fps_to_fls_eq_0_iff has_laurent_expansion_fps
+ has_laurent_expansion_nonzero_imp_eventually_nonzero by blast
+ moreover have "eventually (\<lambda>z. z \<noteq> 0) (at (0 :: complex))"
+ by (auto simp: eventually_at_filter)
+ ultimately have ev: "eventually (\<lambda>z. z \<noteq> 0 \<and> g z \<noteq> 0) (at 0)"
+ by eventually_elim blast
+ from ev1 and lim_g have lim_g': "filterlim g (at 0) (at 0)"
+ by (auto simp: filterlim_at)
+ define g' where "g' = (\<lambda>z. if z = 0 then fps_nth G 0 else g z)"
+
+ show ?thesis
+ proof (cases "F = 0")
+ assume [simp]: "F = 0"
+ have "eventually (\<lambda>z. f z = 0) (at 0)"
+ using F by (auto simp: has_laurent_expansion_def)
+ hence "eventually (\<lambda>z. f (g z) = 0) (at 0)"
+ using lim_g' by (rule eventually_compose_filterlim)
+ thus ?thesis
+ by (auto simp: has_laurent_expansion_def)
+ next
+ assume [simp]: "F \<noteq> 0"
+ define n where "n = fls_subdegree F"
+ define f' where
+ "f' = (\<lambda>z. if z = 0 then fps_nth (fls_base_factor_to_fps F) 0 else f z * z powi -n)"
+ have "((\<lambda>z. (f' \<circ> g') z * g z powi n)) has_laurent_expansion fls_compose_fps F G"
+ unfolding f'_def n_def fls_compose_fps_def g'_def
+ by (intro fps_expansion_intros laurent_expansion_intros has_fps_expansion_fps_to_fls
+ has_fps_expansion_fls_base_factor_to_fps assms has_laurent_expansion_fps)
+ also have "?this \<longleftrightarrow> ?thesis"
+ by (intro has_laurent_expansion_cong eventually_mono[OF ev])
+ (auto simp: f'_def power_int_minus g'_def)
+ finally show ?thesis .
+ qed
+qed
+
+lemma has_laurent_expansion_fls_X_inv [laurent_expansion_intros]:
+ "inverse has_laurent_expansion fls_X_inv"
+ using has_laurent_expansion_inverse[OF has_laurent_expansion_fps_X]
+ by (simp add: fls_inverse_X)
+
+lemma fls_X_power_int [simp]: "fls_X powi n = (fls_X_intpow n :: 'a :: division_ring fls)"
+ by (auto simp: power_int_def fls_X_power_conv_shift_1 fls_inverse_X fls_inverse_shift
+ simp flip: fls_inverse_X_power)
+
+lemma fls_const_power_int: "fls_const (c powi n) = fls_const (c :: 'a :: division_ring) powi n"
+ by (auto simp: power_int_def fls_const_power fls_inverse_const)
+
+lemma fls_nth_fls_compose_fps_linear:
+ fixes c :: "'a :: field"
+ assumes [simp]: "c \<noteq> 0"
+ shows "fls_nth (fls_compose_fps F (fps_const c * fps_X)) n = fls_nth F n * c powi n"
+proof -
+ {
+ assume *: "n \<ge> fls_subdegree F"
+ hence "c ^ nat (n - fls_subdegree F) = c powi int (nat (n - fls_subdegree F))"
+ by (simp add: power_int_def)
+ also have "\<dots> * c powi fls_subdegree F = c powi (int (nat (n - fls_subdegree F)) + fls_subdegree F)"
+ using * by (subst power_int_add) auto
+ also have "\<dots> = c powi n"
+ using * by simp
+ finally have "c ^ nat (n - fls_subdegree F) * c powi fls_subdegree F = c powi n" .
+ }
+ thus ?thesis
+ by (simp add: fls_compose_fps_def fps_compose_linear fls_times_fps_to_fls power_int_mult_distrib
+ fls_shifted_times_simps
+ flip: fls_const_power_int)
+qed
+
+lemma zorder_times_analytic:
+ assumes "f analytic_on {z}" "g analytic_on {z}"
+ assumes "eventually (\<lambda>z. f z * g z \<noteq> 0) (at z)"
+ shows "zorder (\<lambda>z. f z * g z) z = zorder f z + zorder g z"
+proof -
+ have *: "(\<lambda>w. f (z + w)) has_fps_expansion fps_expansion f z"
+ "(\<lambda>w. g (z + w)) has_fps_expansion fps_expansion g z"
+ "(\<lambda>w. f (z + w) * g (z + w)) has_fps_expansion fps_expansion f z * fps_expansion g z"
+ by (intro fps_expansion_intros analytic_at_imp_has_fps_expansion assms)+
+ have [simp]: "fps_expansion f z \<noteq> 0"
+ proof
+ assume "fps_expansion f z = 0"
+ hence "eventually (\<lambda>z. f z * g z = 0) (at z)" using *(1)
+ by (auto simp: has_fps_expansion_0_iff nhds_to_0' eventually_filtermap eventually_at_filter
+ elim: eventually_mono)
+ with assms(3) have "eventually (\<lambda>z. False) (at z)"
+ by eventually_elim auto
+ thus False by simp
+ qed
+ have [simp]: "fps_expansion g z \<noteq> 0"
+ proof
+ assume "fps_expansion g z = 0"
+ hence "eventually (\<lambda>z. f z * g z = 0) (at z)" using *(2)
+ by (auto simp: has_fps_expansion_0_iff nhds_to_0' eventually_filtermap eventually_at_filter
+ elim: eventually_mono)
+ with assms(3) have "eventually (\<lambda>z. False) (at z)"
+ by eventually_elim auto
+ thus False by simp
+ qed
+ from *[THEN has_fps_expansion_zorder] show ?thesis
+ by auto
+qed
+
+lemma analytic_on_prod [analytic_intros]:
+ assumes "\<And>x. x \<in> A \<Longrightarrow> f x analytic_on B"
+ shows "(\<lambda>z. \<Prod>x\<in>A. f x z) analytic_on B"
+ using assms by (induction A rule: infinite_finite_induct) (auto intro!: analytic_intros)
+
+lemma zorder_const [simp]: "c \<noteq> 0 \<Longrightarrow> zorder (\<lambda>_. c) z = 0"
+ by (intro zorder_eqI[where s = UNIV]) auto
+
+lemma zorder_prod_analytic:
+ assumes "\<And>x. x \<in> A \<Longrightarrow> f x analytic_on {z}"
+ assumes "eventually (\<lambda>z. (\<Prod>x\<in>A. f x z) \<noteq> 0) (at z)"
+ shows "zorder (\<lambda>z. \<Prod>x\<in>A. f x z) z = (\<Sum>x\<in>A. zorder (f x) z)"
+ using assms
+proof (induction A rule: infinite_finite_induct)
+ case (insert x A)
+ have "zorder (\<lambda>z. f x z * (\<Prod>x\<in>A. f x z)) z = zorder (f x) z + zorder (\<lambda>z. \<Prod>x\<in>A. f x z) z"
+ using insert.prems insert.hyps by (intro zorder_times_analytic analytic_intros) auto
+ also have "zorder (\<lambda>z. \<Prod>x\<in>A. f x z) z = (\<Sum>x\<in>A. zorder (f x) z)"
+ using insert.prems insert.hyps by (intro insert.IH) (auto elim!: eventually_mono)
+ finally show ?case using insert
+ by simp
+qed auto
+
+lemma zorder_eq_0I:
+ assumes "g analytic_on {z}" "g z \<noteq> 0"
+ shows "zorder g z = 0"
+proof -
+ from assms obtain r where r: "r > 0" "g holomorphic_on ball z r"
+ unfolding analytic_on_def by blast
+ thus ?thesis using assms
+ by (intro zorder_eqI[of "ball z r" _ g]) auto
+qed
+
+lemma zorder_pos_iff:
+ assumes "f holomorphic_on A" "open A" "z \<in> A" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
+ shows "zorder f z > 0 \<longleftrightarrow> f z = 0"
+proof -
+ have "f analytic_on {z}"
+ using assms analytic_at by blast
+ hence *: "(\<lambda>w. f (z + w)) has_fps_expansion fps_expansion f z"
+ using analytic_at_imp_has_fps_expansion by blast
+ have nz: "fps_expansion f z \<noteq> 0"
+ proof
+ assume "fps_expansion f z = 0"
+ hence "eventually (\<lambda>z. f z = 0) (nhds z)"
+ using * by (auto simp: has_fps_expansion_def nhds_to_0' eventually_filtermap add_ac)
+ hence "eventually (\<lambda>z. f z = 0) (at z)"
+ by (auto simp: eventually_at_filter elim: eventually_mono)
+ with assms show False
+ by (auto simp: frequently_def)
+ qed
+ from has_fps_expansion_zorder[OF * this] have eq: "zorder f z = int (subdegree (fps_expansion f z))"
+ by auto
+ moreover have "subdegree (fps_expansion f z) = 0 \<longleftrightarrow> fps_nth (fps_expansion f z) 0 \<noteq> 0"
+ using nz by (auto simp: subdegree_eq_0_iff)
+ moreover have "fps_nth (fps_expansion f z) 0 = f z"
+ by (auto simp: fps_expansion_def)
+ ultimately show ?thesis
+ by auto
+qed
+
+lemma zorder_pos_iff':
+ assumes "f analytic_on {z}" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
+ shows "zorder f z > 0 \<longleftrightarrow> f z = 0"
+proof -
+ from assms(1) obtain A where A: "open A" "{z} \<subseteq> A" "f holomorphic_on A"
+ unfolding analytic_on_holomorphic by auto
+ with zorder_pos_iff [OF A(3,1), of z] assms show ?thesis
+ by auto
+qed
+
+lemma zorder_ge_0:
+ assumes "f analytic_on {z}" "frequently (\<lambda>z. f z \<noteq> 0) (at z)"
+ shows "zorder f z \<ge> 0"
+proof -
+ have *: "(\<lambda>w. f (z + w)) has_laurent_expansion fps_to_fls (fps_expansion f z)"
+ using assms by (simp add: analytic_at_imp_has_fps_expansion has_laurent_expansion_fps)
+ from * assms(2) have "fps_to_fls (fps_expansion f z) \<noteq> 0"
+ by (auto simp: has_laurent_expansion_def frequently_def at_to_0' eventually_filtermap add_ac)
+ with has_laurent_expansion_zorder[OF *] show ?thesis
+ by (simp add: fls_subdegree_fls_to_fps)
+qed
+
+lemma zorder_eq_0_iff:
+ assumes "f analytic_on {z}" "frequently (\<lambda>w. f w \<noteq> 0) (at z)"
+ shows "zorder f z = 0 \<longleftrightarrow> f z \<noteq> 0"
+proof
+ assume "f z \<noteq> 0"
+ thus "zorder f z = 0"
+ using assms zorder_eq_0I by blast
+next
+ assume "zorder f z = 0"
+ thus "f z \<noteq> 0"
+ using assms zorder_pos_iff' by fastforce
+qed
+
+lemma dist_mult_left:
+ "dist (a * b) (a * c :: 'a :: real_normed_field) = norm a * dist b c"
+ unfolding dist_norm right_diff_distrib [symmetric] norm_mult by simp
+
+lemma dist_mult_right:
+ "dist (b * a) (c * a :: 'a :: real_normed_field) = norm a * dist b c"
+ using dist_mult_left[of a b c] by (simp add: mult_ac)
+
+lemma zorder_scale:
+ assumes "f analytic_on {a * z}" "eventually (\<lambda>w. f w \<noteq> 0) (at (a * z))" "a \<noteq> 0"
+ shows "zorder (\<lambda>w. f (a * w)) z = zorder f (a * z)"
+proof -
+ from assms(1) obtain r where r: "r > 0" "f holomorphic_on ball (a * z) r"
+ by (auto simp: analytic_on_def)
+ have *: "open (ball (a * z) r)" "connected (ball (a * z) r)" "a * z \<in> ball (a * z) r"
+ using r \<open>a \<noteq> 0\<close> by (auto simp: dist_norm)
+ from assms(2) have "eventually (\<lambda>w. f w \<noteq> 0 \<and> w \<in> ball (a * z) r - {a * z}) (at (a * z))"
+ using \<open>r > 0\<close> by (intro eventually_conj eventually_at_in_open) auto
+ then obtain z0 where "f z0 \<noteq> 0 \<and> z0 \<in> ball (a * z) r - {a * z}"
+ using eventually_happens[of _ "at (a * z)"] by force
+ hence **: "\<exists>w\<in>ball (a * z) r. f w \<noteq> 0"
+ by blast
+
+ define n where "n = nat (zorder f (a * z))"
+ obtain r' where r':
+ "(if f (a * z) = 0 then 0 < zorder f (a * z) else zorder f (a * z) = 0)"
+ "r' > 0" "cball (a * z) r' \<subseteq> ball (a * z) r" "zor_poly f (a * z) holomorphic_on cball (a * z) r'"
+ "\<And>w. w \<in> cball (a * z) r' \<Longrightarrow>
+ f w = zor_poly f (a * z) w * (w - a * z) ^ n \<and> zor_poly f (a * z) w \<noteq> 0"
+ unfolding n_def using zorder_exist_zero[OF r(2) * **] by blast
+
+ show ?thesis
+ proof (rule zorder_eqI)
+ show "open (ball z (r' / norm a))" "z \<in> ball z (r' / norm a)"
+ using r \<open>r' > 0\<close> \<open>a \<noteq> 0\<close> by auto
+ have "(*) a ` ball z (r' / cmod a) \<subseteq> cball (a * z) r'"
+ proof safe
+ fix w assume "w \<in> ball z (r' / cmod a)"
+ thus "a * w \<in> cball (a * z) r'"
+ using dist_mult_left[of a z w] \<open>a \<noteq> 0\<close> by (auto simp: divide_simps mult_ac)
+ qed
+ thus "(\<lambda>w. a ^ n * (zor_poly f (a * z) \<circ> (\<lambda>w. a * w)) w) holomorphic_on ball z (r' / norm a)"
+ using \<open>a \<noteq> 0\<close> by (intro holomorphic_on_compose_gen[OF _ r'(4)] holomorphic_intros) auto
+ show "a ^ n * (zor_poly f (a * z) \<circ> (\<lambda>w. a * w)) z \<noteq> 0"
+ using r' \<open>a \<noteq> 0\<close> by auto
+ show "f (a * w) = a ^ n * (zor_poly f (a * z) \<circ> (*) a) w * (w - z) powr of_int (zorder f (a * z))"
+ if "w \<in> ball z (r' / norm a)" "w \<noteq> z" for w
+ proof -
+ have "f (a * w) = zor_poly f (a * z) (a * w) * (a * (w - z)) ^ n"
+ using that r'(5)[of "a * w"] dist_mult_left[of a z w] \<open>a \<noteq> 0\<close> unfolding ring_distribs
+ by (auto simp: divide_simps mult_ac)
+ also have "\<dots> = a ^ n * zor_poly f (a * z) (a * w) * (w - z) ^ n"
+ by (subst power_mult_distrib) (auto simp: mult_ac)
+ also have "(w - z) ^ n = (w - z) powr of_nat n"
+ using \<open>w \<noteq> z\<close> by (subst powr_nat') auto
+ also have "of_nat n = of_int (zorder f (a * z))"
+ using r'(1) by (auto simp: n_def split: if_splits)
+ finally show ?thesis
+ unfolding o_def n_def .
+ qed
+ qed
+qed
+
+lemma subdegree_fps_compose [simp]:
+ fixes F G :: "'a :: idom fps"
+ assumes [simp]: "fps_nth G 0 = 0"
+ shows "subdegree (fps_compose F G) = subdegree F * subdegree G"
+proof (cases "G = 0"; cases "F = 0")
+ assume [simp]: "G \<noteq> 0" "F \<noteq> 0"
+ define m where "m = subdegree F"
+ define F' where "F' = fps_shift m F"
+ have F_eq: "F = F' * fps_X ^ m"
+ unfolding F'_def by (simp add: fps_shift_times_fps_X_power m_def)
+ have [simp]: "F' \<noteq> 0"
+ using \<open>F \<noteq> 0\<close> unfolding F_eq by auto
+ have "subdegree (fps_compose F G) = subdegree (fps_compose F' G) + m * subdegree G"
+ by (simp add: F_eq fps_compose_mult_distrib fps_compose_eq_0_iff flip: fps_compose_power)
+ also have "subdegree (fps_compose F' G) = 0"
+ by (intro subdegree_eq_0) (auto simp: F'_def m_def)
+ finally show ?thesis by (simp add: m_def)
+qed auto
+
+lemma fls_subdegree_power_int [simp]:
+ fixes F :: "'a :: field fls"
+ shows "fls_subdegree (F powi n) = n * fls_subdegree F"
+ by (auto simp: power_int_def fls_subdegree_pow)
+
+lemma subdegree_fls_compose_fps [simp]:
+ fixes G :: "'a :: field fps"
+ assumes [simp]: "fps_nth G 0 = 0"
+ shows "fls_subdegree (fls_compose_fps F G) = fls_subdegree F * subdegree G"
+proof (cases "F = 0"; cases "G = 0")
+ assume [simp]: "G \<noteq> 0" "F \<noteq> 0"
+ have nz1: "fls_base_factor_to_fps F \<noteq> 0"
+ using \<open>F \<noteq> 0\<close> fls_base_factor_to_fps_nonzero by blast
+ show ?thesis
+ unfolding fls_compose_fps_def using nz1
+ by (subst fls_subdegree_mult) (simp_all add: fps_compose_eq_0_iff fls_subdegree_fls_to_fps)
+qed (auto simp: fls_compose_fps_0_right)
+
+lemma zorder_compose_aux:
+ assumes "isolated_singularity_at f 0" "not_essential f 0"
+ assumes G: "g has_fps_expansion G" "G \<noteq> 0" "g 0 = 0"
+ assumes "eventually (\<lambda>w. f w \<noteq> 0) (at 0)"
+ shows "zorder (f \<circ> g) 0 = zorder f 0 * subdegree G"
+proof -
+ obtain F where F: "f has_laurent_expansion F"
+ using not_essential_has_laurent_expansion_0[OF assms(1,2)] by blast
+ have [simp]: "fps_nth G 0 = 0"
+ using G \<open>g 0 = 0\<close> by (simp add: has_fps_expansion_imp_0_eq_fps_nth_0)
+ note [simp] = \<open>G \<noteq> 0\<close> \<open>g 0 = 0\<close>
+ have [simp]: "F \<noteq> 0"
+ using has_laurent_expansion_eventually_nonzero_iff[of f 0 F] F assms by simp
+ have FG: "(f \<circ> g) has_laurent_expansion fls_compose_fps F G"
+ by (intro has_laurent_expansion_compose has_laurent_expansion_fps F G) auto
+
+ have "zorder (f \<circ> g) 0 = fls_subdegree (fls_compose_fps F G)"
+ using has_laurent_expansion_zorder_0 [OF FG] by (auto simp: fls_compose_fps_eq_0_iff)
+ also have "\<dots> = fls_subdegree F * int (subdegree G)"
+ by simp
+ also have "fls_subdegree F = zorder f 0"
+ using has_laurent_expansion_zorder_0 [OF F] by auto
+ finally show ?thesis .
+qed
+
+lemma zorder_compose:
+ assumes "isolated_singularity_at f (g z)" "not_essential f (g z)"
+ assumes G: "(\<lambda>x. g (z + x) - g z) has_fps_expansion G" "G \<noteq> 0"
+ assumes "eventually (\<lambda>w. f w \<noteq> 0) (at (g z))"
+ shows "zorder (f \<circ> g) z = zorder f (g z) * subdegree G"
+proof -
+ define f' where "f' = (\<lambda>w. f (g z + w))"
+ define g' where "g' = (\<lambda>w. g (z + w) - g z)"
+ have "zorder f (g z) = zorder f' 0"
+ by (simp add: f'_def zorder_shift' add_ac)
+ have "zorder (\<lambda>x. g x - g z) z = zorder g' 0"
+ by (simp add: g'_def zorder_shift' add_ac)
+ have "zorder (f \<circ> g) z = zorder (f' \<circ> g') 0"
+ by (simp add: zorder_shift' f'_def g'_def add_ac o_def)
+ also have "\<dots> = zorder f' 0 * int (subdegree G)"
+ proof (rule zorder_compose_aux)
+ show "isolated_singularity_at f' 0" unfolding f'_def
+ using assms has_laurent_expansion_isolated_0 not_essential_has_laurent_expansion by blast
+ show "not_essential f' 0" unfolding f'_def
+ using assms has_laurent_expansion_not_essential_0 not_essential_has_laurent_expansion by blast
+ qed (use assms in \<open>auto simp: f'_def g'_def at_to_0' eventually_filtermap add_ac\<close>)
+ also have "zorder f' 0 = zorder f (g z)"
+ by (simp add: f'_def zorder_shift' add_ac)
+ finally show ?thesis .
+qed
+
+lemma fps_to_fls_eq_fls_const_iff [simp]: "fps_to_fls F = fls_const c \<longleftrightarrow> F = fps_const c"
+proof
+ assume "F = fps_const c"
+ thus "fps_to_fls F = fls_const c"
+ by simp
+next
+ assume "fps_to_fls F = fls_const c"
+ thus "F = fps_const c"
+ by (metis fls_regpart_const fls_regpart_fps_trivial)
+qed
+
+lemma zorder_compose':
+ assumes "isolated_singularity_at f (g z)" "not_essential f (g z)"
+ assumes "g analytic_on {z}"
+ assumes "eventually (\<lambda>w. f w \<noteq> 0) (at (g z))"
+ assumes "eventually (\<lambda>w. g w \<noteq> g z) (at z)"
+ shows "zorder (f \<circ> g) z = zorder f (g z) * zorder (\<lambda>x. g x - g z) z"
+proof -
+ obtain G where G [fps_expansion_intros]: "(\<lambda>x. g (z + x)) has_fps_expansion G"
+ using assms analytic_at_imp_has_fps_expansion by blast
+ have G': "(\<lambda>x. g (z + x) - g z) has_fps_expansion G - fps_const (g z)"
+ by (intro fps_expansion_intros)
+ hence G'': "(\<lambda>x. g (z + x) - g z) has_laurent_expansion fps_to_fls (G - fps_const (g z))"
+ using has_laurent_expansion_fps by blast
+ have nz: "G - fps_const (g z) \<noteq> 0"
+ using has_laurent_expansion_eventually_nonzero_iff[OF G''] assms by auto
+ have "zorder (f \<circ> g) z = zorder f (g z) * subdegree (G - fps_const (g z))"
+ proof (rule zorder_compose)
+ show "(\<lambda>x. g (z + x) - g z) has_fps_expansion G - fps_const (g z)"
+ by (intro fps_expansion_intros)
+ qed (use assms nz in auto)
+ also have "int (subdegree (G - fps_const (g z))) = fls_subdegree (fps_to_fls G - fls_const (g z))"
+ by (simp flip: fls_subdegree_fls_to_fps)
+ also have "\<dots> = zorder (\<lambda>x. g x - g z) z"
+ using has_laurent_expansion_zorder [OF G''] nz by auto
+ finally show ?thesis .
+qed
+
+lemma analytic_at_cong:
+ assumes "eventually (\<lambda>x. f x = g x) (nhds x)" "x = y"
+ shows "f analytic_on {x} \<longleftrightarrow> g analytic_on {y}"
+proof -
+ have "g analytic_on {x}" if "f analytic_on {x}" "eventually (\<lambda>x. f x = g x) (nhds x)" for f g
+ proof -
+ have "(\<lambda>y. f (x + y)) has_fps_expansion fps_expansion f x"
+ by (rule analytic_at_imp_has_fps_expansion) fact
+ also have "?this \<longleftrightarrow> (\<lambda>y. g (x + y)) has_fps_expansion fps_expansion f x"
+ using that by (intro has_fps_expansion_cong refl) (auto simp: nhds_to_0' eventually_filtermap)
+ finally show ?thesis
+ by (rule has_fps_expansion_imp_analytic)
+ qed
+ from this[of f g] this[of g f] show ?thesis using assms
+ by (auto simp: eq_commute)
+qed
+
+
+lemma has_laurent_expansion_sin' [laurent_expansion_intros]:
+ "sin has_laurent_expansion fps_to_fls (fps_sin 1)"
+ using has_fps_expansion_sin' has_fps_expansion_to_laurent by blast
+
+lemma has_laurent_expansion_cos' [laurent_expansion_intros]:
+ "cos has_laurent_expansion fps_to_fls (fps_cos 1)"
+ using has_fps_expansion_cos' has_fps_expansion_to_laurent by blast
+
+lemma has_laurent_expansion_sin [laurent_expansion_intros]:
+ "(\<lambda>z. sin (c * z)) has_laurent_expansion fps_to_fls (fps_sin c)"
+ by (intro has_laurent_expansion_fps has_fps_expansion_sin)
+
+lemma has_laurent_expansion_cos [laurent_expansion_intros]:
+ "(\<lambda>z. cos (c * z)) has_laurent_expansion fps_to_fls (fps_cos c)"
+ by (intro has_laurent_expansion_fps has_fps_expansion_cos)
+
+lemma has_laurent_expansion_tan' [laurent_expansion_intros]:
+ "tan has_laurent_expansion fps_to_fls (fps_tan 1)"
+ using has_fps_expansion_tan' has_fps_expansion_to_laurent by blast
+
+lemma has_laurent_expansion_tan [laurent_expansion_intros]:
+ "(\<lambda>z. tan (c * z)) has_laurent_expansion fps_to_fls (fps_tan c)"
+ by (intro has_laurent_expansion_fps has_fps_expansion_tan)
+
+end
--- /dev/null Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Complex_Analysis/Meromorphic.thy Mon Feb 20 13:59:42 2023 +0100
@@ -0,0 +1,2333 @@
+theory Meromorphic
+ imports Laurent_Convergence Riemann_Mapping
+begin
+
+lemma analytic_at_cong:
+ assumes "eventually (\<lambda>x. f x = g x) (nhds x)" "x = y"
+ shows "f analytic_on {x} \<longleftrightarrow> g analytic_on {y}"
+proof -
+ have "g analytic_on {x}" if "f analytic_on {x}" "eventually (\<lambda>x. f x = g x) (nhds x)" for f g
+ proof -
+ have "(\<lambda>y. f (x + y)) has_fps_expansion fps_expansion f x"
+ by (rule analytic_at_imp_has_fps_expansion) fact
+ also have "?this \<longleftrightarrow> (\<lambda>y. g (x + y)) has_fps_expansion fps_expansion f x"
+ using that by (intro has_fps_expansion_cong refl) (auto simp: nhds_to_0' eventually_filtermap)
+ finally show ?thesis
+ by (rule has_fps_expansion_imp_analytic)
+ qed
+ from this[of f g] this[of g f] show ?thesis using assms
+ by (auto simp: eq_commute)
+qed
+
+definition remove_sings :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where
+ "remove_sings f z = (if \<exists>c. f \<midarrow>z\<rightarrow> c then Lim (at z) f else 0)"
+
+lemma remove_sings_eqI [intro]:
+ assumes "f \<midarrow>z\<rightarrow> c"
+ shows "remove_sings f z = c"
+ using assms unfolding remove_sings_def by (auto simp: tendsto_Lim)
+
+lemma remove_sings_at_analytic [simp]:
+ assumes "f analytic_on {z}"
+ shows "remove_sings f z = f z"
+ using assms by (intro remove_sings_eqI) (simp add: analytic_at_imp_isCont isContD)
+
+lemma remove_sings_at_pole [simp]:
+ assumes "is_pole f z"
+ shows "remove_sings f z = 0"
+ using assms unfolding remove_sings_def is_pole_def
+ by (meson at_neq_bot not_tendsto_and_filterlim_at_infinity)
+
+lemma eventually_remove_sings_eq_at:
+ assumes "isolated_singularity_at f z"
+ shows "eventually (\<lambda>w. remove_sings f w = f w) (at z)"
+proof -
+ from assms obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
+ by (auto simp: isolated_singularity_at_def)
+ hence *: "f analytic_on {w}" if "w \<in> ball z r - {z}" for w
+ using r that by (auto intro: analytic_on_subset)
+ have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
+ using r by (intro eventually_at_in_open) auto
+ thus ?thesis
+ by eventually_elim (auto simp: remove_sings_at_analytic *)
+qed
+
+lemma eventually_remove_sings_eq_nhds:
+ assumes "f analytic_on {z}"
+ shows "eventually (\<lambda>w. remove_sings f w = f w) (nhds z)"
+proof -
+ from assms obtain A where A: "open A" "z \<in> A" "f holomorphic_on A"
+ by (auto simp: analytic_at)
+ have "eventually (\<lambda>z. z \<in> A) (nhds z)"
+ by (intro eventually_nhds_in_open A)
+ thus ?thesis
+ proof eventually_elim
+ case (elim w)
+ from elim have "f analytic_on {w}"
+ using A analytic_at by blast
+ thus ?case by auto
+ qed
+qed
+
+lemma remove_sings_compose:
+ assumes "filtermap g (at z) = at z'"
+ shows "remove_sings (f \<circ> g) z = remove_sings f z'"
+proof (cases "\<exists>c. f \<midarrow>z'\<rightarrow> c")
+ case True
+ then obtain c where c: "f \<midarrow>z'\<rightarrow> c"
+ by auto
+ from c have "remove_sings f z' = c"
+ by blast
+ moreover from c have "remove_sings (f \<circ> g) z = c"
+ using c by (intro remove_sings_eqI) (auto simp: filterlim_def filtermap_compose assms)
+ ultimately show ?thesis
+ by simp
+next
+ case False
+ hence "\<not>(\<exists>c. (f \<circ> g) \<midarrow>z\<rightarrow> c)"
+ by (auto simp: filterlim_def filtermap_compose assms)
+ with False show ?thesis
+ by (auto simp: remove_sings_def)
+qed
+
+lemma remove_sings_cong:
+ assumes "eventually (\<lambda>x. f x = g x) (at z)" "z = z'"
+ shows "remove_sings f z = remove_sings g z'"
+proof (cases "\<exists>c. f \<midarrow>z\<rightarrow> c")
+ case True
+ then obtain c where c: "f \<midarrow>z\<rightarrow> c" by blast
+ hence "remove_sings f z = c"
+ by blast
+ moreover have "f \<midarrow>z\<rightarrow> c \<longleftrightarrow> g \<midarrow>z'\<rightarrow> c"
+ using assms by (intro filterlim_cong refl) auto
+ with c have "remove_sings g z' = c"
+ by (intro remove_sings_eqI) auto
+ ultimately show ?thesis
+ by simp
+next
+ case False
+ have "f \<midarrow>z\<rightarrow> c \<longleftrightarrow> g \<midarrow>z'\<rightarrow> c" for c
+ using assms by (intro filterlim_cong) auto
+ with False show ?thesis
+ by (auto simp: remove_sings_def)
+qed
+
+
+lemma deriv_remove_sings_at_analytic [simp]:
+ assumes "f analytic_on {z}"
+ shows "deriv (remove_sings f) z = deriv f z"
+ apply (rule deriv_cong_ev)
+ apply (rule eventually_remove_sings_eq_nhds)
+ using assms by auto
+
+lemma isolated_singularity_at_remove_sings [simp, intro]:
+ assumes "isolated_singularity_at f z"
+ shows "isolated_singularity_at (remove_sings f) z"
+ using isolated_singularity_at_cong[OF eventually_remove_sings_eq_at[OF assms] refl] assms
+ by simp
+
+lemma not_essential_remove_sings_iff [simp]:
+ assumes "isolated_singularity_at f z"
+ shows "not_essential (remove_sings f) z \<longleftrightarrow> not_essential f z"
+ using not_essential_cong[OF eventually_remove_sings_eq_at[OF assms(1)] refl]
+ by simp
+
+lemma not_essential_remove_sings [intro]:
+ assumes "isolated_singularity_at f z" "not_essential f z"
+ shows "not_essential (remove_sings f) z"
+ by (subst not_essential_remove_sings_iff) (use assms in auto)
+
+lemma
+ assumes "isolated_singularity_at f z"
+ shows is_pole_remove_sings_iff [simp]:
+ "is_pole (remove_sings f) z \<longleftrightarrow> is_pole f z"
+ and zorder_remove_sings [simp]:
+ "zorder (remove_sings f) z = zorder f z"
+ and zor_poly_remove_sings [simp]:
+ "zor_poly (remove_sings f) z = zor_poly f z"
+ and has_laurent_expansion_remove_sings_iff [simp]:
+ "(\<lambda>w. remove_sings f (z + w)) has_laurent_expansion F \<longleftrightarrow>
+ (\<lambda>w. f (z + w)) has_laurent_expansion F"
+ and tendsto_remove_sings_iff [simp]:
+ "remove_sings f \<midarrow>z\<rightarrow> c \<longleftrightarrow> f \<midarrow>z\<rightarrow> c"
+ by (intro is_pole_cong eventually_remove_sings_eq_at refl zorder_cong
+ zor_poly_cong has_laurent_expansion_cong' tendsto_cong assms)+
+
+lemma get_all_poles_from_remove_sings:
+ fixes f:: "complex \<Rightarrow> complex"
+ defines "ff\<equiv>remove_sings f"
+ assumes f_holo:"f holomorphic_on s - pts" and "finite pts"
+ "pts\<subseteq>s" "open s" and not_ess:"\<forall>x\<in>pts. not_essential f x"
+ obtains pts' where
+ "pts' \<subseteq> pts" "finite pts'" "ff holomorphic_on s - pts'" "\<forall>x\<in>pts'. is_pole ff x"
+proof -
+ define pts' where "pts' = {x\<in>pts. is_pole f x}"
+
+ have "pts' \<subseteq> pts" unfolding pts'_def by auto
+ then have "finite pts'" using \<open>finite pts\<close>
+ using rev_finite_subset by blast
+ then have "open (s - pts')" using \<open>open s\<close>
+ by (simp add: finite_imp_closed open_Diff)
+
+ have isolated:"isolated_singularity_at f z" if "z\<in>pts" for z
+ proof (rule isolated_singularity_at_holomorphic)
+ show "f holomorphic_on (s-(pts-{z})) - {z}"
+ by (metis Diff_insert f_holo insert_Diff that)
+ show " open (s - (pts - {z}))"
+ by (meson assms(3) assms(5) finite_Diff finite_imp_closed open_Diff)
+ show "z \<in> s - (pts - {z})"
+ using assms(4) that by auto
+ qed
+
+ have "ff holomorphic_on s - pts'"
+ proof (rule no_isolated_singularity')
+ show "(ff \<longlongrightarrow> ff z) (at z within s - pts')" if "z \<in> pts-pts'" for z
+ proof -
+ have "at z within s - pts' = at z"
+ apply (rule at_within_open)
+ using \<open>open (s - pts')\<close> that \<open>pts\<subseteq>s\<close> by auto
+ moreover have "ff \<midarrow>z\<rightarrow> ff z"
+ unfolding ff_def
+ proof (subst tendsto_remove_sings_iff)
+ show "isolated_singularity_at f z"
+ apply (rule isolated)
+ using that by auto
+ have "not_essential f z"
+ using not_ess that by auto
+ moreover have "\<not>is_pole f z"
+ using that unfolding pts'_def by auto
+ ultimately have "\<exists>c. f \<midarrow>z\<rightarrow> c"
+ unfolding not_essential_def by auto
+ then show "f \<midarrow>z\<rightarrow> remove_sings f z"
+ using remove_sings_eqI by blast
+ qed
+ ultimately show ?thesis by auto
+ qed
+ have "ff holomorphic_on s - pts"
+ using f_holo
+ proof (elim holomorphic_transform)
+ fix x assume "x \<in> s - pts"
+ then have "f analytic_on {x}"
+ using assms(3) assms(5) f_holo
+ by (meson finite_imp_closed
+ holomorphic_on_imp_analytic_at open_Diff)
+ from remove_sings_at_analytic[OF this]
+ show "f x = ff x" unfolding ff_def by auto
+ qed
+ then show "ff holomorphic_on s - pts' - (pts - pts')"
+ apply (elim holomorphic_on_subset)
+ by blast
+ show "open (s - pts')"
+ by (simp add: \<open>open (s - pts')\<close>)
+ show "finite (pts - pts')"
+ by (simp add: assms(3))
+ qed
+ moreover have "\<forall>x\<in>pts'. is_pole ff x"
+ unfolding pts'_def
+ using ff_def is_pole_remove_sings_iff isolated by blast
+ moreover note \<open>pts' \<subseteq> pts\<close> \<open>finite pts'\<close>
+ ultimately show ?thesis using that by auto
+qed
+
+lemma remove_sings_eq_0_iff:
+ assumes "not_essential f w"
+ shows "remove_sings f w = 0 \<longleftrightarrow> is_pole f w \<or> f \<midarrow>w\<rightarrow> 0"
+proof (cases "is_pole f w")
+ case True
+ then show ?thesis by simp
+next
+ case False
+ then obtain c where c:"f \<midarrow>w\<rightarrow> c"
+ using \<open>not_essential f w\<close> unfolding not_essential_def by auto
+ then show ?thesis
+ using False remove_sings_eqI by auto
+qed
+
+definition meromorphic_on:: "[complex \<Rightarrow> complex, complex set, complex set] \<Rightarrow> bool"
+ ("_ (meromorphic'_on) _ _" [50,50,50]50) where
+ "f meromorphic_on D pts \<equiv>
+ open D \<and> pts \<subseteq> D \<and> (\<forall>z\<in>pts. isolated_singularity_at f z \<and> not_essential f z) \<and>
+ (\<forall>z\<in>D. \<not>(z islimpt pts)) \<and> (f holomorphic_on D-pts)"
+
+lemma meromorphic_imp_holomorphic: "f meromorphic_on D pts \<Longrightarrow> f holomorphic_on (D - pts)"
+ unfolding meromorphic_on_def by auto
+
+lemma meromorphic_imp_closedin_pts:
+ assumes "f meromorphic_on D pts"
+ shows "closedin (top_of_set D) pts"
+ by (meson assms closedin_limpt meromorphic_on_def)
+
+lemma meromorphic_imp_open_diff':
+ assumes "f meromorphic_on D pts" "pts' \<subseteq> pts"
+ shows "open (D - pts')"
+proof -
+ have "D - pts' = D - closure pts'"
+ proof safe
+ fix x assume x: "x \<in> D" "x \<in> closure pts'" "x \<notin> pts'"
+ hence "x islimpt pts'"
+ by (subst islimpt_in_closure) auto
+ hence "x islimpt pts"
+ by (rule islimpt_subset) fact
+ with assms x show False
+ by (auto simp: meromorphic_on_def)
+ qed (use closure_subset in auto)
+ then show ?thesis
+ using assms meromorphic_on_def by auto
+qed
+
+lemma meromorphic_imp_open_diff: "f meromorphic_on D pts \<Longrightarrow> open (D - pts)"
+ by (erule meromorphic_imp_open_diff') auto
+
+lemma meromorphic_pole_subset:
+ assumes merf: "f meromorphic_on D pts"
+ shows "{x\<in>D. is_pole f x} \<subseteq> pts"
+ by (smt (verit) Diff_iff assms mem_Collect_eq meromorphic_imp_open_diff
+ meromorphic_on_def not_is_pole_holomorphic subsetI)
+
+named_theorems meromorphic_intros
+
+lemma meromorphic_on_subset:
+ assumes "f meromorphic_on A pts" "open B" "B \<subseteq> A" "pts' = pts \<inter> B"
+ shows "f meromorphic_on B pts'"
+ unfolding meromorphic_on_def
+proof (intro ballI conjI)
+ fix z assume "z \<in> B"
+ show "\<not>z islimpt pts'"
+ proof
+ assume "z islimpt pts'"
+ hence "z islimpt pts"
+ by (rule islimpt_subset) (use \<open>pts' = _\<close> in auto)
+ thus False using \<open>z \<in> B\<close> \<open>B \<subseteq> A\<close> assms(1)
+ by (auto simp: meromorphic_on_def)
+ qed
+qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+
+lemma meromorphic_on_superset_pts:
+ assumes "f meromorphic_on A pts" "pts \<subseteq> pts'" "pts' \<subseteq> A" "\<forall>x\<in>A. \<not>x islimpt pts'"
+ shows "f meromorphic_on A pts'"
+ unfolding meromorphic_on_def
+proof (intro conjI ballI impI)
+ fix z assume "z \<in> pts'"
+ from assms(1) have holo: "f holomorphic_on A - pts" and "open A"
+ unfolding meromorphic_on_def by blast+
+ have "open (A - pts)"
+ by (intro meromorphic_imp_open_diff[OF assms(1)])
+
+ show "isolated_singularity_at f z"
+ proof (cases "z \<in> pts")
+ case False
+ thus ?thesis
+ using \<open>open (A - pts)\<close> assms \<open>z \<in> pts'\<close>
+ by (intro isolated_singularity_at_holomorphic[of _ "A - pts"] holomorphic_on_subset[OF holo])
+ auto
+ qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+
+ show "not_essential f z"
+ proof (cases "z \<in> pts")
+ case False
+ thus ?thesis
+ using \<open>open (A - pts)\<close> assms \<open>z \<in> pts'\<close>
+ by (intro not_essential_holomorphic[of _ "A - pts"] holomorphic_on_subset[OF holo])
+ auto
+ qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+
+lemma meromorphic_on_no_singularities: "f meromorphic_on A {} \<longleftrightarrow> f holomorphic_on A \<and> open A"
+ by (auto simp: meromorphic_on_def)
+
+lemma holomorphic_on_imp_meromorphic_on:
+ "f holomorphic_on A \<Longrightarrow> pts \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> \<forall>x\<in>A. \<not>x islimpt pts \<Longrightarrow> f meromorphic_on A pts"
+ by (rule meromorphic_on_superset_pts[where pts = "{}"])
+ (auto simp: meromorphic_on_no_singularities)
+
+lemma meromorphic_on_const [meromorphic_intros]:
+ assumes "open A" "\<forall>x\<in>A. \<not>x islimpt pts" "pts \<subseteq> A"
+ shows "(\<lambda>_. c) meromorphic_on A pts"
+ by (rule holomorphic_on_imp_meromorphic_on) (use assms in auto)
+
+lemma meromorphic_on_ident [meromorphic_intros]:
+ assumes "open A" "\<forall>x\<in>A. \<not>x islimpt pts" "pts \<subseteq> A"
+ shows "(\<lambda>x. x) meromorphic_on A pts"
+ by (rule holomorphic_on_imp_meromorphic_on) (use assms in auto)
+
+lemma meromorphic_on_id [meromorphic_intros]:
+ assumes "open A" "\<forall>x\<in>A. \<not>x islimpt pts" "pts \<subseteq> A"
+ shows "id meromorphic_on A pts"
+ using meromorphic_on_ident assms unfolding id_def .
+
+lemma not_essential_add [singularity_intros]:
+ assumes f_ness: "not_essential f z" and g_ness: "not_essential g z"
+ assumes f_iso: "isolated_singularity_at f z" and g_iso: "isolated_singularity_at g z"
+ shows "not_essential (\<lambda>w. f w + g w) z"
+proof -
+ have "(\<lambda>w. f (z + w) + g (z + w)) has_laurent_expansion laurent_expansion f z + laurent_expansion g z"
+ by (intro not_essential_has_laurent_expansion laurent_expansion_intros assms)
+ hence "not_essential (\<lambda>w. f (z + w) + g (z + w)) 0"
+ using has_laurent_expansion_not_essential_0 by blast
+ thus ?thesis
+ by (simp add: not_essential_shift_0)
+qed
+
+lemma meromorphic_on_uminus [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>z. -f z) meromorphic_on A pts"
+ unfolding meromorphic_on_def
+ by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>)
+
+lemma meromorphic_on_add [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" "g meromorphic_on A pts"
+ shows "(\<lambda>z. f z + g z) meromorphic_on A pts"
+ unfolding meromorphic_on_def
+ by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>)
+
+lemma meromorphic_on_add':
+ assumes "f meromorphic_on A pts1" "g meromorphic_on A pts2"
+ shows "(\<lambda>z. f z + g z) meromorphic_on A (pts1 \<union> pts2)"
+proof (rule meromorphic_intros)
+ show "f meromorphic_on A (pts1 \<union> pts2)"
+ by (rule meromorphic_on_superset_pts[OF assms(1)])
+ (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>)
+ show "g meromorphic_on A (pts1 \<union> pts2)"
+ by (rule meromorphic_on_superset_pts[OF assms(2)])
+ (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>)
+qed
+
+lemma meromorphic_on_add_const [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>z. f z + c) meromorphic_on A pts"
+ unfolding meromorphic_on_def
+ by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>)
+
+lemma meromorphic_on_minus_const [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>z. f z - c) meromorphic_on A pts"
+ using meromorphic_on_add_const[OF assms,of "-c"] by simp
+
+lemma meromorphic_on_diff [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" "g meromorphic_on A pts"
+ shows "(\<lambda>z. f z - g z) meromorphic_on A pts"
+ using meromorphic_on_add[OF assms(1) meromorphic_on_uminus[OF assms(2)]] by simp
+
+lemma meromorphic_on_diff':
+ assumes "f meromorphic_on A pts1" "g meromorphic_on A pts2"
+ shows "(\<lambda>z. f z - g z) meromorphic_on A (pts1 \<union> pts2)"
+proof (rule meromorphic_intros)
+ show "f meromorphic_on A (pts1 \<union> pts2)"
+ by (rule meromorphic_on_superset_pts[OF assms(1)])
+ (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>)
+ show "g meromorphic_on A (pts1 \<union> pts2)"
+ by (rule meromorphic_on_superset_pts[OF assms(2)])
+ (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>)
+qed
+
+lemma meromorphic_on_mult [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" "g meromorphic_on A pts"
+ shows "(\<lambda>z. f z * g z) meromorphic_on A pts"
+ unfolding meromorphic_on_def
+ by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>)
+
+lemma meromorphic_on_mult':
+ assumes "f meromorphic_on A pts1" "g meromorphic_on A pts2"
+ shows "(\<lambda>z. f z * g z) meromorphic_on A (pts1 \<union> pts2)"
+proof (rule meromorphic_intros)
+ show "f meromorphic_on A (pts1 \<union> pts2)"
+ by (rule meromorphic_on_superset_pts[OF assms(1)])
+ (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>)
+ show "g meromorphic_on A (pts1 \<union> pts2)"
+ by (rule meromorphic_on_superset_pts[OF assms(2)])
+ (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>)
+qed
+
+
+
+lemma meromorphic_on_imp_not_essential:
+ assumes "f meromorphic_on A pts" "z \<in> A"
+ shows "not_essential f z"
+proof (cases "z \<in> pts")
+ case False
+ thus ?thesis
+ using not_essential_holomorphic[of f "A - pts" z] meromorphic_imp_open_diff[OF assms(1)] assms
+ by (auto simp: meromorphic_on_def)
+qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+
+lemma meromorphic_imp_analytic: "f meromorphic_on D pts \<Longrightarrow> f analytic_on (D - pts)"
+ unfolding meromorphic_on_def
+ apply (subst analytic_on_open)
+ using meromorphic_imp_open_diff meromorphic_on_id apply blast
+ apply auto
+ done
+
+lemma not_islimpt_isolated_zeros:
+ assumes mero: "f meromorphic_on A pts" and "z \<in> A"
+ shows "\<not>z islimpt {w\<in>A. isolated_zero f w}"
+proof
+ assume islimpt: "z islimpt {w\<in>A. isolated_zero f w}"
+ have holo: "f holomorphic_on A - pts" and "open A"
+ using assms by (auto simp: meromorphic_on_def)
+ have open': "open (A - (pts - {z}))"
+ by (intro meromorphic_imp_open_diff'[OF mero]) auto
+ then obtain r where r: "r > 0" "ball z r \<subseteq> A - (pts - {z})"
+ using meromorphic_imp_open_diff[OF mero] \<open>z \<in> A\<close> openE by blast
+
+ have "not_essential f z"
+ using assms by (rule meromorphic_on_imp_not_essential)
+ then consider c where "f \<midarrow>z\<rightarrow> c" | "is_pole f z"
+ unfolding not_essential_def by blast
+ thus False
+ proof cases
+ assume "is_pole f z"
+ hence "eventually (\<lambda>w. f w \<noteq> 0) (at z)"
+ by (rule non_zero_neighbour_pole)
+ hence "\<not>z islimpt {w. f w = 0}"
+ by (simp add: islimpt_conv_frequently_at frequently_def)
+ moreover have "z islimpt {w. f w = 0}"
+ using islimpt by (rule islimpt_subset) (auto simp: isolated_zero_def)
+ ultimately show False by contradiction
+ next
+ fix c assume c: "f \<midarrow>z\<rightarrow> c"
+ define g where "g = (\<lambda>w. if w = z then c else f w)"
+ have holo': "g holomorphic_on A - (pts - {z})" unfolding g_def
+ by (intro removable_singularity holomorphic_on_subset[OF holo] open' c) auto
+
+ have eq_zero: "g w = 0" if "w \<in> ball z r" for w
+ proof (rule analytic_continuation[where f = g])
+ show "open (ball z r)" "connected (ball z r)" "{w\<in>ball z r. isolated_zero f w} \<subseteq> ball z r"
+ by auto
+ have "z islimpt {w\<in>A. isolated_zero f w} \<inter> ball z r"
+ using islimpt \<open>r > 0\<close> by (intro islimpt_Int_eventually eventually_at_in_open') auto
+ also have "\<dots> = {w\<in>ball z r. isolated_zero f w}"
+ using r by auto
+ finally show "z islimpt {w\<in>ball z r. isolated_zero f w}"
+ by simp
+ next
+ fix w assume w: "w \<in> {w\<in>ball z r. isolated_zero f w}"
+ show "g w = 0"
+ proof (cases "w = z")
+ case False
+ thus ?thesis using w by (auto simp: g_def isolated_zero_def)
+ next
+ case True
+ have "z islimpt {z. f z = 0}"
+ using islimpt by (rule islimpt_subset) (auto simp: isolated_zero_def)
+ thus ?thesis
+ using w by (simp add: isolated_zero_altdef True)
+ qed
+ qed (use r that in \<open>auto intro!: holomorphic_on_subset[OF holo'] simp: isolated_zero_def\<close>)
+
+ have "infinite ({w\<in>A. isolated_zero f w} \<inter> ball z r)"
+ using islimpt \<open>r > 0\<close> unfolding islimpt_eq_infinite_ball by blast
+ hence "{w\<in>A. isolated_zero f w} \<inter> ball z r \<noteq> {}"
+ by force
+ then obtain z0 where z0: "z0 \<in> A" "isolated_zero f z0" "z0 \<in> ball z r"
+ by blast
+ have "\<forall>\<^sub>F y in at z0. y \<in> ball z r - (if z = z0 then {} else {z}) - {z0}"
+ using r z0 by (intro eventually_at_in_open) auto
+ hence "eventually (\<lambda>w. f w = 0) (at z0)"
+ proof eventually_elim
+ case (elim w)
+ show ?case
+ using eq_zero[of w] elim by (auto simp: g_def split: if_splits)
+ qed
+ hence "eventually (\<lambda>w. f w = 0) (at z0)"
+ by (auto simp: g_def eventually_at_filter elim!: eventually_mono split: if_splits)
+ moreover from z0 have "eventually (\<lambda>w. f w \<noteq> 0) (at z0)"
+ by (simp add: isolated_zero_def)
+ ultimately have "eventually (\<lambda>_. False) (at z0)"
+ by eventually_elim auto
+ thus False
+ by simp
+ qed
+qed
+
+lemma closedin_isolated_zeros:
+ assumes "f meromorphic_on A pts"
+ shows "closedin (top_of_set A) {z\<in>A. isolated_zero f z}"
+ unfolding closedin_limpt using not_islimpt_isolated_zeros[OF assms] by auto
+
+lemma meromorphic_on_deriv':
+ assumes "f meromorphic_on A pts" "open A"
+ assumes "\<And>x. x \<in> A - pts \<Longrightarrow> (f has_field_derivative f' x) (at x)"
+ shows "f' meromorphic_on A pts"
+ unfolding meromorphic_on_def
+proof (intro conjI ballI)
+ have "open (A - pts)"
+ by (intro meromorphic_imp_open_diff[OF assms(1)])
+ thus "f' holomorphic_on A - pts"
+ by (rule derivative_is_holomorphic) (use assms in auto)
+next
+ fix z assume "z \<in> pts"
+ hence "z \<in> A"
+ using assms(1) by (auto simp: meromorphic_on_def)
+ from \<open>z \<in> pts\<close> obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
+ using assms(1) by (auto simp: meromorphic_on_def isolated_singularity_at_def)
+
+ have "open (ball z r \<inter> (A - (pts - {z})))"
+ by (intro open_Int assms meromorphic_imp_open_diff'[OF assms(1)]) auto
+ then obtain r' where r': "r' > 0" "ball z r' \<subseteq> ball z r \<inter> (A - (pts - {z}))"
+ using r \<open>z \<in> A\<close> by (subst (asm) open_contains_ball) fastforce
+
+ have "open (ball z r' - {z})"
+ by auto
+ hence "f' holomorphic_on ball z r' - {z}"
+ by (rule derivative_is_holomorphic[of _ f]) (use r' in \<open>auto intro!: assms(3)\<close>)
+ moreover have "open (ball z r' - {z})"
+ by auto
+ ultimately show "isolated_singularity_at f' z"
+ unfolding isolated_singularity_at_def using \<open>r' > 0\<close>
+ by (auto simp: analytic_on_open intro!: exI[of _ r'])
+next
+ fix z assume z: "z \<in> pts"
+ hence z': "not_essential f z" "z \<in> A"
+ using assms by (auto simp: meromorphic_on_def)
+ from z'(1) show "not_essential f' z"
+ proof (rule not_essential_deriv')
+ show "z \<in> A - (pts - {z})"
+ using \<open>z \<in> A\<close> by blast
+ show "open (A - (pts - {z}))"
+ by (intro meromorphic_imp_open_diff'[OF assms(1)]) auto
+ qed (use assms in auto)
+qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+
+lemma meromorphic_on_deriv [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" "open A"
+ shows "deriv f meromorphic_on A pts"
+proof (intro meromorphic_on_deriv'[OF assms(1)])
+ have *: "open (A - pts)"
+ by (intro meromorphic_imp_open_diff[OF assms(1)])
+ show "(f has_field_derivative deriv f x) (at x)" if "x \<in> A - pts" for x
+ using assms(1) by (intro holomorphic_derivI[OF _ * that]) (auto simp: meromorphic_on_def)
+qed fact
+
+lemma meromorphic_on_imp_analytic_at:
+ assumes "f meromorphic_on A pts" "z \<in> A - pts"
+ shows "f analytic_on {z}"
+ using assms by (metis analytic_at meromorphic_imp_open_diff meromorphic_on_def)
+
+lemma meromorphic_compact_finite_pts:
+ assumes "f meromorphic_on D pts" "compact S" "S \<subseteq> D"
+ shows "finite (S \<inter> pts)"
+proof -
+ { assume "infinite (S \<inter> pts)"
+ then obtain z where "z \<in> S" and z: "z islimpt (S \<inter> pts)"
+ using assms by (metis compact_eq_Bolzano_Weierstrass inf_le1)
+ then have False
+ using assms by (meson in_mono inf_le2 islimpt_subset meromorphic_on_def) }
+ then show ?thesis by metis
+qed
+
+lemma meromorphic_imp_countable:
+ assumes "f meromorphic_on D pts"
+ shows "countable pts"
+proof -
+ obtain K :: "nat \<Rightarrow> complex set" where K: "D = (\<Union>n. K n)" "\<And>n. compact (K n)"
+ using assms unfolding meromorphic_on_def by (metis open_Union_compact_subsets)
+ then have "pts = (\<Union>n. K n \<inter> pts)"
+ using assms meromorphic_on_def by auto
+ moreover have "\<And>n. finite (K n \<inter> pts)"
+ by (metis K(1) K(2) UN_I assms image_iff meromorphic_compact_finite_pts rangeI subset_eq)
+ ultimately show ?thesis
+ by (metis countableI_type countable_UN countable_finite)
+qed
+
+lemma meromorphic_imp_connected_diff':
+ assumes "f meromorphic_on D pts" "connected D" "pts' \<subseteq> pts"
+ shows "connected (D - pts')"
+proof (rule connected_open_diff_countable)
+ show "countable pts'"
+ by (rule countable_subset [OF assms(3)]) (use assms(1) in \<open>auto simp: meromorphic_imp_countable\<close>)
+qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+
+lemma meromorphic_imp_connected_diff:
+ assumes "f meromorphic_on D pts" "connected D"
+ shows "connected (D - pts)"
+ using meromorphic_imp_connected_diff'[OF assms order.refl] .
+
+lemma meromorphic_on_compose [meromorphic_intros]:
+ assumes f: "f meromorphic_on A pts" and g: "g holomorphic_on B"
+ assumes "open B" and "g ` B \<subseteq> A"
+ shows "(\<lambda>x. f (g x)) meromorphic_on B (isolated_points_of (g -` pts \<inter> B))"
+ unfolding meromorphic_on_def
+proof (intro ballI conjI)
+ fix z assume z: "z \<in> isolated_points_of (g -` pts \<inter> B)"
+ hence z': "z \<in> B" "g z \<in> pts"
+ using isolated_points_of_subset by blast+
+ have g': "g analytic_on {z}"
+ using g z' \<open>open B\<close> analytic_at by blast
+
+ show "isolated_singularity_at (\<lambda>x. f (g x)) z"
+ by (rule isolated_singularity_at_compose[OF _ g']) (use f z' in \<open>auto simp: meromorphic_on_def\<close>)
+ show "not_essential (\<lambda>x. f (g x)) z"
+ by (rule not_essential_compose[OF _ g']) (use f z' in \<open>auto simp: meromorphic_on_def\<close>)
+next
+ fix z assume z: "z \<in> B"
+ hence "g z \<in> A"
+ using assms by auto
+ hence "\<not>g z islimpt pts"
+ using f by (auto simp: meromorphic_on_def)
+ hence ev: "eventually (\<lambda>w. w \<notin> pts) (at (g z))"
+ by (auto simp: islimpt_conv_frequently_at frequently_def)
+ have g': "g analytic_on {z}"
+ by (rule holomorphic_on_imp_analytic_at[OF g]) (use assms z in auto)
+
+ (* TODO: There's probably a useful lemma somewhere in here to extract... *)
+ have "eventually (\<lambda>w. w \<notin> isolated_points_of (g -` pts \<inter> B)) (at z)"
+ proof (cases "isolated_zero (\<lambda>w. g w - g z) z")
+ case True
+ have "eventually (\<lambda>w. w \<notin> pts) (at (g z))"
+ using ev by (auto simp: islimpt_conv_frequently_at frequently_def)
+ moreover have "g \<midarrow>z\<rightarrow> g z"
+ using analytic_at_imp_isCont[OF g'] isContD by blast
+ hence lim: "filterlim g (at (g z)) (at z)"
+ using True by (auto simp: filterlim_at isolated_zero_def)
+ have "eventually (\<lambda>w. g w \<notin> pts) (at z)"
+ using ev lim by (rule eventually_compose_filterlim)
+ thus ?thesis
+ by eventually_elim (auto simp: isolated_points_of_def)
+ next
+ case False
+ have "eventually (\<lambda>w. g w - g z = 0) (nhds z)"
+ using False by (rule non_isolated_zero) (auto intro!: analytic_intros g')
+ hence "eventually (\<lambda>w. g w = g z \<and> w \<in> B) (nhds z)"
+ using eventually_nhds_in_open[OF \<open>open B\<close> \<open>z \<in> B\<close>]
+ by eventually_elim auto
+ then obtain X where X: "open X" "z \<in> X" "X \<subseteq> B" "\<forall>x\<in>X. g x = g z"
+ unfolding eventually_nhds by blast
+
+ have "z0 \<notin> isolated_points_of (g -` pts \<inter> B)" if "z0 \<in> X" for z0
+ proof (cases "g z \<in> pts")
+ case False
+ with that have "g z0 \<notin> pts"
+ using X by metis
+ thus ?thesis
+ by (auto simp: isolated_points_of_def)
+ next
+ case True
+ have "eventually (\<lambda>w. w \<in> X) (at z0)"
+ by (intro eventually_at_in_open') fact+
+ hence "eventually (\<lambda>w. w \<in> g -` pts \<inter> B) (at z0)"
+ by eventually_elim (use X True in fastforce)
+ hence "frequently (\<lambda>w. w \<in> g -` pts \<inter> B) (at z0)"
+ by (meson at_neq_bot eventually_frequently)
+ thus "z0 \<notin> isolated_points_of (g -` pts \<inter> B)"
+ unfolding isolated_points_of_def by (auto simp: frequently_def)
+ qed
+ moreover have "eventually (\<lambda>x. x \<in> X) (at z)"
+ by (intro eventually_at_in_open') fact+
+ ultimately show ?thesis
+ by (auto elim!: eventually_mono)
+ qed
+ thus "\<not>z islimpt isolated_points_of (g -` pts \<inter> B)"
+ by (auto simp: islimpt_conv_frequently_at frequently_def)
+next
+ have "f \<circ> g analytic_on (\<Union>z\<in>B - isolated_points_of (g -` pts \<inter> B). {z})"
+ unfolding analytic_on_UN
+ proof
+ fix z assume z: "z \<in> B - isolated_points_of (g -` pts \<inter> B)"
+ hence "z \<in> B" by blast
+ have g': "g analytic_on {z}"
+ by (rule holomorphic_on_imp_analytic_at[OF g]) (use assms z in auto)
+ show "f \<circ> g analytic_on {z}"
+ proof (cases "g z \<in> pts")
+ case False
+ show ?thesis
+ proof (rule analytic_on_compose)
+ show "f analytic_on g ` {z}" using False z assms
+ by (auto intro!: meromorphic_on_imp_analytic_at[OF f])
+ qed fact
+ next
+ case True
+ show ?thesis
+ proof (cases "isolated_zero (\<lambda>w. g w - g z) z")
+ case False
+ hence "eventually (\<lambda>w. g w - g z = 0) (nhds z)"
+ by (rule non_isolated_zero) (auto intro!: analytic_intros g')
+ hence "f \<circ> g analytic_on {z} \<longleftrightarrow> (\<lambda>_. f (g z)) analytic_on {z}"
+ by (intro analytic_at_cong) (auto elim!: eventually_mono)
+ thus ?thesis
+ by simp
+ next
+ case True
+ hence ev: "eventually (\<lambda>w. g w \<noteq> g z) (at z)"
+ by (auto simp: isolated_zero_def)
+
+ have "\<not>g z islimpt pts"
+ using \<open>g z \<in> pts\<close> f by (auto simp: meromorphic_on_def)
+ hence "eventually (\<lambda>w. w \<notin> pts) (at (g z))"
+ by (auto simp: islimpt_conv_frequently_at frequently_def)
+ moreover have "g \<midarrow>z\<rightarrow> g z"
+ using analytic_at_imp_isCont[OF g'] isContD by blast
+ with ev have "filterlim g (at (g z)) (at z)"
+ by (auto simp: filterlim_at)
+ ultimately have "eventually (\<lambda>w. g w \<notin> pts) (at z)"
+ using eventually_compose_filterlim by blast
+ hence "z \<in> isolated_points_of (g -` pts \<inter> B)"
+ using \<open>g z \<in> pts\<close> \<open>z \<in> B\<close>
+ by (auto simp: isolated_points_of_def elim!: eventually_mono)
+ with z show ?thesis by simp
+ qed
+ qed
+ qed
+ also have "\<dots> = B - isolated_points_of (g -` pts \<inter> B)"
+ by blast
+ finally show "(\<lambda>x. f (g x)) holomorphic_on B - isolated_points_of (g -` pts \<inter> B)"
+ unfolding o_def using analytic_imp_holomorphic by blast
+qed (auto simp: isolated_points_of_def \<open>open B\<close>)
+
+lemma meromorphic_on_compose':
+ assumes f: "f meromorphic_on A pts" and g: "g holomorphic_on B"
+ assumes "open B" and "g ` B \<subseteq> A" and "pts' = (isolated_points_of (g -` pts \<inter> B))"
+ shows "(\<lambda>x. f (g x)) meromorphic_on B pts'"
+ using meromorphic_on_compose[OF assms(1-4)] assms(5) by simp
+
+lemma meromorphic_on_inverse': "inverse meromorphic_on UNIV 0"
+ unfolding meromorphic_on_def
+ by (auto intro!: holomorphic_intros singularity_intros not_essential_inverse
+ isolated_singularity_at_inverse simp: islimpt_finite)
+
+lemma meromorphic_on_inverse [meromorphic_intros]:
+ assumes mero: "f meromorphic_on A pts"
+ shows "(\<lambda>z. inverse (f z)) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero f z})"
+proof -
+ have "open A"
+ using mero by (auto simp: meromorphic_on_def)
+ have open': "open (A - pts)"
+ by (intro meromorphic_imp_open_diff[OF mero])
+ have holo: "f holomorphic_on A - pts"
+ using assms by (auto simp: meromorphic_on_def)
+ have ana: "f analytic_on A - pts"
+ using open' holo by (simp add: analytic_on_open)
+
+ show ?thesis
+ unfolding meromorphic_on_def
+ proof (intro conjI ballI)
+ fix z assume z: "z \<in> pts \<union> {z\<in>A. isolated_zero f z}"
+ have "isolated_singularity_at f z \<and> not_essential f z"
+ proof (cases "z \<in> pts")
+ case False
+ have "f holomorphic_on A - pts - {z}"
+ by (intro holomorphic_on_subset[OF holo]) auto
+ hence "isolated_singularity_at f z"
+ by (rule isolated_singularity_at_holomorphic)
+ (use z False in \<open>auto intro!: meromorphic_imp_open_diff[OF mero]\<close>)
+ moreover have "not_essential f z"
+ using z False
+ by (intro not_essential_holomorphic[OF holo] meromorphic_imp_open_diff[OF mero]) auto
+ ultimately show ?thesis by blast
+ qed (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+ thus "isolated_singularity_at (\<lambda>z. inverse (f z)) z" "not_essential (\<lambda>z. inverse (f z)) z"
+ by (auto intro!: isolated_singularity_at_inverse not_essential_inverse)
+ next
+ fix z assume "z \<in> A"
+ hence "\<not> z islimpt {z\<in>A. isolated_zero f z}"
+ by (rule not_islimpt_isolated_zeros[OF mero])
+ thus "\<not> z islimpt pts \<union> {z \<in> A. isolated_zero f z}" using \<open>z \<in> A\<close>
+ using mero by (auto simp: islimpt_Un meromorphic_on_def)
+ next
+ show "pts \<union> {z \<in> A. isolated_zero f z} \<subseteq> A"
+ using mero by (auto simp: meromorphic_on_def)
+ next
+ have "(\<lambda>z. inverse (f z)) analytic_on (\<Union>w\<in>A - (pts \<union> {z \<in> A. isolated_zero f z}) . {w})"
+ unfolding analytic_on_UN
+ proof (intro ballI)
+ fix w assume w: "w \<in> A - (pts \<union> {z \<in> A. isolated_zero f z})"
+ show "(\<lambda>z. inverse (f z)) analytic_on {w}"
+ proof (cases "f w = 0")
+ case False
+ thus ?thesis using w
+ by (intro analytic_intros analytic_on_subset[OF ana]) auto
+ next
+ case True
+ have "eventually (\<lambda>w. f w = 0) (nhds w)"
+ using True w by (intro non_isolated_zero analytic_on_subset[OF ana]) auto
+ hence "(\<lambda>z. inverse (f z)) analytic_on {w} \<longleftrightarrow> (\<lambda>_. 0) analytic_on {w}"
+ using w by (intro analytic_at_cong refl) auto
+ thus ?thesis
+ by simp
+ qed
+ qed
+ also have "\<dots> = A - (pts \<union> {z \<in> A. isolated_zero f z})"
+ by blast
+ finally have "(\<lambda>z. inverse (f z)) analytic_on \<dots>" .
+ moreover have "open (A - (pts \<union> {z \<in> A. isolated_zero f z}))"
+ using closedin_isolated_zeros[OF mero] open' \<open>open A\<close>
+ by (metis (no_types, lifting) Diff_Diff_Int Diff_Un closedin_closed open_Diff open_Int)
+ ultimately show "(\<lambda>z. inverse (f z)) holomorphic_on A - (pts \<union> {z \<in> A. isolated_zero f z})"
+ by (subst (asm) analytic_on_open) auto
+ qed (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un
+ intro!: holomorphic_intros singularity_intros\<close>)
+qed
+
+lemma meromorphic_on_inverse'' [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" "{z\<in>A. f z = 0} \<subseteq> pts"
+ shows "(\<lambda>z. inverse (f z)) meromorphic_on A pts"
+proof -
+ have "(\<lambda>z. inverse (f z)) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
+ by (intro meromorphic_on_inverse assms)
+ also have "(pts \<union> {z \<in> A. isolated_zero f z}) = pts"
+ using assms(2) by (auto simp: isolated_zero_def)
+ finally show ?thesis .
+qed
+
+lemma meromorphic_on_divide [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" and "g meromorphic_on A pts"
+ shows "(\<lambda>z. f z / g z) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
+proof -
+ have mero1: "(\<lambda>z. inverse (g z)) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
+ by (intro meromorphic_intros assms)
+ have sparse: "\<forall>x\<in>A. \<not> x islimpt pts \<union> {z\<in>A. isolated_zero g z}" and "pts \<subseteq> A"
+ using mero1 by (auto simp: meromorphic_on_def)
+ have mero2: "f meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
+ by (rule meromorphic_on_superset_pts[OF assms(1)]) (use sparse \<open>pts \<subseteq> A\<close> in auto)
+ have "(\<lambda>z. f z * inverse (g z)) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
+ by (intro meromorphic_on_mult mero1 mero2)
+ thus ?thesis
+ by (simp add: field_simps)
+qed
+
+lemma meromorphic_on_divide' [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" "g meromorphic_on A pts" "{z\<in>A. g z = 0} \<subseteq> pts"
+ shows "(\<lambda>z. f z / g z) meromorphic_on A pts"
+proof -
+ have "(\<lambda>z. f z * inverse (g z)) meromorphic_on A pts"
+ by (intro meromorphic_intros assms)
+ thus ?thesis
+ by (simp add: field_simps)
+qed
+
+lemma meromorphic_on_cmult_left [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>x. c * f x) meromorphic_on A pts"
+ using assms by (intro meromorphic_intros) (auto simp: meromorphic_on_def)
+
+lemma meromorphic_on_cmult_right [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>x. f x * c) meromorphic_on A pts"
+ using assms by (intro meromorphic_intros) (auto simp: meromorphic_on_def)
+
+lemma meromorphic_on_scaleR [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>x. c *\<^sub>R f x) meromorphic_on A pts"
+ using assms unfolding scaleR_conv_of_real
+ by (intro meromorphic_intros) (auto simp: meromorphic_on_def)
+
+lemma meromorphic_on_sum [meromorphic_intros]:
+ assumes "\<And>y. y \<in> I \<Longrightarrow> f y meromorphic_on A pts"
+ assumes "I \<noteq> {} \<or> open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)"
+ shows "(\<lambda>x. \<Sum>y\<in>I. f y x) meromorphic_on A pts"
+proof -
+ have *: "open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)"
+ using assms(2)
+ proof
+ assume "I \<noteq> {}"
+ then obtain x where "x \<in> I"
+ by blast
+ from assms(1)[OF this] show ?thesis
+ by (auto simp: meromorphic_on_def)
+ qed auto
+ show ?thesis
+ using assms(1)
+ by (induction I rule: infinite_finite_induct) (use * in \<open>auto intro!: meromorphic_intros\<close>)
+qed
+
+lemma meromorphic_on_prod [meromorphic_intros]:
+ assumes "\<And>y. y \<in> I \<Longrightarrow> f y meromorphic_on A pts"
+ assumes "I \<noteq> {} \<or> open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)"
+ shows "(\<lambda>x. \<Prod>y\<in>I. f y x) meromorphic_on A pts"
+proof -
+ have *: "open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)"
+ using assms(2)
+ proof
+ assume "I \<noteq> {}"
+ then obtain x where "x \<in> I"
+ by blast
+ from assms(1)[OF this] show ?thesis
+ by (auto simp: meromorphic_on_def)
+ qed auto
+ show ?thesis
+ using assms(1)
+ by (induction I rule: infinite_finite_induct) (use * in \<open>auto intro!: meromorphic_intros\<close>)
+qed
+
+lemma meromorphic_on_power [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>x. f x ^ n) meromorphic_on A pts"
+proof -
+ have "(\<lambda>x. \<Prod>i\<in>{..<n}. f x) meromorphic_on A pts"
+ by (intro meromorphic_intros assms(1)) (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+ thus ?thesis
+ by simp
+qed
+
+lemma meromorphic_on_power_int [meromorphic_intros]:
+ assumes "f meromorphic_on A pts"
+ shows "(\<lambda>z. f z powi n) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
+proof -
+ have inv: "(\<lambda>x. inverse (f x)) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
+ by (intro meromorphic_intros assms)
+ have *: "f meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
+ by (intro meromorphic_on_superset_pts [OF assms(1)])
+ (use inv in \<open>auto simp: meromorphic_on_def\<close>)
+ show ?thesis
+ proof (cases "n \<ge> 0")
+ case True
+ have "(\<lambda>x. f x ^ nat n) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
+ by (intro meromorphic_intros *)
+ thus ?thesis
+ using True by (simp add: power_int_def)
+ next
+ case False
+ have "(\<lambda>x. inverse (f x) ^ nat (-n)) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
+ by (intro meromorphic_intros assms)
+ thus ?thesis
+ using False by (simp add: power_int_def)
+ qed
+qed
+
+lemma meromorphic_on_power_int' [meromorphic_intros]:
+ assumes "f meromorphic_on A pts" "n \<ge> 0 \<or> (\<forall>z\<in>A. isolated_zero f z \<longrightarrow> z \<in> pts)"
+ shows "(\<lambda>z. f z powi n) meromorphic_on A pts"
+proof (cases "n \<ge> 0")
+ case True
+ have "(\<lambda>z. f z ^ nat n) meromorphic_on A pts"
+ by (intro meromorphic_intros assms)
+ thus ?thesis
+ using True by (simp add: power_int_def)
+next
+ case False
+ have "(\<lambda>z. f z powi n) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero f z})"
+ by (rule meromorphic_on_power_int) fact
+ also from assms(2) False have "pts \<union> {z\<in>A. isolated_zero f z} = pts"
+ by auto
+ finally show ?thesis .
+qed
+
+lemma has_laurent_expansion_on_imp_meromorphic_on:
+ assumes "open A"
+ assumes laurent: "\<And>z. z \<in> A \<Longrightarrow> \<exists>F. (\<lambda>w. f (z + w)) has_laurent_expansion F"
+ shows "f meromorphic_on A {z\<in>A. \<not>f analytic_on {z}}"
+ unfolding meromorphic_on_def
+proof (intro conjI ballI)
+ fix z assume "z \<in> {z\<in>A. \<not>f analytic_on {z}}"
+ then obtain F where F: "(\<lambda>w. f (z + w)) has_laurent_expansion F"
+ using laurent[of z] by blast
+ from F show "not_essential f z" "isolated_singularity_at f z"
+ using has_laurent_expansion_not_essential has_laurent_expansion_isolated by blast+
+next
+ fix z assume z: "z \<in> A"
+ obtain F where F: "(\<lambda>w. f (z + w)) has_laurent_expansion F"
+ using laurent[of z] \<open>z \<in> A\<close> by blast
+ from F have "isolated_singularity_at f z"
+ using has_laurent_expansion_isolated z by blast
+ then obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
+ unfolding isolated_singularity_at_def by blast
+ have "f analytic_on {w}" if "w \<in> ball z r - {z}" for w
+ by (rule analytic_on_subset[OF r(2)]) (use that in auto)
+ hence "eventually (\<lambda>w. f analytic_on {w}) (at z)"
+ using eventually_at_in_open[of "ball z r" z] \<open>r > 0\<close> by (auto elim!: eventually_mono)
+ hence "\<not>z islimpt {w. \<not>f analytic_on {w}}"
+ by (auto simp: islimpt_conv_frequently_at frequently_def)
+ thus "\<not>z islimpt {w\<in>A. \<not>f analytic_on {w}}"
+ using islimpt_subset[of z "{w\<in>A. \<not>f analytic_on {w}}" "{w. \<not>f analytic_on {w}}"] by blast
+next
+ have "f analytic_on A - {w\<in>A. \<not>f analytic_on {w}}"
+ by (subst analytic_on_analytic_at) auto
+ thus "f holomorphic_on A - {w\<in>A. \<not>f analytic_on {w}}"
+ by (meson analytic_imp_holomorphic)
+qed (use assms in auto)
+
+lemma meromorphic_on_imp_has_laurent_expansion:
+ assumes "f meromorphic_on A pts" "z \<in> A"
+ shows "(\<lambda>w. f (z + w)) has_laurent_expansion laurent_expansion f z"
+proof (cases "z \<in> pts")
+ case True
+ thus ?thesis
+ using assms by (intro not_essential_has_laurent_expansion) (auto simp: meromorphic_on_def)
+next
+ case False
+ have "f holomorphic_on (A - pts)"
+ using assms by (auto simp: meromorphic_on_def)
+ moreover have "z \<in> A - pts" "open (A - pts)"
+ using assms(2) False by (auto intro!: meromorphic_imp_open_diff[OF assms(1)])
+ ultimately have "f analytic_on {z}"
+ unfolding analytic_at by blast
+ thus ?thesis
+ using isolated_singularity_at_analytic not_essential_analytic
+ not_essential_has_laurent_expansion by blast
+qed
+
+lemma
+ assumes "isolated_singularity_at f z" "f \<midarrow>z\<rightarrow> c"
+ shows eventually_remove_sings_eq_nhds':
+ "eventually (\<lambda>w. remove_sings f w = (if w = z then c else f w)) (nhds z)"
+ and remove_sings_analytic_at_singularity: "remove_sings f analytic_on {z}"
+proof -
+ have "eventually (\<lambda>w. w \<noteq> z) (at z)"
+ by (auto simp: eventually_at_filter)
+ hence "eventually (\<lambda>w. remove_sings f w = (if w = z then c else f w)) (at z)"
+ using eventually_remove_sings_eq_at[OF assms(1)]
+ by eventually_elim auto
+ moreover have "remove_sings f z = c"
+ using assms by auto
+ ultimately show ev: "eventually (\<lambda>w. remove_sings f w = (if w = z then c else f w)) (nhds z)"
+ by (simp add: eventually_at_filter)
+
+ have "(\<lambda>w. if w = z then c else f w) analytic_on {z}"
+ by (intro removable_singularity' assms)
+ also have "?this \<longleftrightarrow> remove_sings f analytic_on {z}"
+ using ev by (intro analytic_at_cong) (auto simp: eq_commute)
+ finally show \<dots> .
+qed
+
+lemma remove_sings_meromorphic_on:
+ assumes "f meromorphic_on A pts" "\<And>z. z \<in> pts - pts' \<Longrightarrow> \<not>is_pole f z" "pts' \<subseteq> pts"
+ shows "remove_sings f meromorphic_on A pts'"
+ unfolding meromorphic_on_def
+proof safe
+ have "remove_sings f analytic_on {z}" if "z \<in> A - pts'" for z
+ proof (cases "z \<in> pts")
+ case False
+ hence *: "f analytic_on {z}"
+ using assms meromorphic_imp_open_diff[OF assms(1)] that
+ by (force simp: meromorphic_on_def analytic_at)
+ have "remove_sings f analytic_on {z} \<longleftrightarrow> f analytic_on {z}"
+ by (intro analytic_at_cong eventually_remove_sings_eq_nhds * refl)
+ thus ?thesis using * by simp
+ next
+ case True
+ have isol: "isolated_singularity_at f z"
+ using True using assms by (auto simp: meromorphic_on_def)
+ from assms(1) have "not_essential f z"
+ using True by (auto simp: meromorphic_on_def)
+ with assms(2) True that obtain c where "f \<midarrow>z\<rightarrow> c"
+ by (auto simp: not_essential_def)
+ thus "remove_sings f analytic_on {z}"
+ by (intro remove_sings_analytic_at_singularity isol)
+ qed
+ hence "remove_sings f analytic_on A - pts'"
+ by (subst analytic_on_analytic_at) auto
+ thus "remove_sings f holomorphic_on A - pts'"
+ using meromorphic_imp_open_diff'[OF assms(1,3)] by (subst (asm) analytic_on_open)
+qed (use assms islimpt_subset[OF _ assms(3)] in \<open>auto simp: meromorphic_on_def\<close>)
+
+lemma remove_sings_holomorphic_on:
+ assumes "f meromorphic_on A pts" "\<And>z. z \<in> pts \<Longrightarrow> \<not>is_pole f z"
+ shows "remove_sings f holomorphic_on A"
+ using remove_sings_meromorphic_on[OF assms(1), of "{}"] assms(2)
+ by (auto simp: meromorphic_on_no_singularities)
+
+lemma meromorphic_on_Ex_iff:
+ "(\<exists>pts. f meromorphic_on A pts) \<longleftrightarrow>
+ open A \<and> (\<forall>z\<in>A. \<exists>F. (\<lambda>w. f (z + w)) has_laurent_expansion F)"
+proof safe
+ fix pts assume *: "f meromorphic_on A pts"
+ from * show "open A"
+ by (auto simp: meromorphic_on_def)
+ show "\<exists>F. (\<lambda>w. f (z + w)) has_laurent_expansion F" if "z \<in> A" for z
+ using that *
+ by (intro exI[of _ "laurent_expansion f z"] meromorphic_on_imp_has_laurent_expansion)
+qed (blast intro!: has_laurent_expansion_on_imp_meromorphic_on)
+
+lemma is_pole_inverse_holomorphic_pts:
+ fixes pts::"complex set" and f::"complex \<Rightarrow> complex"
+ defines "g \<equiv> \<lambda>x. (if x\<in>pts then 0 else inverse (f x))"
+ assumes mer: "f meromorphic_on D pts"
+ and non_z: "\<And>z. z \<in> D - pts \<Longrightarrow> f z \<noteq> 0"
+ and all_poles:"\<forall>x. is_pole f x \<longleftrightarrow> x\<in>pts"
+ shows "g holomorphic_on D"
+proof -
+ have "open D" and f_holo: "f holomorphic_on (D-pts)"
+ using mer by (auto simp: meromorphic_on_def)
+ have "\<exists>r. r>0 \<and> f analytic_on ball z r - {z}
+ \<and> (\<forall>x \<in> ball z r - {z}. f x\<noteq>0)" if "z\<in>pts" for z
+ proof -
+ have "isolated_singularity_at f z" "is_pole f z"
+ using mer meromorphic_on_def that all_poles by blast+
+ then obtain r1 where "r1>0" and fan: "f analytic_on ball z r1 - {z}"
+ by (meson isolated_singularity_at_def)
+ obtain r2 where "r2>0" "\<forall>x \<in> ball z r2 - {z}. f x\<noteq>0"
+ using non_zero_neighbour_pole[OF \<open>is_pole f z\<close>]
+ unfolding eventually_at by (metis Diff_iff UNIV_I dist_commute insertI1 mem_ball)
+ define r where "r = min r1 r2"
+ have "r>0" by (simp add: \<open>0 < r2\<close> \<open>r1>0\<close> r_def)
+ moreover have "f analytic_on ball z r - {z}"
+ using r_def by (force intro: analytic_on_subset [OF fan])
+ moreover have "\<forall>x \<in> ball z r - {z}. f x\<noteq>0"
+ by (simp add: \<open>\<forall>x\<in>ball z r2 - {z}. f x \<noteq> 0\<close> r_def)
+ ultimately show ?thesis by auto
+ qed
+ then obtain get_r where r_pos:"get_r z>0"
+ and r_ana:"f analytic_on ball z (get_r z) - {z}"
+ and r_nz:"\<forall>x \<in> ball z (get_r z) - {z}. f x\<noteq>0"
+ if "z\<in>pts" for z
+ by metis
+ define p_balls where "p_balls \<equiv> \<Union>z\<in>pts. ball z (get_r z)"
+ have g_ball:"g holomorphic_on ball z (get_r z)" if "z\<in>pts" for z
+ proof -
+ have "(\<lambda>x. if x = z then 0 else inverse (f x)) holomorphic_on ball z (get_r z)"
+ proof (rule is_pole_inverse_holomorphic)
+ show "f holomorphic_on ball z (get_r z) - {z}"
+ using analytic_imp_holomorphic r_ana that by blast
+ show "is_pole f z"
+ using mer meromorphic_on_def that all_poles by force
+ show "\<forall>x\<in>ball z (get_r z) - {z}. f x \<noteq> 0"
+ using r_nz that by metis
+ qed auto
+ then show ?thesis unfolding g_def
+ by (smt (verit, ccfv_SIG) Diff_iff Elementary_Metric_Spaces.open_ball
+ all_poles analytic_imp_holomorphic empty_iff
+ holomorphic_transform insert_iff not_is_pole_holomorphic
+ open_delete r_ana that)
+ qed
+ then have "g holomorphic_on p_balls"
+ proof -
+ have "g analytic_on p_balls"
+ unfolding p_balls_def analytic_on_UN
+ using g_ball by (simp add: analytic_on_open)
+ moreover have "open p_balls" using p_balls_def by blast
+ ultimately show ?thesis
+ by (simp add: analytic_imp_holomorphic)
+ qed
+ moreover have "g holomorphic_on D-pts"
+ proof -
+ have "(\<lambda>z. inverse (f z)) holomorphic_on D - pts"
+ using f_holo holomorphic_on_inverse non_z by blast
+ then show ?thesis
+ by (metis DiffD2 g_def holomorphic_transform)
+ qed
+ moreover have "open p_balls"
+ using p_balls_def by blast
+ ultimately have "g holomorphic_on (p_balls \<union> (D-pts))"
+ by (simp add: holomorphic_on_Un meromorphic_imp_open_diff[OF mer])
+ moreover have "D \<subseteq> p_balls \<union> (D-pts)"
+ unfolding p_balls_def using \<open>\<And>z. z \<in> pts \<Longrightarrow> 0 < get_r z\<close> by force
+ ultimately show "g holomorphic_on D" by (meson holomorphic_on_subset)
+qed
+
+lemma meromorphic_imp_analytic_on:
+ assumes "f meromorphic_on D pts"
+ shows "f analytic_on (D - pts)"
+ by (metis assms analytic_on_open meromorphic_imp_open_diff meromorphic_on_def)
+
+lemma meromorphic_imp_constant_on:
+ assumes merf: "f meromorphic_on D pts"
+ and "f constant_on (D - pts)"
+ and "\<forall>x\<in>pts. is_pole f x"
+ shows "f constant_on D"
+proof -
+ obtain c where c:"\<And>z. z \<in> D-pts \<Longrightarrow> f z = c"
+ by (meson assms constant_on_def)
+
+ have "f z = c" if "z \<in> D" for z
+ proof (cases "is_pole f z")
+ case True
+ then obtain r0 where "r0 > 0" and r0: "f analytic_on ball z r0 - {z}" and pol: "is_pole f z"
+ using merf unfolding meromorphic_on_def isolated_singularity_at_def
+ by (metis \<open>z \<in> D\<close> insert_Diff insert_Diff_if insert_iff merf
+ meromorphic_imp_open_diff not_is_pole_holomorphic)
+ have "open D"
+ using merf meromorphic_on_def by auto
+ then obtain r where "r > 0" "ball z r \<subseteq> D" "r \<le> r0"
+ by (smt (verit, best) \<open>0 < r0\<close> \<open>z \<in> D\<close> openE order_subst2 subset_ball)
+ have r: "f analytic_on ball z r - {z}"
+ by (meson Diff_mono \<open>r \<le> r0\<close> analytic_on_subset order_refl r0 subset_ball)
+ have "ball z r - {z} \<subseteq> -pts"
+ using merf r unfolding meromorphic_on_def
+ by (meson ComplI Elementary_Metric_Spaces.open_ball
+ analytic_imp_holomorphic assms(3) not_is_pole_holomorphic open_delete subsetI)
+ with \<open>ball z r \<subseteq> D\<close> have "ball z r - {z} \<subseteq> D-pts"
+ by fastforce
+ with c have c': "\<And>u. u \<in> ball z r - {z} \<Longrightarrow> f u = c"
+ by blast
+ have False if "\<forall>\<^sub>F x in at z. cmod c + 1 \<le> cmod (f x)"
+ proof -
+ have "\<forall>\<^sub>F x in at z within ball z r - {z}. cmod c + 1 \<le> cmod (f x)"
+ by (smt (verit, best) Diff_UNIV Diff_eq_empty_iff eventually_at_topological insert_subset that)
+ with \<open>r > 0\<close> show ?thesis
+ apply (simp add: c' eventually_at_filter topological_space_class.eventually_nhds open_dist)
+ by (metis dist_commute min_less_iff_conj perfect_choose_dist)
+ qed
+ with pol show ?thesis
+ by (auto simp: is_pole_def filterlim_at_infinity_conv_norm_at_top filterlim_at_top)
+ next
+ case False
+ then show ?thesis by (meson DiffI assms(3) c that)
+ qed
+ then show ?thesis
+ by (simp add: constant_on_def)
+qed
+
+
+lemma meromorphic_isolated:
+ assumes merf: "f meromorphic_on D pts" and "p\<in>pts"
+ obtains r where "r>0" "ball p r \<subseteq> D" "ball p r \<inter> pts = {p}"
+proof -
+ have "\<forall>z\<in>D. \<exists>e>0. finite (pts \<inter> ball z e)"
+ using merf unfolding meromorphic_on_def islimpt_eq_infinite_ball
+ by auto
+ then obtain r0 where r0:"r0>0" "finite (pts \<inter> ball p r0)"
+ by (metis assms(2) in_mono merf meromorphic_on_def)
+ moreover define pts' where "pts' = pts \<inter> ball p r0 - {p}"
+ ultimately have "finite pts'"
+ by simp
+
+ define r1 where "r1=(if pts'={} then r0 else
+ min (Min {dist p' p |p'. p'\<in>pts'}/2) r0)"
+ have "r1>0 \<and> pts \<inter> ball p r1 - {p} = {}"
+ proof (cases "pts'={}")
+ case True
+ then show ?thesis
+ using pts'_def r0(1) r1_def by presburger
+ next
+ case False
+ define S where "S={dist p' p |p'. p'\<in>pts'}"
+
+ have nempty:"S \<noteq> {}"
+ using False S_def by blast
+ have finite:"finite S"
+ using \<open>finite pts'\<close> S_def by simp
+
+ have "r1>0"
+ proof -
+ have "r1=min (Min S/2) r0"
+ using False unfolding S_def r1_def by auto
+ moreover have "Min S\<in>S"
+ using \<open>S\<noteq>{}\<close> \<open>finite S\<close> Min_in by auto
+ then have "Min S>0" unfolding S_def
+ using pts'_def by force
+ ultimately show ?thesis using \<open>r0>0\<close> by auto
+ qed
+ moreover have "pts \<inter> ball p r1 - {p} = {}"
+ proof (rule ccontr)
+ assume "pts \<inter> ball p r1 - {p} \<noteq> {}"
+ then obtain p' where "p'\<in>pts \<inter> ball p r1 - {p}" by blast
+ moreover have "r1\<le>r0" using r1_def by auto
+ ultimately have "p'\<in>pts'" unfolding pts'_def
+ by auto
+ then have "dist p' p\<ge>Min S"
+ using S_def eq_Min_iff local.finite by blast
+ moreover have "dist p' p < Min S"
+ using \<open>p'\<in>pts \<inter> ball p r1 - {p}\<close> False unfolding r1_def
+ apply (fold S_def)
+ by (smt (verit, ccfv_threshold) DiffD1 Int_iff dist_commute
+ dist_triangle_half_l mem_ball)
+ ultimately show False by auto
+ qed
+ ultimately show ?thesis by auto
+ qed
+ then have "r1>0" and r1_pts:"pts \<inter> ball p r1 - {p} = {}" by auto
+
+ obtain r2 where "r2>0" "ball p r2 \<subseteq> D"
+ by (metis assms(2) merf meromorphic_on_def openE subset_eq)
+ define r where "r=min r1 r2"
+ have "r > 0" unfolding r_def
+ by (simp add: \<open>0 < r1\<close> \<open>0 < r2\<close>)
+ moreover have "ball p r \<subseteq> D"
+ using \<open>ball p r2 \<subseteq> D\<close> r_def by auto
+ moreover have "ball p r \<inter> pts = {p}"
+ using assms(2) \<open>r>0\<close> r1_pts
+ unfolding r_def by auto
+ ultimately show ?thesis using that by auto
+qed
+
+lemma meromorphic_pts_closure:
+ assumes merf: "f meromorphic_on D pts"
+ shows "pts \<subseteq> closure (D - pts)"
+proof -
+ have "p islimpt (D - pts)" if "p\<in>pts" for p
+ proof -
+ obtain r where "r>0" "ball p r \<subseteq> D" "ball p r \<inter> pts = {p}"
+ using meromorphic_isolated[OF merf \<open>p\<in>pts\<close>] by auto
+ from \<open>r>0\<close>
+ have "p islimpt ball p r - {p}"
+ by (meson open_ball ball_subset_cball in_mono islimpt_ball
+ islimpt_punctured le_less open_contains_ball_eq)
+ moreover have " ball p r - {p} \<subseteq> D - pts"
+ using \<open>ball p r \<inter> pts = {p}\<close> \<open>ball p r \<subseteq> D\<close> by fastforce
+ ultimately show ?thesis
+ using islimpt_subset by auto
+ qed
+ then show ?thesis by (simp add: islimpt_in_closure subset_eq)
+qed
+
+lemma nconst_imp_nzero_neighbour:
+ assumes merf: "f meromorphic_on D pts"
+ and f_nconst:"\<not>(\<forall>w\<in>D-pts. f w=0)"
+ and "z\<in>D" and "connected D"
+ shows "(\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w \<in> D - pts)"
+proof -
+ obtain \<beta> where \<beta>:"\<beta> \<in> D - pts" "f \<beta>\<noteq>0"
+ using f_nconst by auto
+
+ have ?thesis if "z\<notin>pts"
+ proof -
+ have "\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w \<in> D - pts"
+ apply (rule non_zero_neighbour_alt[of f "D-pts" z \<beta>])
+ subgoal using merf meromorphic_on_def by blast
+ subgoal using merf meromorphic_imp_open_diff by auto
+ subgoal using assms(4) merf meromorphic_imp_connected_diff by blast
+ subgoal by (simp add: assms(3) that)
+ using \<beta> by auto
+ then show ?thesis by (auto elim:eventually_mono)
+ qed
+ moreover have ?thesis if "z\<in>pts" "\<not> f \<midarrow>z\<rightarrow> 0"
+ proof -
+ have "\<forall>\<^sub>F w in at z. w \<in> D - pts"
+ using merf[unfolded meromorphic_on_def islimpt_iff_eventually] \<open>z\<in>D\<close>
+ using eventually_at_in_open' eventually_elim2 by fastforce
+ moreover have "\<forall>\<^sub>F w in at z. f w \<noteq> 0"
+ proof (cases "is_pole f z")
+ case True
+ then show ?thesis using non_zero_neighbour_pole by auto
+ next
+ case False
+ moreover have "not_essential f z"
+ using merf meromorphic_on_def that(1) by fastforce
+ ultimately obtain c where "c\<noteq>0" "f \<midarrow>z\<rightarrow> c"
+ by (metis \<open>\<not> f \<midarrow>z\<rightarrow> 0\<close> not_essential_def)
+ then show ?thesis
+ using tendsto_imp_eventually_ne by auto
+ qed
+ ultimately show ?thesis by eventually_elim auto
+ qed
+ moreover have ?thesis if "z\<in>pts" "f \<midarrow>z\<rightarrow> 0"
+ proof -
+ define ff where "ff=(\<lambda>x. if x=z then 0 else f x)"
+ define A where "A=D - (pts - {z})"
+
+ have "f holomorphic_on A - {z}"
+ by (metis A_def Diff_insert analytic_imp_holomorphic
+ insert_Diff merf meromorphic_imp_analytic_on that(1))
+ moreover have "open A"
+ using A_def merf meromorphic_imp_open_diff' by force
+ ultimately have "ff holomorphic_on A"
+ using \<open>f \<midarrow>z\<rightarrow> 0\<close> unfolding ff_def
+ by (rule removable_singularity)
+ moreover have "connected A"
+ proof -
+ have "connected (D - pts)"
+ using assms(4) merf meromorphic_imp_connected_diff by auto
+ moreover have "D - pts \<subseteq> A"
+ unfolding A_def by auto
+ moreover have "A \<subseteq> closure (D - pts)" unfolding A_def
+ by (smt (verit, ccfv_SIG) Diff_empty Diff_insert
+ closure_subset insert_Diff_single insert_absorb
+ insert_subset merf meromorphic_pts_closure that(1))
+ ultimately show ?thesis using connected_intermediate_closure
+ by auto
+ qed
+ moreover have "z \<in> A" using A_def assms(3) by blast
+ moreover have "ff z = 0" unfolding ff_def by auto
+ moreover have "\<beta> \<in> A " using A_def \<beta>(1) by blast
+ moreover have "ff \<beta> \<noteq> 0" using \<beta>(1) \<beta>(2) ff_def that(1) by auto
+ ultimately obtain r where "0 < r"
+ "ball z r \<subseteq> A" "\<And>x. x \<in> ball z r - {z} \<Longrightarrow> ff x \<noteq> 0"
+ using \<open>open A\<close> isolated_zeros[of ff A z \<beta>] by auto
+ then show ?thesis unfolding eventually_at ff_def
+ by (intro exI[of _ r]) (auto simp: A_def dist_commute ball_def)
+ qed
+ ultimately show ?thesis by auto
+qed
+
+lemma nconst_imp_nzero_neighbour':
+ assumes merf: "f meromorphic_on D pts"
+ and f_nconst:"\<not>(\<forall>w\<in>D-pts. f w=0)"
+ and "z\<in>D" and "connected D"
+ shows "\<forall>\<^sub>F w in at z. f w \<noteq> 0"
+ using nconst_imp_nzero_neighbour[OF assms]
+ by (auto elim:eventually_mono)
+
+lemma meromorphic_compact_finite_zeros:
+ assumes merf:"f meromorphic_on D pts"
+ and "compact S" "S \<subseteq> D" "connected D"
+ and f_nconst:"\<not>(\<forall>w\<in>D-pts. f w=0)"
+ shows "finite ({x\<in>S. f x=0})"
+proof -
+ have "finite ({x\<in>S. f x=0 \<and> x \<notin> pts})"
+ proof (rule ccontr)
+ assume "infinite {x \<in> S. f x = 0 \<and> x \<notin> pts}"
+ then obtain z where "z\<in>S" and z_lim:"z islimpt {x \<in> S. f x = 0
+ \<and> x \<notin> pts}"
+ using \<open>compact S\<close> unfolding compact_eq_Bolzano_Weierstrass
+ by auto
+
+ from z_lim
+ have "\<exists>\<^sub>F x in at z. f x = 0 \<and> x \<in> S \<and> x \<notin> pts"
+ unfolding islimpt_iff_eventually not_eventually by simp
+ moreover have "\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w \<in> D - pts"
+ using nconst_imp_nzero_neighbour[OF merf f_nconst _ \<open>connected D\<close>]
+ \<open>z\<in>S\<close> \<open>S \<subseteq> D\<close>
+ by auto
+ ultimately have "\<exists>\<^sub>F x in at z. False"
+ by (simp add: eventually_mono frequently_def)
+ then show False by auto
+ qed
+ moreover have "finite (S \<inter> pts)"
+ using meromorphic_compact_finite_pts[OF merf \<open>compact S\<close> \<open>S \<subseteq> D\<close>] .
+ ultimately have "finite ({x\<in>S. f x=0 \<and> x \<notin> pts} \<union> (S \<inter> pts))"
+ unfolding finite_Un by auto
+ then show ?thesis by (elim rev_finite_subset) auto
+qed
+
+lemma meromorphic_onI [intro?]:
+ assumes "open A" "pts \<subseteq> A"
+ assumes "f holomorphic_on A - pts" "\<And>z. z \<in> A \<Longrightarrow> \<not>z islimpt pts"
+ assumes "\<And>z. z \<in> pts \<Longrightarrow> isolated_singularity_at f z"
+ assumes "\<And>z. z \<in> pts \<Longrightarrow> not_essential f z"
+ shows "f meromorphic_on A pts"
+ using assms unfolding meromorphic_on_def by blast
+
+lemma Polygamma_plus_of_nat:
+ assumes "\<forall>k<m. z \<noteq> -of_nat k"
+ shows "Polygamma n (z + of_nat m) =
+ Polygamma n z + (-1) ^ n * fact n * (\<Sum>k<m. 1 / (z + of_nat k) ^ Suc n)"
+ using assms
+proof (induction m)
+ case (Suc m)
+ have "Polygamma n (z + of_nat (Suc m)) = Polygamma n (z + of_nat m + 1)"
+ by (simp add: add_ac)
+ also have "\<dots> = Polygamma n (z + of_nat m) + (-1) ^ n * fact n * (1 / ((z + of_nat m) ^ Suc n))"
+ using Suc.prems by (subst Polygamma_plus1) (auto simp: add_eq_0_iff2)
+ also have "Polygamma n (z + of_nat m) =
+ Polygamma n z + (-1) ^ n * (\<Sum>k<m. 1 / (z + of_nat k) ^ Suc n) * fact n"
+ using Suc.prems by (subst Suc.IH) auto
+ finally show ?case
+ by (simp add: algebra_simps)
+qed auto
+
+lemma tendsto_Gamma [tendsto_intros]:
+ assumes "(f \<longlongrightarrow> c) F" "c \<notin> \<int>\<^sub>\<le>\<^sub>0"
+ shows "((\<lambda>z. Gamma (f z)) \<longlongrightarrow> Gamma c) F"
+ by (intro isCont_tendsto_compose[OF _ assms(1)] continuous_intros assms)
+
+lemma tendsto_Polygamma [tendsto_intros]:
+ fixes f :: "_ \<Rightarrow> 'a :: {real_normed_field,euclidean_space}"
+ assumes "(f \<longlongrightarrow> c) F" "c \<notin> \<int>\<^sub>\<le>\<^sub>0"
+ shows "((\<lambda>z. Polygamma n (f z)) \<longlongrightarrow> Polygamma n c) F"
+ by (intro isCont_tendsto_compose[OF _ assms(1)] continuous_intros assms)
+
+lemma analytic_on_Gamma' [analytic_intros]:
+ assumes "f analytic_on A" "\<forall>x\<in>A. f x \<notin> \<int>\<^sub>\<le>\<^sub>0"
+ shows "(\<lambda>z. Gamma (f z)) analytic_on A"
+ using analytic_on_compose_gen[OF assms(1) analytic_Gamma[of "f ` A"]] assms(2)
+ by (auto simp: o_def)
+
+lemma analytic_on_Polygamma' [analytic_intros]:
+ assumes "f analytic_on A" "\<forall>x\<in>A. f x \<notin> \<int>\<^sub>\<le>\<^sub>0"
+ shows "(\<lambda>z. Polygamma n (f z)) analytic_on A"
+ using analytic_on_compose_gen[OF assms(1) analytic_on_Polygamma[of "f ` A" n]] assms(2)
+ by (auto simp: o_def)
+
+lemma
+ shows is_pole_Polygamma: "is_pole (Polygamma n) (-of_nat m :: complex)"
+ and zorder_Polygamma: "zorder (Polygamma n) (-of_nat m) = -int (Suc n)"
+ and residue_Polygamma: "residue (Polygamma n) (-of_nat m) = (if n = 0 then -1 else 0)"
+proof -
+ define g1 :: "complex \<Rightarrow> complex" where
+ "g1 = (\<lambda>z. Polygamma n (z + of_nat (Suc m)) +
+ (-1) ^ Suc n * fact n * (\<Sum>k<m. 1 / (z + of_nat k) ^ Suc n))"
+ define g :: "complex \<Rightarrow> complex" where
+ "g = (\<lambda>z. g1 z + (-1) ^ Suc n * fact n / (z + of_nat m) ^ Suc n)"
+ define F where "F = fps_to_fls (fps_expansion g1 (-of_nat m)) + fls_const ((-1) ^ Suc n * fact n) / fls_X ^ Suc n"
+ have F_altdef: "F = fps_to_fls (fps_expansion g1 (-of_nat m)) + fls_shift (n+1) (fls_const ((-1) ^ Suc n * fact n))"
+ by (simp add: F_def del: power_Suc)
+
+ have "\<not>(-of_nat m) islimpt (\<int>\<^sub>\<le>\<^sub>0 :: complex set)"
+ by (intro discrete_imp_not_islimpt[where e = 1])
+ (auto elim!: nonpos_Ints_cases simp: dist_of_int)
+ hence "eventually (\<lambda>z::complex. z \<notin> \<int>\<^sub>\<le>\<^sub>0) (at (-of_nat m))"
+ by (auto simp: islimpt_conv_frequently_at frequently_def)
+ hence ev: "eventually (\<lambda>z. Polygamma n z = g z) (at (-of_nat m))"
+ proof eventually_elim
+ case (elim z)
+ hence *: "\<forall>k<Suc m. z \<noteq> - of_nat k"
+ by auto
+ thus ?case
+ using Polygamma_plus_of_nat[of "Suc m" z n, OF *]
+ by (auto simp: g_def g1_def algebra_simps)
+ qed
+
+ have "(\<lambda>w. g (-of_nat m + w)) has_laurent_expansion F"
+ unfolding g_def F_def
+ by (intro laurent_expansion_intros has_laurent_expansion_fps analytic_at_imp_has_fps_expansion)
+ (auto simp: g1_def intro!: laurent_expansion_intros analytic_intros)
+ also have "?this \<longleftrightarrow> (\<lambda>w. Polygamma n (-of_nat m + w)) has_laurent_expansion F"
+ using ev by (intro has_laurent_expansion_cong refl)
+ (simp_all add: eq_commute at_to_0' eventually_filtermap)
+ finally have *: "(\<lambda>w. Polygamma n (-of_nat m + w)) has_laurent_expansion F" .
+
+ have subdegree: "fls_subdegree F = -int (Suc n)" unfolding F_def
+ by (subst fls_subdegree_add_eq2) (simp_all add: fls_subdegree_fls_to_fps fls_divide_subdegree)
+ have [simp]: "F \<noteq> 0"
+ using subdegree by auto
+
+ show "is_pole (Polygamma n) (-of_nat m :: complex)"
+ using * by (rule has_laurent_expansion_imp_is_pole) (auto simp: subdegree)
+ show "zorder (Polygamma n) (-of_nat m :: complex) = -int (Suc n)"
+ by (subst has_laurent_expansion_zorder[OF *]) (auto simp: subdegree)
+ show "residue (Polygamma n) (-of_nat m :: complex) = (if n = 0 then -1 else 0)"
+ by (subst has_laurent_expansion_residue[OF *]) (auto simp: F_altdef)
+qed
+
+lemma Gamma_meromorphic_on [meromorphic_intros]: "Gamma meromorphic_on UNIV \<int>\<^sub>\<le>\<^sub>0"
+proof
+ show "\<not>z islimpt \<int>\<^sub>\<le>\<^sub>0" for z :: complex
+ by (intro discrete_imp_not_islimpt[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int)
+next
+ fix z :: complex assume z: "z \<in> \<int>\<^sub>\<le>\<^sub>0"
+ then obtain n where n: "z = -of_nat n"
+ by (elim nonpos_Ints_cases')
+ show "not_essential Gamma z"
+ by (auto simp: n intro!: is_pole_imp_not_essential is_pole_Gamma)
+ have *: "open (-(\<int>\<^sub>\<le>\<^sub>0 - {z}))"
+ by (intro open_Compl discrete_imp_closed[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int)
+ have "Gamma holomorphic_on -(\<int>\<^sub>\<le>\<^sub>0 - {z}) - {z}"
+ by (intro holomorphic_intros) auto
+ thus "isolated_singularity_at Gamma z"
+ by (rule isolated_singularity_at_holomorphic) (use z * in auto)
+qed (auto intro!: holomorphic_intros)
+
+lemma Polygamma_meromorphic_on [meromorphic_intros]: "Polygamma n meromorphic_on UNIV \<int>\<^sub>\<le>\<^sub>0"
+proof
+ show "\<not>z islimpt \<int>\<^sub>\<le>\<^sub>0" for z :: complex
+ by (intro discrete_imp_not_islimpt[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int)
+next
+ fix z :: complex assume z: "z \<in> \<int>\<^sub>\<le>\<^sub>0"
+ then obtain m where n: "z = -of_nat m"
+ by (elim nonpos_Ints_cases')
+ show "not_essential (Polygamma n) z"
+ by (auto simp: n intro!: is_pole_imp_not_essential is_pole_Polygamma)
+ have *: "open (-(\<int>\<^sub>\<le>\<^sub>0 - {z}))"
+ by (intro open_Compl discrete_imp_closed[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int)
+ have "Polygamma n holomorphic_on -(\<int>\<^sub>\<le>\<^sub>0 - {z}) - {z}"
+ by (intro holomorphic_intros) auto
+ thus "isolated_singularity_at (Polygamma n) z"
+ by (rule isolated_singularity_at_holomorphic) (use z * in auto)
+qed (auto intro!: holomorphic_intros)
+
+
+theorem argument_principle':
+ fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set"
+ \<comment> \<open>\<^term>\<open>pz\<close> is the set of non-essential singularities and zeros\<close>
+ defines "pz \<equiv> {w\<in>s. f w = 0 \<or> w \<in> poles}"
+ assumes "open s" and
+ "connected s" and
+ f_holo:"f holomorphic_on s-poles" and
+ h_holo:"h holomorphic_on s" and
+ "valid_path g" and
+ loop:"pathfinish g = pathstart g" and
+ path_img:"path_image g \<subseteq> s - pz" and
+ homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and
+ finite:"finite pz" and
+ poles:"\<forall>p\<in>s\<inter>poles. not_essential f p"
+ shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> *
+ (\<Sum>p\<in>pz. winding_number g p * h p * zorder f p)"
+proof -
+ define ff where "ff = remove_sings f"
+
+ have finite':"finite (s \<inter> poles)"
+ using finite unfolding pz_def by (auto elim:rev_finite_subset)
+
+ have isolated:"isolated_singularity_at f z" if "z\<in>s" for z
+ proof (rule isolated_singularity_at_holomorphic)
+ show "f holomorphic_on (s-(poles-{z})) - {z}"
+ by (metis Diff_empty Diff_insert Diff_insert0 Diff_subset
+ f_holo holomorphic_on_subset insert_Diff)
+ show "open (s - (poles - {z}))"
+ by (metis Diff_Diff_Int Int_Diff assms(2) finite' finite_Diff
+ finite_imp_closed inf.idem open_Diff)
+ show "z \<in> s - (poles - {z})"
+ using assms(4) that by auto
+ qed
+
+ have not_ess:"not_essential f w" if "w\<in>s" for w
+ by (metis Diff_Diff_Int Diff_iff Int_Diff Int_absorb assms(2)
+ f_holo finite' finite_imp_closed not_essential_holomorphic
+ open_Diff poles that)
+
+ have nzero:"\<forall>\<^sub>F x in at w. f x \<noteq> 0" if "w\<in>s" for w
+ proof (rule ccontr)
+ assume "\<not> (\<forall>\<^sub>F x in at w. f x \<noteq> 0)"
+ then have "\<exists>\<^sub>F x in at w. f x = 0"
+ unfolding not_eventually by simp
+ moreover have "\<forall>\<^sub>F x in at w. x\<in>s"
+ by (simp add: assms(2) eventually_at_in_open' that)
+ ultimately have "\<exists>\<^sub>F x in at w. x\<in>{w\<in>s. f w = 0}"
+ apply (elim frequently_rev_mp)
+ by (auto elim:eventually_mono)
+ from frequently_at_imp_islimpt[OF this]
+ have "w islimpt {w \<in> s. f w = 0}" .
+ then have "infinite({w \<in> s. f w = 0} \<inter> ball w 1)"
+ unfolding islimpt_eq_infinite_ball by auto
+ then have "infinite({w \<in> s. f w = 0})"
+ by auto
+ then have "infinite pz" unfolding pz_def
+ by (smt (verit) Collect_mono_iff rev_finite_subset)
+ then show False using finite by auto
+ qed
+
+ obtain pts' where pts':"pts' \<subseteq> s \<inter> poles"
+ "finite pts'" "ff holomorphic_on s - pts'" "\<forall>x\<in>pts'. is_pole ff x"
+ apply (elim get_all_poles_from_remove_sings
+ [of f,folded ff_def,rotated -1])
+ subgoal using f_holo by fastforce
+ using \<open>open s\<close> poles finite' by auto
+
+ have pts'_sub_pz:"{w \<in> s. ff w = 0 \<or> w \<in> pts'} \<subseteq> pz"
+ proof -
+ have "w\<in>poles" if "w\<in>s" "w\<in>pts'" for w
+ by (meson in_mono le_infE pts'(1) that(2))
+ moreover have "f w=0" if" w\<in>s" "w\<notin>poles" "ff w=0" for w
+ proof -
+ have "\<not> is_pole f w"
+ by (metis DiffI Diff_Diff_Int Diff_subset assms(2) f_holo
+ finite' finite_imp_closed inf.absorb_iff2
+ not_is_pole_holomorphic open_Diff that(1) that(2))
+ then have "f \<midarrow>w\<rightarrow> 0"
+ using remove_sings_eq_0_iff[OF not_ess[OF \<open>w\<in>s\<close>]] \<open>ff w=0\<close>
+ unfolding ff_def by auto
+ moreover have "f analytic_on {w}"
+ using that(1,2) finite' f_holo assms(2)
+ by (metis Diff_Diff_Int Diff_empty Diff_iff Diff_subset
+ double_diff finite_imp_closed
+ holomorphic_on_imp_analytic_at open_Diff)
+ ultimately show ?thesis
+ using ff_def remove_sings_at_analytic that(3) by presburger
+ qed
+ ultimately show ?thesis unfolding pz_def by auto
+ qed
+
+
+ have "contour_integral g (\<lambda>x. deriv f x * h x / f x)
+ = contour_integral g (\<lambda>x. deriv ff x * h x / ff x)"
+ proof (rule contour_integral_eq)
+ fix x assume "x \<in> path_image g"
+ have "f analytic_on {x}"
+ proof (rule holomorphic_on_imp_analytic_at[of _ "s-poles"])
+ from finite'
+ show "open (s - poles)"
+ using \<open>open s\<close>
+ by (metis Diff_Compl Diff_Diff_Int Diff_eq finite_imp_closed
+ open_Diff)
+ show "x \<in> s - poles"
+ using path_img \<open>x \<in> path_image g\<close> unfolding pz_def by auto
+ qed (use f_holo in simp)
+ then show "deriv f x * h x / f x = deriv ff x * h x / ff x"
+ unfolding ff_def by auto
+ qed
+ also have "... = complex_of_real (2 * pi) * \<i> *
+ (\<Sum>p\<in>{w \<in> s. ff w = 0 \<or> w \<in> pts'}.
+ winding_number g p * h p * of_int (zorder ff p))"
+ proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close>, of ff pts' h g])
+ show "path_image g \<subseteq> s - {w \<in> s. ff w = 0 \<or> w \<in> pts'}"
+ using path_img pts'_sub_pz by auto
+ show "finite {w \<in> s. ff w = 0 \<or> w \<in> pts'}"
+ using pts'_sub_pz finite
+ using rev_finite_subset by blast
+ qed (use pts' assms in auto)
+ also have "... = 2 * pi * \<i> *
+ (\<Sum>p\<in>pz. winding_number g p * h p * zorder f p)"
+ proof -
+ have "(\<Sum>p\<in>{w \<in> s. ff w = 0 \<or> w \<in> pts'}.
+ winding_number g p * h p * of_int (zorder ff p)) =
+ (\<Sum>p\<in>pz. winding_number g p * h p * of_int (zorder f p))"
+ proof (rule sum.mono_neutral_cong_left)
+ have "zorder f w = 0"
+ if "w\<in>s" " f w = 0 \<or> w \<in> poles" "ff w \<noteq> 0" " w \<notin> pts'"
+ for w
+ proof -
+ define F where "F=laurent_expansion f w"
+ have has_l:"(\<lambda>x. f (w + x)) has_laurent_expansion F"
+ unfolding F_def
+ apply (rule not_essential_has_laurent_expansion)
+ using isolated not_ess \<open>w\<in>s\<close> by auto
+ from has_laurent_expansion_eventually_nonzero_iff[OF this]
+ have "F \<noteq>0"
+ using nzero \<open>w\<in>s\<close> by auto
+ from tendsto_0_subdegree_iff[OF has_l this]
+ have "f \<midarrow>w\<rightarrow> 0 = (0 < fls_subdegree F)" .
+ moreover have "\<not> (is_pole f w \<or> f \<midarrow>w\<rightarrow> 0)"
+ using remove_sings_eq_0_iff[OF not_ess[OF \<open>w\<in>s\<close>]] \<open>ff w \<noteq> 0\<close>
+ unfolding ff_def by auto
+ moreover have "is_pole f w = (fls_subdegree F < 0)"
+ using is_pole_fls_subdegree_iff[OF has_l] .
+ ultimately have "fls_subdegree F = 0" by auto
+ then show ?thesis
+ using has_laurent_expansion_zorder[OF has_l \<open>F\<noteq>0\<close>] by auto
+ qed
+ then show "\<forall>i\<in>pz - {w \<in> s. ff w = 0 \<or> w \<in> pts'}.
+ winding_number g i * h i * of_int (zorder f i) = 0"
+ unfolding pz_def by auto
+ show "\<And>x. x \<in> {w \<in> s. ff w = 0 \<or> w \<in> pts'} \<Longrightarrow>
+ winding_number g x * h x * of_int (zorder ff x) =
+ winding_number g x * h x * of_int (zorder f x)"
+ using isolated zorder_remove_sings[of f,folded ff_def] by auto
+ qed (use pts'_sub_pz finite in auto)
+ then show ?thesis by auto
+ qed
+ finally show ?thesis .
+qed
+
+lemma meromorphic_on_imp_isolated_singularity:
+ assumes "f meromorphic_on D pts" "z \<in> D"
+ shows "isolated_singularity_at f z"
+ by (meson DiffI assms(1) assms(2) holomorphic_on_imp_analytic_at isolated_singularity_at_analytic
+ meromorphic_imp_open_diff meromorphic_on_def)
+
+lemma meromorphic_imp_not_is_pole:
+ assumes "f meromorphic_on D pts" "z \<in> D - pts"
+ shows "\<not>is_pole f z"
+proof -
+ from assms have "f analytic_on {z}"
+ using meromorphic_on_imp_analytic_at by blast
+ thus ?thesis
+ using analytic_at not_is_pole_holomorphic by blast
+qed
+
+lemma meromorphic_all_poles_iff_empty [simp]: "f meromorphic_on pts pts \<longleftrightarrow> pts = {}"
+ by (auto simp: meromorphic_on_def holomorphic_on_def open_imp_islimpt)
+
+lemma meromorphic_imp_nonsingular_point_exists:
+ assumes "f meromorphic_on A pts" "A \<noteq> {}"
+ obtains x where "x \<in> A - pts"
+proof -
+ have "A \<noteq> pts"
+ using assms by auto
+ moreover have "pts \<subseteq> A"
+ using assms by (auto simp: meromorphic_on_def)
+ ultimately show ?thesis
+ using that by blast
+qed
+
+lemma meromorphic_frequently_const_imp_const:
+ assumes "f meromorphic_on A pts" "connected A"
+ assumes "frequently (\<lambda>w. f w = c) (at z)"
+ assumes "z \<in> A - pts"
+ assumes "w \<in> A - pts"
+ shows "f w = c"
+proof -
+ have "f w - c = 0"
+ proof (rule analytic_continuation[where f = "\<lambda>z. f z - c"])
+ show "(\<lambda>z. f z - c) holomorphic_on (A - pts)"
+ by (intro holomorphic_intros meromorphic_imp_holomorphic[OF assms(1)])
+ show [intro]: "open (A - pts)"
+ using assms meromorphic_imp_open_diff by blast
+ show "connected (A - pts)"
+ using assms meromorphic_imp_connected_diff by blast
+ show "{z\<in>A-pts. f z = c} \<subseteq> A - pts"
+ by blast
+ have "eventually (\<lambda>z. z \<in> A - pts) (at z)"
+ using assms by (intro eventually_at_in_open') auto
+ hence "frequently (\<lambda>z. f z = c \<and> z \<in> A - pts) (at z)"
+ by (intro frequently_eventually_frequently assms)
+ thus "z islimpt {z\<in>A-pts. f z = c}"
+ by (simp add: islimpt_conv_frequently_at conj_commute)
+ qed (use assms in auto)
+ thus ?thesis
+ by simp
+qed
+
+lemma meromorphic_imp_eventually_neq:
+ assumes "f meromorphic_on A pts" "connected A" "\<not>f constant_on A - pts"
+ assumes "z \<in> A - pts"
+ shows "eventually (\<lambda>z. f z \<noteq> c) (at z)"
+proof (rule ccontr)
+ assume "\<not>eventually (\<lambda>z. f z \<noteq> c) (at z)"
+ hence *: "frequently (\<lambda>z. f z = c) (at z)"
+ by (auto simp: frequently_def)
+ have "\<forall>w\<in>A-pts. f w = c"
+ using meromorphic_frequently_const_imp_const [OF assms(1,2) * assms(4)] by blast
+ hence "f constant_on A - pts"
+ by (auto simp: constant_on_def)
+ thus False
+ using assms(3) by contradiction
+qed
+
+lemma meromorphic_frequently_const_imp_const':
+ assumes "f meromorphic_on A pts" "connected A" "\<forall>w\<in>pts. is_pole f w"
+ assumes "frequently (\<lambda>w. f w = c) (at z)"
+ assumes "z \<in> A"
+ assumes "w \<in> A"
+ shows "f w = c"
+proof -
+ have "\<not>is_pole f z"
+ using frequently_const_imp_not_is_pole[OF assms(4)] .
+ with assms have z: "z \<in> A - pts"
+ by auto
+ have *: "f w = c" if "w \<in> A - pts" for w
+ using that meromorphic_frequently_const_imp_const [OF assms(1,2,4) z] by auto
+ have "\<not>is_pole f u" if "u \<in> A" for u
+ proof -
+ have "is_pole f u \<longleftrightarrow> is_pole (\<lambda>_. c) u"
+ proof (rule is_pole_cong)
+ have "eventually (\<lambda>w. w \<in> A - (pts - {u}) - {u}) (at u)"
+ by (intro eventually_at_in_open meromorphic_imp_open_diff' [OF assms(1)]) (use that in auto)
+ thus "eventually (\<lambda>w. f w = c) (at u)"
+ by eventually_elim (use * in auto)
+ qed auto
+ thus ?thesis
+ by auto
+ qed
+ moreover have "pts \<subseteq> A"
+ using assms(1) by (simp add: meromorphic_on_def)
+ ultimately have "pts = {}"
+ using assms(3) by auto
+ with * and \<open>w \<in> A\<close> show ?thesis
+ by blast
+qed
+
+lemma meromorphic_imp_eventually_neq':
+ assumes "f meromorphic_on A pts" "connected A" "\<forall>w\<in>pts. is_pole f w" "\<not>f constant_on A"
+ assumes "z \<in> A"
+ shows "eventually (\<lambda>z. f z \<noteq> c) (at z)"
+proof (rule ccontr)
+ assume "\<not>eventually (\<lambda>z. f z \<noteq> c) (at z)"
+ hence *: "frequently (\<lambda>z. f z = c) (at z)"
+ by (auto simp: frequently_def)
+ have "\<forall>w\<in>A. f w = c"
+ using meromorphic_frequently_const_imp_const' [OF assms(1,2,3) * assms(5)] by blast
+ hence "f constant_on A"
+ by (auto simp: constant_on_def)
+ thus False
+ using assms(4) by contradiction
+qed
+
+lemma zorder_eq_0_iff_meromorphic:
+ assumes "f meromorphic_on A pts" "\<forall>z\<in>pts. is_pole f z" "z \<in> A"
+ assumes "eventually (\<lambda>x. f x \<noteq> 0) (at z)"
+ shows "zorder f z = 0 \<longleftrightarrow> \<not>is_pole f z \<and> f z \<noteq> 0"
+proof (cases "z \<in> pts")
+ case True
+ from assms obtain F where F: "(\<lambda>x. f (z + x)) has_laurent_expansion F"
+ by (metis True meromorphic_on_def not_essential_has_laurent_expansion) (* TODO: better lemmas *)
+ from F and assms(4) have [simp]: "F \<noteq> 0"
+ using has_laurent_expansion_eventually_nonzero_iff by blast
+ show ?thesis using True assms(2)
+ using is_pole_fls_subdegree_iff [OF F] has_laurent_expansion_zorder [OF F]
+ by auto
+next
+ case False
+ have ana: "f analytic_on {z}"
+ using meromorphic_on_imp_analytic_at False assms by blast
+ hence "\<not>is_pole f z"
+ using analytic_at not_is_pole_holomorphic by blast
+ moreover have "frequently (\<lambda>w. f w \<noteq> 0) (at z)"
+ using assms(4) by (intro eventually_frequently) auto
+ ultimately show ?thesis using zorder_eq_0_iff[OF ana] False
+ by auto
+qed
+
+lemma zorder_pos_iff_meromorphic:
+ assumes "f meromorphic_on A pts" "\<forall>z\<in>pts. is_pole f z" "z \<in> A"
+ assumes "eventually (\<lambda>x. f x \<noteq> 0) (at z)"
+ shows "zorder f z > 0 \<longleftrightarrow> \<not>is_pole f z \<and> f z = 0"
+proof (cases "z \<in> pts")
+ case True
+ from assms obtain F where F: "(\<lambda>x. f (z + x)) has_laurent_expansion F"
+ by (metis True meromorphic_on_def not_essential_has_laurent_expansion) (* TODO: better lemmas *)
+ from F and assms(4) have [simp]: "F \<noteq> 0"
+ using has_laurent_expansion_eventually_nonzero_iff by blast
+ show ?thesis using True assms(2)
+ using is_pole_fls_subdegree_iff [OF F] has_laurent_expansion_zorder [OF F]
+ by auto
+next
+ case False
+ have ana: "f analytic_on {z}"
+ using meromorphic_on_imp_analytic_at False assms by blast
+ hence "\<not>is_pole f z"
+ using analytic_at not_is_pole_holomorphic by blast
+ moreover have "frequently (\<lambda>w. f w \<noteq> 0) (at z)"
+ using assms(4) by (intro eventually_frequently) auto
+ ultimately show ?thesis using zorder_pos_iff'[OF ana] False
+ by auto
+qed
+
+lemma zorder_neg_iff_meromorphic:
+ assumes "f meromorphic_on A pts" "\<forall>z\<in>pts. is_pole f z" "z \<in> A"
+ assumes "eventually (\<lambda>x. f x \<noteq> 0) (at z)"
+ shows "zorder f z < 0 \<longleftrightarrow> is_pole f z"
+proof -
+ have "frequently (\<lambda>x. f x \<noteq> 0) (at z)"
+ using assms by (intro eventually_frequently) auto
+ moreover from assms have "isolated_singularity_at f z" "not_essential f z"
+ using meromorphic_on_imp_isolated_singularity meromorphic_on_imp_not_essential by blast+
+ ultimately show ?thesis
+ using isolated_pole_imp_neg_zorder neg_zorder_imp_is_pole by blast
+qed
+
+lemma meromorphic_on_imp_discrete:
+ assumes mero:"f meromorphic_on S pts" and "connected S"
+ and nconst:"\<not> (\<forall>w\<in>S - pts. f w = c)"
+ shows "discrete {x\<in>S. f x=c}"
+proof -
+ define g where "g=(\<lambda>x. f x - c)"
+ have "\<forall>\<^sub>F w in at z. g w \<noteq> 0" if "z \<in> S" for z
+ proof (rule nconst_imp_nzero_neighbour'[of g S pts z])
+ show "g meromorphic_on S pts" using mero unfolding g_def
+ by (auto intro:meromorphic_intros)
+ show "\<not> (\<forall>w\<in>S - pts. g w = 0)" using nconst unfolding g_def by auto
+ qed fact+
+ then show ?thesis
+ unfolding discrete_altdef g_def
+ using eventually_mono by fastforce
+qed
+
+lemma meromorphic_isolated_in:
+ assumes merf: "f meromorphic_on D pts" "p\<in>pts"
+ shows "p isolated_in pts"
+ by (meson assms isolated_in_islimpt_iff meromorphic_on_def subsetD)
+
+lemma remove_sings_constant_on:
+ assumes merf: "f meromorphic_on D pts" and "connected D"
+ and const:"f constant_on (D - pts)"
+ shows "(remove_sings f) constant_on D"
+proof -
+ have remove_sings_const: "remove_sings f constant_on D - pts"
+ using const
+ by (metis constant_onE merf meromorphic_on_imp_analytic_at remove_sings_at_analytic)
+
+ have ?thesis if "D = {}"
+ using that unfolding constant_on_def by auto
+ moreover have ?thesis if "D\<noteq>{}" "{x\<in>pts. is_pole f x} = {}"
+ proof -
+ obtain \<xi> where "\<xi> \<in> (D - pts)" "\<xi> islimpt (D - pts)"
+ proof -
+ have "open (D - pts)"
+ using meromorphic_imp_open_diff[OF merf] .
+ moreover have "(D - pts) \<noteq> {}" using \<open>D\<noteq>{}\<close>
+ by (metis Diff_empty closure_empty merf
+ meromorphic_pts_closure subset_empty)
+ ultimately show ?thesis using open_imp_islimpt that by auto
+ qed
+ moreover have "remove_sings f holomorphic_on D"
+ using remove_sings_holomorphic_on[OF merf] that by auto
+ moreover note remove_sings_const
+ moreover have "open D"
+ using assms(1) meromorphic_on_def by blast
+ ultimately show ?thesis
+ using Conformal_Mappings.analytic_continuation'
+ [of "remove_sings f" D "D-pts" \<xi>] \<open>connected D\<close>
+ by auto
+ qed
+ moreover have ?thesis if "D\<noteq>{}" "{x\<in>pts. is_pole f x} \<noteq> {}"
+ proof -
+ define PP where "PP={x\<in>D. is_pole f x}"
+ have "remove_sings f meromorphic_on D PP"
+ using merf unfolding PP_def
+ apply (elim remove_sings_meromorphic_on)
+ subgoal using assms(1) meromorphic_on_def by force
+ subgoal using meromorphic_pole_subset merf by auto
+ done
+ moreover have "remove_sings f constant_on D - PP"
+ proof -
+ obtain \<xi> where "\<xi> \<in> f ` (D - pts)"
+ by (metis Diff_empty Diff_eq_empty_iff \<open>D \<noteq> {}\<close> assms(1)
+ closure_empty ex_in_conv imageI meromorphic_pts_closure)
+ have \<xi>:"\<forall>x\<in>D - pts. f x = \<xi>"
+ by (metis \<open>\<xi> \<in> f ` (D - pts)\<close> assms(3) constant_on_def image_iff)
+
+ have "remove_sings f x = \<xi>" if "x\<in>D - PP" for x
+ proof (cases "x\<in>pts")
+ case True
+ then have"x isolated_in pts"
+ using meromorphic_isolated_in[OF merf] by auto
+ then obtain T0 where T0:"open T0" "T0 \<inter> pts = {x}"
+ unfolding isolated_in_def by auto
+ obtain T1 where T1:"open T1" "x\<in>T1" "T1 \<subseteq> D"
+ using merf unfolding meromorphic_on_def
+ using True by blast
+ define T2 where "T2 = T1 \<inter> T0"
+ have "open T2" "x\<in>T2" "T2 - {x} \<subseteq> D - pts"
+ using T0 T1 unfolding T2_def by auto
+ then have "\<forall>w\<in>T2. w\<noteq>x \<longrightarrow> f w =\<xi>"
+ using \<xi> by auto
+ then have "\<forall>\<^sub>F x in at x. f x = \<xi>"
+ unfolding eventually_at_topological
+ using \<open>open T2\<close> \<open>x\<in>T2\<close> by auto
+ then have "f \<midarrow>x\<rightarrow> \<xi>"
+ using tendsto_eventually by auto
+ then show ?thesis by blast
+ next
+ case False
+ then show ?thesis
+ using \<open>\<forall>x\<in>D - pts. f x = \<xi>\<close> assms(1)
+ meromorphic_on_imp_analytic_at that by auto
+ qed
+
+ then show ?thesis unfolding constant_on_def by auto
+ qed
+
+ moreover have "is_pole (remove_sings f) x" if "x\<in>PP" for x
+ proof -
+ have "isolated_singularity_at f x"
+ by (metis (mono_tags, lifting) DiffI PP_def assms(1)
+ isolated_singularity_at_analytic mem_Collect_eq
+ meromorphic_on_def meromorphic_on_imp_analytic_at that)
+ then show ?thesis using that unfolding PP_def by simp
+ qed
+ ultimately show ?thesis
+ using meromorphic_imp_constant_on
+ [of "remove_sings f" D PP]
+ by auto
+ qed
+ ultimately show ?thesis by auto
+qed
+
+lemma meromorphic_eq_meromorphic_extend:
+ assumes "f meromorphic_on A pts1" "g meromorphic_on A pts1" "\<not>z islimpt pts2"
+ assumes "\<And>z. z \<in> A - pts2 \<Longrightarrow> f z = g z" "pts1 \<subseteq> pts2" "z \<in> A - pts1"
+ shows "f z = g z"
+proof -
+ have "g analytic_on {z}"
+ using assms by (intro meromorphic_on_imp_analytic_at[OF assms(2)]) auto
+ hence "g \<midarrow>z\<rightarrow> g z"
+ using analytic_at_imp_isCont isContD by blast
+ also have "?this \<longleftrightarrow> f \<midarrow>z\<rightarrow> g z"
+ proof (intro filterlim_cong)
+ have "eventually (\<lambda>w. w \<notin> pts2) (at z)"
+ using assms by (auto simp: islimpt_conv_frequently_at frequently_def)
+ moreover have "eventually (\<lambda>w. w \<in> A) (at z)"
+ using assms by (intro eventually_at_in_open') (auto simp: meromorphic_on_def)
+ ultimately show "\<forall>\<^sub>F x in at z. g x = f x"
+ by eventually_elim (use assms in auto)
+ qed auto
+ finally have "f \<midarrow>z\<rightarrow> g z" .
+ moreover have "f analytic_on {z}"
+ using assms by (intro meromorphic_on_imp_analytic_at[OF assms(1)]) auto
+ hence "f \<midarrow>z\<rightarrow> f z"
+ using analytic_at_imp_isCont isContD by blast
+ ultimately show ?thesis
+ using tendsto_unique by force
+qed
+
+lemma meromorphic_constant_on_extend:
+ assumes "f constant_on A - pts1" "f meromorphic_on A pts1" "f meromorphic_on A pts2" "pts2 \<subseteq> pts1"
+ shows "f constant_on A - pts2"
+proof -
+ from assms(1) obtain c where c: "\<And>z. z \<in> A - pts1 \<Longrightarrow> f z = c"
+ unfolding constant_on_def by auto
+ have "f z = c" if "z \<in> A - pts2" for z
+ using assms(3)
+ proof (rule meromorphic_eq_meromorphic_extend[where z = z])
+ show "(\<lambda>a. c) meromorphic_on A pts2"
+ by (intro meromorphic_on_const) (use assms in \<open>auto simp: meromorphic_on_def\<close>)
+ show "\<not>z islimpt pts1"
+ using that assms by (auto simp: meromorphic_on_def)
+ qed (use assms c that in auto)
+ thus ?thesis
+ by (auto simp: constant_on_def)
+qed
+
+lemma meromorphic_remove_sings_constant_on_imp_constant_on:
+ assumes "f meromorphic_on A pts"
+ assumes "remove_sings f constant_on A"
+ shows "f constant_on A - pts"
+proof -
+ from assms(2) obtain c where c: "\<And>z. z \<in> A \<Longrightarrow> remove_sings f z = c"
+ by (auto simp: constant_on_def)
+ have "f z = c" if "z \<in> A - pts" for z
+ using meromorphic_on_imp_analytic_at[OF assms(1) that] c[of z] that
+ by auto
+ thus ?thesis
+ by (auto simp: constant_on_def)
+qed
+
+
+
+
+definition singularities_on :: "complex set \<Rightarrow> (complex \<Rightarrow> complex) \<Rightarrow> complex set" where
+ "singularities_on A f =
+ {z\<in>A. isolated_singularity_at f z \<and> not_essential f z \<and> \<not>f analytic_on {z}}"
+
+lemma singularities_on_subset: "singularities_on A f \<subseteq> A"
+ by (auto simp: singularities_on_def)
+
+lemma pole_in_singularities_on:
+ assumes "f meromorphic_on A pts" "z \<in> A" "is_pole f z"
+ shows "z \<in> singularities_on A f"
+ unfolding singularities_on_def not_essential_def using assms
+ using analytic_at_imp_no_pole meromorphic_on_imp_isolated_singularity by force
+
+
+lemma meromorphic_on_subset_pts:
+ assumes "f meromorphic_on A pts" "pts' \<subseteq> pts" "f analytic_on pts - pts'"
+ shows "f meromorphic_on A pts'"
+proof
+ show "open A" "pts' \<subseteq> A"
+ using assms by (auto simp: meromorphic_on_def)
+ show "isolated_singularity_at f z" "not_essential f z" if "z \<in> pts'" for z
+ using assms that by (auto simp: meromorphic_on_def)
+ show "\<not>z islimpt pts'" if "z \<in> A" for z
+ using assms that islimpt_subset unfolding meromorphic_on_def by blast
+ have "f analytic_on A - pts"
+ using assms(1) meromorphic_imp_analytic by blast
+ with assms have "f analytic_on (A - pts) \<union> (pts - pts')"
+ by (subst analytic_on_Un) auto
+ also have "(A - pts) \<union> (pts - pts') = A - pts'"
+ using assms by (auto simp: meromorphic_on_def)
+ finally show "f holomorphic_on A - pts'"
+ using analytic_imp_holomorphic by blast
+qed
+
+lemma meromorphic_on_imp_superset_singularities_on:
+ assumes "f meromorphic_on A pts"
+ shows "singularities_on A f \<subseteq> pts"
+proof
+ fix z assume "z \<in> singularities_on A f"
+ hence "z \<in> A" "\<not>f analytic_on {z}"
+ by (auto simp: singularities_on_def)
+ with assms show "z \<in> pts"
+ by (meson DiffI meromorphic_on_imp_analytic_at)
+qed
+
+lemma meromorphic_on_singularities_on:
+ assumes "f meromorphic_on A pts"
+ shows "f meromorphic_on A (singularities_on A f)"
+ using assms meromorphic_on_imp_superset_singularities_on[OF assms]
+proof (rule meromorphic_on_subset_pts)
+ have "f analytic_on {z}" if "z \<in> pts - singularities_on A f" for z
+ using that assms by (auto simp: singularities_on_def meromorphic_on_def)
+ thus "f analytic_on pts - singularities_on A f"
+ using analytic_on_analytic_at by blast
+qed
+
+theorem Residue_theorem_inside:
+ assumes f: "f meromorphic_on s pts"
+ "simply_connected s"
+ assumes g: "valid_path g"
+ "pathfinish g = pathstart g"
+ "path_image g \<subseteq> s - pts"
+ defines "pts1 \<equiv> pts \<inter> inside (path_image g)"
+ shows "finite pts1"
+ and "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts1. winding_number g p * residue f p)"
+proof -
+ note [dest] = valid_path_imp_path
+ have cl_g [intro]: "closed (path_image g)"
+ using g by (auto intro!: closed_path_image)
+ have "open s"
+ using f(1) by (auto simp: meromorphic_on_def)
+ define pts2 where "pts2 = pts - pts1"
+
+ define A where "A = path_image g \<union> inside (path_image g)"
+ have "closed A"
+ unfolding A_def using g by (intro closed_path_image_Un_inside) auto
+ moreover have "bounded A"
+ unfolding A_def using g by (auto intro!: bounded_path_image bounded_inside)
+ ultimately have 1: "compact A"
+ using compact_eq_bounded_closed by blast
+ have 2: "open (s - pts2)"
+ using f by (auto intro!: meromorphic_imp_open_diff' [OF f(1)] simp: pts2_def)
+ have 3: "A \<subseteq> s - pts2"
+ unfolding A_def pts2_def pts1_def
+ using f(2) g(3) 2 subset_simply_connected_imp_inside_subset[of s "path_image g"] \<open>open s\<close>
+ by auto
+
+ obtain \<epsilon> where \<epsilon>: "\<epsilon> > 0" "(\<Union>x\<in>A. ball x \<epsilon>) \<subseteq> s - pts2"
+ using compact_subset_open_imp_ball_epsilon_subset[OF 1 2 3] by blast
+ define B where "B = (\<Union>x\<in>A. ball x \<epsilon>)"
+
+ have "finite (A \<inter> pts)"
+ using 1 3 by (intro meromorphic_compact_finite_pts[OF f(1)]) auto
+ also have "A \<inter> pts = pts1"
+ unfolding pts1_def using g by (auto simp: A_def)
+ finally show fin: "finite pts1" .
+
+ show "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts1. winding_number g p * residue f p)"
+ proof (rule Residue_theorem)
+ show "open B"
+ by (auto simp: B_def)
+ next
+ have "connected A"
+ unfolding A_def using g
+ by (intro connected_with_inside closed_path_image connected_path_image) auto
+ hence "connected (A \<union> B)"
+ unfolding B_def using g \<open>\<epsilon> > 0\<close> f(2)
+ by (intro connected_Un_UN connected_path_image valid_path_imp_path)
+ (auto simp: simply_connected_imp_connected)
+ also have "A \<union> B = B"
+ using \<epsilon>(1) by (auto simp: B_def)
+ finally show "connected B" .
+ next
+ have "f holomorphic_on (s - pts)"
+ by (intro meromorphic_imp_holomorphic f)
+ moreover have "B - pts1 \<subseteq> s - pts"
+ using \<epsilon> unfolding B_def by (auto simp: pts1_def pts2_def)
+ ultimately show "f holomorphic_on (B - pts1)"
+ by (rule holomorphic_on_subset)
+ next
+ have "path_image g \<subseteq> A - pts1"
+ using g unfolding pts1_def by (auto simp: A_def)
+ also have "\<dots> \<subseteq> B - pts1"
+ unfolding B_def using \<epsilon>(1) by auto
+ finally show "path_image g \<subseteq> B - pts1" .
+ next
+ show "\<forall>z. z \<notin> B \<longrightarrow> winding_number g z = 0"
+ proof safe
+ fix z assume z: "z \<notin> B"
+ hence "z \<notin> A"
+ using \<epsilon>(1) by (auto simp: B_def)
+ hence "z \<in> outside (path_image g)"
+ unfolding A_def by (simp add: union_with_inside)
+ thus "winding_number g z = 0"
+ using g by (intro winding_number_zero_in_outside) auto
+ qed
+ qed (use g fin in auto)
+qed
+
+theorem Residue_theorem':
+ assumes f: "f meromorphic_on s pts"
+ "simply_connected s"
+ assumes g: "valid_path g"
+ "pathfinish g = pathstart g"
+ "path_image g \<subseteq> s - pts"
+ assumes pts': "finite pts'"
+ "pts' \<subseteq> s"
+ "\<And>z. z \<in> pts - pts' \<Longrightarrow> winding_number g z = 0"
+ shows "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts'. winding_number g p * residue f p)"
+proof -
+ note [dest] = valid_path_imp_path
+ define pts1 where "pts1 = pts \<inter> inside (path_image g)"
+
+ have "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts1. winding_number g p * residue f p)"
+ unfolding pts1_def by (intro Residue_theorem_inside[OF f g])
+ also have "(\<Sum>p\<in>pts1. winding_number g p * residue f p) =
+ (\<Sum>p\<in>pts'. winding_number g p * residue f p)"
+ proof (intro sum.mono_neutral_cong refl)
+ show "finite pts1"
+ unfolding pts1_def by (intro Residue_theorem_inside[OF f g])
+ show "finite pts'"
+ by fact
+ next
+ fix z assume z: "z \<in> pts' - pts1"
+ show "winding_number g z * residue f z = 0"
+ proof (cases "z \<in> pts")
+ case True
+ with z have "z \<notin> path_image g \<union> inside (path_image g)"
+ using g(3) by (auto simp: pts1_def)
+ hence "z \<in> outside (path_image g)"
+ by (simp add: union_with_inside)
+ hence "winding_number g z = 0"
+ using g by (intro winding_number_zero_in_outside) auto
+ thus ?thesis
+ by simp
+ next
+ case False
+ with z pts' have "z \<in> s - pts"
+ by auto
+ with f(1) have "f analytic_on {z}"
+ by (intro meromorphic_on_imp_analytic_at)
+ hence "residue f z = 0"
+ using analytic_at residue_holo by blast
+ thus ?thesis
+ by simp
+ qed
+ next
+ fix z assume z: "z \<in> pts1 - pts'"
+ hence "winding_number g z = 0"
+ using pts' by (auto simp: pts1_def)
+ thus "winding_number g z * residue f z = 0"
+ by simp
+ qed
+ finally show ?thesis .
+qed
+
+end
\ No newline at end of file
--- a/src/HOL/Complex_Analysis/Residue_Theorem.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Complex_Analysis/Residue_Theorem.thy Mon Feb 20 13:59:42 2023 +0100
@@ -3,6 +3,34 @@
imports Complex_Residues "HOL-Library.Landau_Symbols"
begin
+text \<open>Could be moved to a previous theory importing both Landau Symbols and Elementary Metric Spaces\<close>
+lemma continuous_bounded_at_infinity_imp_bounded:
+ fixes f :: "real \<Rightarrow> 'a :: real_normed_field"
+ assumes "f \<in> O[at_bot](\<lambda>_. 1)"
+ assumes "f \<in> O[at_top](\<lambda>_. 1)"
+ assumes "continuous_on UNIV f"
+ shows "bounded (range f)"
+proof -
+ from assms(1) obtain c1 where "eventually (\<lambda>x. norm (f x) \<le> c1) at_bot"
+ by (auto elim!: landau_o.bigE)
+ then obtain x1 where x1: "\<And>x. x \<le> x1 \<Longrightarrow> norm (f x) \<le> c1"
+ by (auto simp: eventually_at_bot_linorder)
+ from assms(2) obtain c2 where "eventually (\<lambda>x. norm (f x) \<le> c2) at_top"
+ by (auto elim!: landau_o.bigE)
+ then obtain x2 where x2: "\<And>x. x \<ge> x2 \<Longrightarrow> norm (f x) \<le> c2"
+ by (auto simp: eventually_at_top_linorder)
+ have "compact (f ` {x1..x2})"
+ by (intro compact_continuous_image continuous_on_subset[OF assms(3)]) auto
+ hence "bounded (f ` {x1..x2})"
+ by (rule compact_imp_bounded)
+ then obtain c3 where c3: "\<And>x. x \<in> {x1..x2} \<Longrightarrow> norm (f x) \<le> c3"
+ unfolding bounded_iff by fast
+ have "norm (f x) \<le> Max {c1, c2, c3}" for x
+ by (cases "x \<le> x1"; cases "x \<ge> x2") (use x1 x2 c3 in \<open>auto simp: le_max_iff_disj\<close>)
+ thus ?thesis
+ unfolding bounded_iff by blast
+qed
+
subsection \<open>Cauchy's residue theorem\<close>
lemma get_integrable_path:
--- a/src/HOL/Complex_Analysis/Riemann_Mapping.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Complex_Analysis/Riemann_Mapping.thy Mon Feb 20 13:59:42 2023 +0100
@@ -1245,21 +1245,9 @@
interpret SC_Chain
using assms by (simp add: SC_Chain_def)
have "?fp \<and> ?ucc \<and> ?ei"
-proof -
- have *: "\<lbrakk>\<alpha> \<Longrightarrow> \<beta>; \<beta> \<Longrightarrow> \<gamma>; \<gamma> \<Longrightarrow> \<delta>; \<delta> \<Longrightarrow> \<alpha>\<rbrakk>
- \<Longrightarrow> (\<alpha> \<longleftrightarrow> \<beta>) \<and> (\<alpha> \<longleftrightarrow> \<gamma>) \<and> (\<alpha> \<longleftrightarrow> \<delta>)" for \<alpha> \<beta> \<gamma> \<delta>
- by blast
- show ?thesis
- apply (rule *)
- using frontier_properties simply_connected_imp_connected apply blast
-apply clarify
- using unbounded_complement_components simply_connected_imp_connected apply blast
- using empty_inside apply blast
- using empty_inside_imp_simply_connected apply blast
- done
-qed
+ using empty_inside empty_inside_imp_simply_connected frontier_properties unbounded_complement_components winding_number_zero by blast
then show ?fp ?ucc ?ei
- by safe
+ by blast+
qed
lemma simply_connected_iff_simple:
@@ -1270,6 +1258,12 @@
apply (metis DIM_complex assms(2) cobounded_has_bounded_component double_compl order_refl)
by (meson assms inside_bounded_complement_connected_empty simply_connected_eq_empty_inside simply_connected_eq_unbounded_complement_components)
+lemma subset_simply_connected_imp_inside_subset:
+ fixes A :: "complex set"
+ assumes "simply_connected A" "open A" "B \<subseteq> A"
+ shows "inside B \<subseteq> A"
+by (metis assms Diff_eq_empty_iff inside_mono subset_empty simply_connected_eq_empty_inside)
+
subsection\<open>Further equivalences based on continuous logs and sqrts\<close>
context SC_Chain
@@ -1299,9 +1293,7 @@
and expg: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> (f \<circ> k) z = exp (g z)"
proof (rule continuous_logarithm_on_ball)
show "continuous_on (ball 0 1) (f \<circ> k)"
- apply (rule continuous_on_compose [OF contk])
- using kim continuous_on_subset [OF contf]
- by blast
+ using contf continuous_on_compose contk kim by blast
show "\<And>z. z \<in> ball 0 1 \<Longrightarrow> (f \<circ> k) z \<noteq> 0"
using kim nz by auto
qed auto
@@ -1438,9 +1430,7 @@
lemma Borsukian_eq_simply_connected:
fixes S :: "complex set"
shows "open S \<Longrightarrow> Borsukian S \<longleftrightarrow> (\<forall>C \<in> components S. simply_connected C)"
-apply (auto simp: Borsukian_componentwise_eq open_imp_locally_connected)
- using in_components_connected open_components simply_connected_eq_Borsukian apply blast
- using open_components simply_connected_eq_Borsukian by blast
+ by (meson Borsukian_componentwise_eq in_components_connected open_components open_imp_locally_connected simply_connected_eq_Borsukian)
lemma Borsukian_separation_open_closed:
fixes S :: "complex set"
@@ -1451,7 +1441,7 @@
assume "open S"
show ?thesis
unfolding Borsukian_eq_simply_connected [OF \<open>open S\<close>]
- by (meson \<open>open S\<close> \<open>bounded S\<close> bounded_subset in_components_connected in_components_subset nonseparation_by_component_eq open_components simply_connected_iff_simple)
+ by (metis \<open>open S\<close> \<open>bounded S\<close> bounded_subset in_components_maximal nonseparation_by_component_eq open_components simply_connected_iff_simple)
next
assume "closed S"
with \<open>bounded S\<close> show ?thesis
@@ -1659,10 +1649,8 @@
proof
assume ?lhs
then show ?rhs
- apply (auto simp: winding_number_homotopic_loops_null_eq [OF assms])
- apply (rule homotopic_loops_imp_homotopic_paths_null)
- apply (simp add: linepath_refl)
- done
+ using homotopic_loops_imp_homotopic_paths_null
+ by (force simp add: linepath_refl winding_number_homotopic_loops_null_eq [OF assms])
next
assume ?rhs
then show ?lhs
@@ -1701,11 +1689,7 @@
then show ?rhs
using homotopic_paths_imp_pathstart assms
by (fastforce simp add: dest: homotopic_paths_imp_homotopic_loops homotopic_paths_loop_parts)
-next
- assume ?rhs
- then show ?lhs
- by (simp add: winding_number_homotopic_paths)
-qed
+qed (simp add: winding_number_homotopic_paths)
lemma winding_number_homotopic_loops_eq:
assumes "path p" and \<zeta>p: "\<zeta> \<notin> path_image p"
@@ -1725,17 +1709,20 @@
then have "pathstart r \<noteq> \<zeta>" by blast
have "homotopic_loops (- {\<zeta>}) p (r +++ q +++ reversepath r)"
proof (rule homotopic_paths_imp_homotopic_loops)
- show "homotopic_paths (- {\<zeta>}) p (r +++ q +++ reversepath r)"
- by (metis (mono_tags, opaque_lifting) \<open>path r\<close> L \<zeta>p \<zeta>q \<open>path p\<close> \<open>path q\<close> homotopic_loops_conjugate loops not_in_path_image_join paf pas path_image_reversepath path_imp_reversepath path_join_eq pathfinish_join pathfinish_reversepath pathstart_join pathstart_reversepath rim subset_Compl_singleton winding_number_homotopic_loops winding_number_homotopic_paths_eq)
+ have "path (r +++ q +++ reversepath r)"
+ by (simp add: \<open>path r\<close> \<open>path q\<close> loops paf)
+ moreover have "\<zeta> \<notin> path_image (r +++ q +++ reversepath r)"
+ by (metis \<zeta>q not_in_path_image_join path_image_reversepath rim subset_Compl_singleton)
+ moreover have "homotopic_loops (- {\<zeta>}) (r +++ q +++ reversepath r) q"
+ using \<open>path q\<close> \<open>path r\<close> \<zeta>q homotopic_loops_conjugate loops(2) paf rim by blast
+ ultimately show "homotopic_paths (- {\<zeta>}) p (r +++ q +++ reversepath r)"
+ using loops pathfinish_join pathfinish_reversepath pathstart_join
+ by (metis L \<zeta>p \<open>path p\<close> pas winding_number_homotopic_loops winding_number_homotopic_paths_eq)
qed (use loops pas in auto)
moreover have "homotopic_loops (- {\<zeta>}) (r +++ q +++ reversepath r) q"
using rim \<zeta>q by (auto simp: homotopic_loops_conjugate paf \<open>path q\<close> \<open>path r\<close> loops)
ultimately show ?rhs
using homotopic_loops_trans by metis
-next
- assume ?rhs
- then show ?lhs
- by (simp add: winding_number_homotopic_loops)
-qed
+qed (simp add: winding_number_homotopic_loops)
end
--- a/src/HOL/Complex_Analysis/Winding_Numbers.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Complex_Analysis/Winding_Numbers.thy Mon Feb 20 13:59:42 2023 +0100
@@ -901,6 +901,16 @@
by metis
qed
+lemma bounded_winding_number_nz:
+ assumes "path g" "pathfinish g = pathstart g"
+ shows "bounded {z. winding_number g z \<noteq> 0}"
+proof -
+ obtain B where "\<And>x. norm x \<ge> B \<Longrightarrow> winding_number g x = 0"
+ using winding_number_zero_at_infinity[OF assms] by auto
+ thus ?thesis
+ unfolding bounded_iff by (intro exI[of _ "B + 1"]) force
+qed
+
lemma winding_number_zero_point:
"\<lbrakk>path \<gamma>; convex S; pathfinish \<gamma> = pathstart \<gamma>; open S; path_image \<gamma> \<subseteq> S\<rbrakk>
\<Longrightarrow> \<exists>z. z \<in> S \<and> winding_number \<gamma> z = 0"
--- a/src/HOL/Computational_Algebra/Formal_Laurent_Series.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Computational_Algebra/Formal_Laurent_Series.thy Mon Feb 20 13:59:42 2023 +0100
@@ -332,6 +332,9 @@
lemma fls_shift_eq0_iff: "fls_shift m f = 0 \<longleftrightarrow> f = 0"
using fls_shift_eq_iff[of m f 0] by simp
+lemma fls_shift_eq_1_iff: "fls_shift n f = 1 \<longleftrightarrow> f = fls_shift (-n) 1"
+ by (metis add_minus_cancel fls_shift_eq_iff fls_shift_fls_shift)
+
lemma fls_shift_nonneg_subdegree: "m \<le> fls_subdegree f \<Longrightarrow> fls_subdegree (fls_shift m f) \<ge> 0"
by (cases "f=0") (auto intro: fls_subdegree_geI)
@@ -699,6 +702,9 @@
thus "f $ n = g $ n" by simp
qed
+lemma fps_to_fls_eq_iff [simp]: "fps_to_fls f = fps_to_fls g \<longleftrightarrow> f = g"
+ using fps_to_fls_eq_imp_fps_eq by blast
+
lemma fps_zero_to_fls [simp]: "fps_to_fls 0 = 0"
by (intro fls_zero_eqI) simp
@@ -723,9 +729,12 @@
lemma fps_X_to_fls [simp]: "fps_to_fls fps_X = fls_X"
by (fastforce intro: fls_eqI)
-lemma fps_to_fls_eq_zero_iff: "(fps_to_fls f = 0) \<longleftrightarrow> (f=0)"
+lemma fps_to_fls_eq_0_iff [simp]: "(fps_to_fls f = 0) \<longleftrightarrow> (f=0)"
using fps_to_fls_nonzeroI by auto
+lemma fps_to_fls_eq_1_iff [simp]: "fps_to_fls f = 1 \<longleftrightarrow> f = 1"
+ using fps_to_fls_eq_iff by fastforce
+
lemma fls_subdegree_fls_to_fps_gt0: "fls_subdegree (fps_to_fls f) \<ge> 0"
proof (cases "f=0")
case False show ?thesis
@@ -780,6 +789,25 @@
using fls_base_factor_fps_to_fls[of f] fls_regpart_fps_trivial[of "unit_factor f"]
by simp
+lemma fls_as_fps:
+ fixes f :: "'a :: zero fls" and n :: int
+ assumes n: "n \<ge> -fls_subdegree f"
+ obtains f' where "f = fls_shift n (fps_to_fls f')"
+proof -
+ have "fls_subdegree (fls_shift (- n) f) \<ge> 0"
+ by (rule fls_shift_nonneg_subdegree) (use n in simp)
+ hence "f = fls_shift n (fps_to_fls (fls_regpart (fls_shift (-n) f)))"
+ by (subst fls_regpart_to_fls_trivial) simp_all
+ thus ?thesis
+ by (rule that)
+qed
+
+lemma fls_as_fps':
+ fixes f :: "'a :: zero fls" and n :: int
+ assumes n: "n \<ge> -fls_subdegree f"
+ shows "\<exists>f'. f = fls_shift n (fps_to_fls f')"
+ using fls_as_fps[OF assms] by metis
+
abbreviation
"fls_regpart_as_fls f \<equiv> fps_to_fls (fls_regpart f)"
abbreviation
@@ -1719,10 +1747,12 @@
by (simp add: fls_times_conv_fps_times)
qed simp
+lemma fps_to_fls_power: "fps_to_fls (f ^ n) = fps_to_fls f ^ n"
+ by (simp add: fls_pow_conv_fps_pow fls_subdegree_fls_to_fps_gt0)
+
lemma fls_pow_conv_regpart:
"fls_subdegree f \<ge> 0 \<Longrightarrow> fls_regpart (f ^ n) = (fls_regpart f) ^ n"
- using fls_pow_subdegree_ge0[of f n] fls_pow_conv_fps_pow[of f n]
- by (intro fps_to_fls_eq_imp_fps_eq) simp
+ by (simp add: fls_pow_conv_fps_pow)
text \<open>These two lemmas show that shifting 1 is equivalent to powers of the implied variable.\<close>
@@ -1995,17 +2025,7 @@
lemma fls_lr_inverse_eq0_imp_starting0:
"fls_left_inverse f x = 0 \<Longrightarrow> x = 0"
"fls_right_inverse f x = 0 \<Longrightarrow> x = 0"
-proof-
- assume "fls_left_inverse f x = 0"
- hence "fps_left_inverse (fls_base_factor_to_fps f) x = 0"
- using fls_shift_eq_iff fps_to_fls_eq_zero_iff by fastforce
- thus "x = 0" using fps_lr_inverse_eq0_imp_starting0(1) by fast
-next
- assume "fls_right_inverse f x = 0"
- hence "fps_right_inverse (fls_base_factor_to_fps f) x = 0"
- using fls_shift_eq_iff fps_to_fls_eq_zero_iff by fastforce
- thus "x = 0" using fps_lr_inverse_eq0_imp_starting0(2) by fast
-qed
+ by (metis fls_lr_inverse_base fls_nonzeroI)+
lemma fls_lr_inverse_eq_0_iff:
fixes x :: "'a::{comm_monoid_add,mult_zero,uminus}"
@@ -3231,10 +3251,146 @@
thus "\<exists>g. 1 = g * f" by fast
qed
+subsection \<open>Composition\<close>
+
+definition fls_compose_fps :: "'a :: field fls \<Rightarrow> 'a fps \<Rightarrow> 'a fls" where
+ "fls_compose_fps f g =
+ (if f = 0 then 0
+ else if fls_subdegree f \<ge> 0 then fps_to_fls (fps_compose (fls_regpart f) g)
+ else fps_to_fls (fps_compose (fls_base_factor_to_fps f) g) /
+ fps_to_fls g ^ nat (-fls_subdegree f))"
+
+lemma fls_compose_fps_fps [simp]:
+ "fls_compose_fps (fps_to_fls f) g = fps_to_fls (fps_compose f g)"
+ by (simp add: fls_compose_fps_def fls_subdegree_fls_to_fps_gt0 fps_to_fls_eq_0_iff)
+
+lemma fls_const_transfer [transfer_rule]:
+ "rel_fun (=) (pcr_fls (=))
+ (\<lambda>c n. if n = 0 then c else 0) fls_const"
+ by (auto simp: fls_const_def rel_fun_def pcr_fls_def OO_def cr_fls_def)
+
+lemma fls_shift_transfer [transfer_rule]:
+ "rel_fun (=) (rel_fun (pcr_fls (=)) (pcr_fls (=)))
+ (\<lambda>n f k. f (k+n)) fls_shift"
+ by (auto simp: fls_const_def rel_fun_def pcr_fls_def OO_def cr_fls_def)
+
+lift_definition fls_compose_power :: "'a :: zero fls \<Rightarrow> nat \<Rightarrow> 'a fls" is
+ "\<lambda>f d n. if d > 0 \<and> int d dvd n then f (n div int d) else 0"
+proof -
+ fix f :: "int \<Rightarrow> 'a" and d :: nat
+ assume *: "eventually (\<lambda>n. f (-int n) = 0) cofinite"
+ show "eventually (\<lambda>n. (if d > 0 \<and> int d dvd -int n then f (-int n div int d) else 0) = 0) cofinite"
+ proof (cases "d = 0")
+ case False
+ from * have "eventually (\<lambda>n. f (-int n) = 0) at_top"
+ by (simp add: cofinite_eq_sequentially)
+ hence "eventually (\<lambda>n. f (-int (n div d)) = 0) at_top"
+ by (rule eventually_compose_filterlim[OF _ filterlim_at_top_div_const_nat]) (use False in auto)
+ hence "eventually (\<lambda>n. (if d > 0 \<and> int d dvd -int n then f (-int n div int d) else 0) = 0) at_top"
+ by eventually_elim (auto simp: zdiv_int dvd_neg_div)
+ thus ?thesis
+ by (simp add: cofinite_eq_sequentially)
+ qed auto
+qed
+
+lemma fls_nth_compose_power:
+ assumes "d > 0"
+ shows "fls_nth (fls_compose_power f d) n = (if int d dvd n then fls_nth f (n div int d) else 0)"
+ using assms by transfer auto
+
+
+lemma fls_compose_power_0_left [simp]: "fls_compose_power 0 d = 0"
+ by transfer auto
+
+lemma fls_compose_power_1_left [simp]: "d > 0 \<Longrightarrow> fls_compose_power 1 d = 1"
+ by transfer (auto simp: fun_eq_iff)
+
+lemma fls_compose_power_const_left [simp]:
+ "d > 0 \<Longrightarrow> fls_compose_power (fls_const c) d = fls_const c"
+ by transfer (auto simp: fun_eq_iff)
+
+lemma fls_compose_power_shift [simp]:
+ "d > 0 \<Longrightarrow> fls_compose_power (fls_shift n f) d = fls_shift (d * n) (fls_compose_power f d)"
+ by transfer (auto simp: fun_eq_iff add_ac mult_ac)
+
+lemma fls_compose_power_X_intpow [simp]:
+ "d > 0 \<Longrightarrow> fls_compose_power (fls_X_intpow n) d = fls_X_intpow (int d * n)"
+ by simp
+
+lemma fls_compose_power_X [simp]:
+ "d > 0 \<Longrightarrow> fls_compose_power fls_X d = fls_X_intpow (int d)"
+ by transfer (auto simp: fun_eq_iff)
+
+lemma fls_compose_power_X_inv [simp]:
+ "d > 0 \<Longrightarrow> fls_compose_power fls_X_inv d = fls_X_intpow (-int d)"
+ by (simp add: fls_X_inv_conv_shift_1)
+
+lemma fls_compose_power_0_right [simp]: "fls_compose_power f 0 = 0"
+ by transfer auto
+
+lemma fls_compose_power_add [simp]:
+ "fls_compose_power (f + g) d = fls_compose_power f d + fls_compose_power g d"
+ by transfer auto
+
+lemma fls_compose_power_diff [simp]:
+ "fls_compose_power (f - g) d = fls_compose_power f d - fls_compose_power g d"
+ by transfer auto
+
+lemma fls_compose_power_uminus [simp]:
+ "fls_compose_power (-f) d = -fls_compose_power f d"
+ by transfer auto
+
+lemma fps_nth_compose_X_power:
+ "fps_nth (f oo (fps_X ^ d)) n = (if d dvd n then fps_nth f (n div d) else 0)"
+proof -
+ have "fps_nth (f oo (fps_X ^ d)) n = (\<Sum>i = 0..n. f $ i * (fps_X ^ (d * i)) $ n)"
+ unfolding fps_compose_def by (simp add: power_mult)
+ also have "\<dots> = (\<Sum>i\<in>(if d dvd n then {n div d} else {}). f $ i * (fps_X ^ (d * i)) $ n)"
+ by (intro sum.mono_neutral_right) auto
+ also have "\<dots> = (if d dvd n then fps_nth f (n div d) else 0)"
+ by auto
+ finally show ?thesis .
+qed
+
+lemma fls_compose_power_fps_to_fls:
+ assumes "d > 0"
+ shows "fls_compose_power (fps_to_fls f) d = fps_to_fls (fps_compose f (fps_X ^ d))"
+ using assms
+ by (intro fls_eqI) (auto simp: fls_nth_compose_power fps_nth_compose_X_power
+ pos_imp_zdiv_neg_iff div_neg_pos_less0 nat_div_distrib
+ simp flip: int_dvd_int_iff)
+
+lemma fls_compose_power_mult [simp]:
+ "fls_compose_power (f * g :: 'a :: idom fls) d = fls_compose_power f d * fls_compose_power g d"
+proof (cases "d > 0")
+ case True
+ define n where "n = nat (max 0 (max (- fls_subdegree f) (- fls_subdegree g)))"
+ have n_ge: "-fls_subdegree f \<le> int n" "-fls_subdegree g \<le> int n"
+ unfolding n_def by auto
+ obtain f' where f': "f = fls_shift n (fps_to_fls f')"
+ using fls_as_fps[OF n_ge(1)] by (auto simp: n_def)
+ obtain g' where g': "g = fls_shift n (fps_to_fls g')"
+ using fls_as_fps[OF n_ge(2)] by (auto simp: n_def)
+ show ?thesis using \<open>d > 0\<close>
+ by (simp add: f' g' fls_shifted_times_simps mult_ac fls_compose_power_fps_to_fls
+ fps_compose_mult_distrib flip: fls_times_fps_to_fls)
+qed auto
+
+lemma fls_compose_power_power [simp]:
+ assumes "d > 0 \<or> n > 0"
+ shows "fls_compose_power (f ^ n :: 'a :: idom fls) d = fls_compose_power f d ^ n"
+proof (cases "d > 0")
+ case True
+ thus ?thesis by (induction n) auto
+qed (use assms in auto)
+
+lemma fls_nth_compose_power' [simp]:
+ "d = 0 \<or> \<not>d dvd n \<Longrightarrow> fls_nth (fls_compose_power f d) n = 0"
+ "d dvd n \<Longrightarrow> d > 0 \<Longrightarrow> fls_nth (fls_compose_power f d) n = fls_nth f (n div d)"
+ by (transfer; force; fail)+
subsection \<open>Formal differentiation and integration\<close>
-
subsubsection \<open>Derivative definition and basic properties\<close>
definition "fls_deriv f = Abs_fls (\<lambda>n. of_int (n+1) * f$$(n+1))"
--- a/src/HOL/Computational_Algebra/Formal_Power_Series.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Computational_Algebra/Formal_Power_Series.thy Mon Feb 20 13:59:42 2023 +0100
@@ -59,12 +59,9 @@
by (rule LeastI_ex)
moreover have "\<forall>m<?n. f $ m = 0"
by (auto dest: not_less_Least)
- ultimately have "f $ ?n \<noteq> 0 \<and> (\<forall>m<?n. f $ m = 0)" ..
- then show ?thesis ..
+ ultimately show ?thesis by metis
qed
- show ?lhs if ?rhs
- using that by (auto simp add: expand_fps_eq)
-qed
+qed (auto simp: expand_fps_eq)
lemma fps_nonzeroI: "f$n \<noteq> 0 \<Longrightarrow> f \<noteq> 0"
by auto
@@ -165,71 +162,36 @@
lemma subdegreeI:
assumes "f $ d \<noteq> 0" and "\<And>i. i < d \<Longrightarrow> f $ i = 0"
shows "subdegree f = d"
-proof-
- from assms(1) have "f \<noteq> 0" by auto
- moreover from assms(1) have "(LEAST i. f $ i \<noteq> 0) = d"
- proof (rule Least_equality)
- fix e assume "f $ e \<noteq> 0"
- with assms(2) have "\<not>(e < d)" by blast
- thus "e \<ge> d" by simp
- qed
- ultimately show ?thesis unfolding subdegree_def by simp
-qed
+ by (smt (verit) LeastI_ex assms fps_zero_nth linorder_cases not_less_Least subdegree_def)
lemma nth_subdegree_nonzero [simp,intro]: "f \<noteq> 0 \<Longrightarrow> f $ subdegree f \<noteq> 0"
-proof-
- assume "f \<noteq> 0"
- hence "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
- also from \<open>f \<noteq> 0\<close> have "\<exists>n. f$n \<noteq> 0" using fps_nonzero_nth by blast
- from LeastI_ex[OF this] have "f $ (LEAST n. f $ n \<noteq> 0) \<noteq> 0" .
- finally show ?thesis .
-qed
+ using fps_nonzero_nth_minimal subdegreeI by blast
lemma nth_less_subdegree_zero [dest]: "n < subdegree f \<Longrightarrow> f $ n = 0"
-proof (cases "f = 0")
- assume "f \<noteq> 0" and less: "n < subdegree f"
- note less
- also from \<open>f \<noteq> 0\<close> have "subdegree f = (LEAST n. f $ n \<noteq> 0)" by (simp add: subdegree_def)
- finally show "f $ n = 0" using not_less_Least by blast
-qed simp_all
+ by (metis fps_nonzero_nth_minimal fps_zero_nth subdegreeI)
lemma subdegree_geI:
assumes "f \<noteq> 0" "\<And>i. i < n \<Longrightarrow> f$i = 0"
shows "subdegree f \<ge> n"
-proof (rule ccontr)
- assume "\<not>(subdegree f \<ge> n)"
- with assms(2) have "f $ subdegree f = 0" by simp
- moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
- ultimately show False by contradiction
-qed
+ by (meson assms leI nth_subdegree_nonzero)
lemma subdegree_greaterI:
assumes "f \<noteq> 0" "\<And>i. i \<le> n \<Longrightarrow> f$i = 0"
shows "subdegree f > n"
-proof (rule ccontr)
- assume "\<not>(subdegree f > n)"
- with assms(2) have "f $ subdegree f = 0" by simp
- moreover from assms(1) have "f $ subdegree f \<noteq> 0" by simp
- ultimately show False by contradiction
-qed
+ by (meson assms leI nth_subdegree_nonzero)
lemma subdegree_leI:
"f $ n \<noteq> 0 \<Longrightarrow> subdegree f \<le> n"
- by (rule leI) auto
+ using linorder_not_less by blast
lemma subdegree_0 [simp]: "subdegree 0 = 0"
by (simp add: subdegree_def)
lemma subdegree_1 [simp]: "subdegree 1 = 0"
- by (cases "(1::'a) = 0")
- (auto intro: subdegreeI fps_ext simp: subdegree_def)
+ by (metis fps_one_nth nth_subdegree_nonzero subdegree_0)
lemma subdegree_eq_0_iff: "subdegree f = 0 \<longleftrightarrow> f = 0 \<or> f $ 0 \<noteq> 0"
-proof (cases "f = 0")
- assume "f \<noteq> 0"
- thus ?thesis
- using nth_subdegree_nonzero[OF \<open>f \<noteq> 0\<close>] by (fastforce intro!: subdegreeI)
-qed simp_all
+ using nth_subdegree_nonzero subdegree_leI by fastforce
lemma subdegree_eq_0 [simp]: "f $ 0 \<noteq> 0 \<Longrightarrow> subdegree f = 0"
by (simp add: subdegree_eq_0_iff)
@@ -247,12 +209,11 @@
qed simp
lemma subdegree_minus_commute [simp]:
- "subdegree (f-(g::('a::group_add) fps)) = subdegree (g - f)"
-proof (-, cases "g-f=0")
- case True
- have "\<And>n. (f - g) $ n = -((g - f) $ n)" by simp
- with True have "f - g = 0" by (intro fps_ext) simp
- with True show ?thesis by simp
+ fixes f :: "'a::group_add fps"
+ shows "subdegree (f-g) = subdegree (g - f)"
+proof (cases "g-f=0")
+ case True then show ?thesis
+ by (metis fps_sub_nth nth_subdegree_nonzero right_minus_eq)
next
case False show ?thesis
using nth_subdegree_nonzero[OF False] by (fastforce intro: subdegreeI)
@@ -273,12 +234,7 @@
proof-
assume fg: "f + g = 0"
have "\<And>n. f $ n = - g $ n"
- proof-
- fix n
- from fg have "(f + g) $ n = 0" by simp
- hence "f $ n + g $ n - g $ n = - g $ n" by simp
- thus "f $ n = - g $ n" by simp
- qed
+ by (metis add_eq_0_iff equation_minus_iff fg fps_add_nth fps_neg_nth fps_zero_nth)
with assms show False by (auto intro: fps_ext)
qed
thus "f + g \<noteq> 0" by fast
@@ -288,43 +244,38 @@
assumes "f \<noteq> 0"
and "subdegree f < subdegree (g :: 'a::monoid_add fps)"
shows "subdegree (f + g) = subdegree f"
- using assms
- by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
+ using assms by(auto intro: subdegreeI simp: nth_less_subdegree_zero)
lemma subdegree_add_eq2:
assumes "g \<noteq> 0"
and "subdegree g < subdegree (f :: 'a :: monoid_add fps)"
shows "subdegree (f + g) = subdegree g"
- using assms
- by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
+ using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
lemma subdegree_diff_eq1:
assumes "f \<noteq> 0"
and "subdegree f < subdegree (g :: 'a :: group_add fps)"
shows "subdegree (f - g) = subdegree f"
- using assms
- by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
+ using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
lemma subdegree_diff_eq1_cancel:
assumes "f \<noteq> 0"
and "subdegree f < subdegree (g :: 'a :: cancel_comm_monoid_add fps)"
shows "subdegree (f - g) = subdegree f"
- using assms
- by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
+ using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
lemma subdegree_diff_eq2:
assumes "g \<noteq> 0"
and "subdegree g < subdegree (f :: 'a :: group_add fps)"
shows "subdegree (f - g) = subdegree g"
- using assms
- by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
+ using assms by (auto intro: subdegreeI simp: nth_less_subdegree_zero)
lemma subdegree_diff_ge [simp]:
assumes "f \<noteq> (g :: 'a :: group_add fps)"
shows "subdegree (f - g) \<ge> min (subdegree f) (subdegree g)"
proof-
- from assms have "f = - (- g) \<Longrightarrow> False" using expand_fps_eq by fastforce
- hence "f \<noteq> - (- g)" by fast
+ have "f \<noteq> - (- g)"
+ using assms expand_fps_eq by fastforce
moreover have "f + - g = f - g" by (simp add: fps_ext)
ultimately show ?thesis
using subdegree_add_ge[of f "-g"] by simp
@@ -334,8 +285,7 @@
fixes f g :: "'a :: comm_monoid_diff fps"
assumes "f - g \<noteq> 0"
shows "subdegree (f - g) \<ge> subdegree f"
- using assms
- by (auto intro: subdegree_geI simp: nth_less_subdegree_zero)
+ using assms by (auto intro: subdegree_geI simp: nth_less_subdegree_zero)
lemma nth_subdegree_mult_left [simp]:
fixes f g :: "('a :: {mult_zero,comm_monoid_add}) fps"
--- a/src/HOL/Computational_Algebra/Fundamental_Theorem_Algebra.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Computational_Algebra/Fundamental_Theorem_Algebra.thy Mon Feb 20 13:59:42 2023 +0100
@@ -39,11 +39,9 @@
using norm_ge_zero[of z] by arith
have "norm (poly (pCons c cs) z) \<le> norm c + norm (z * poly cs z)"
using norm_triangle_ineq[of c "z* poly cs z"] by simp
- also have "\<dots> \<le> norm c + r * m"
+ also have "\<dots> \<le> ?k"
using mult_mono[OF H th rp norm_ge_zero[of "poly cs z"]]
by (simp add: norm_mult)
- also have "\<dots> \<le> ?k"
- by simp
finally show ?thesis .
qed
with kp show ?case by blast
@@ -63,37 +61,40 @@
smult h (offset_poly p h) + pCons a (offset_poly p h)"
by (cases "p = 0 \<and> a = 0") (auto simp add: offset_poly_def)
-lemma offset_poly_single: "offset_poly [:a:] h = [:a:]"
+lemma offset_poly_single [simp]: "offset_poly [:a:] h = [:a:]"
by (simp add: offset_poly_pCons offset_poly_0)
lemma poly_offset_poly: "poly (offset_poly p h) x = poly p (h + x)"
- apply (induct p)
- apply (simp add: offset_poly_0)
- apply (simp add: offset_poly_pCons algebra_simps)
- done
+ by (induct p) (auto simp add: offset_poly_0 offset_poly_pCons algebra_simps)
lemma offset_poly_eq_0_lemma: "smult c p + pCons a p = 0 \<Longrightarrow> p = 0"
by (induct p arbitrary: a) (simp, force)
-lemma offset_poly_eq_0_iff: "offset_poly p h = 0 \<longleftrightarrow> p = 0"
- apply (safe intro!: offset_poly_0)
- apply (induct p)
- apply simp
- apply (simp add: offset_poly_pCons)
- apply (frule offset_poly_eq_0_lemma, simp)
- done
+lemma offset_poly_eq_0_iff [simp]: "offset_poly p h = 0 \<longleftrightarrow> p = 0"
+proof
+ show "offset_poly p h = 0 \<Longrightarrow> p = 0"
+ proof(induction p)
+ case 0
+ then show ?case by blast
+ next
+ case (pCons a p)
+ then show ?case
+ by (metis offset_poly_eq_0_lemma offset_poly_pCons offset_poly_single)
+ qed
+qed (simp add: offset_poly_0)
-lemma degree_offset_poly: "degree (offset_poly p h) = degree p"
- apply (induct p)
- apply (simp add: offset_poly_0)
- apply (case_tac "p = 0")
- apply (simp add: offset_poly_0 offset_poly_pCons)
- apply (simp add: offset_poly_pCons)
- apply (subst degree_add_eq_right)
- apply (rule le_less_trans [OF degree_smult_le])
- apply (simp add: offset_poly_eq_0_iff)
- apply (simp add: offset_poly_eq_0_iff)
- done
+lemma degree_offset_poly [simp]: "degree (offset_poly p h) = degree p"
+proof(induction p)
+ case 0
+ then show ?case
+ by (simp add: offset_poly_0)
+next
+ case (pCons a p)
+ have "p \<noteq> 0 \<Longrightarrow> degree (offset_poly (pCons a p) h) = Suc (degree p)"
+ by (metis degree_add_eq_right degree_pCons_eq degree_smult_le le_imp_less_Suc offset_poly_eq_0_iff offset_poly_pCons pCons.IH)
+ then show ?case
+ by simp
+qed
definition "psize p = (if p = 0 then 0 else Suc (degree p))"
@@ -103,13 +104,7 @@
lemma poly_offset:
fixes p :: "'a::comm_ring_1 poly"
shows "\<exists>q. psize q = psize p \<and> (\<forall>x. poly q x = poly p (a + x))"
-proof (intro exI conjI)
- show "psize (offset_poly p a) = psize p"
- unfolding psize_def
- by (simp add: offset_poly_eq_0_iff degree_offset_poly)
- show "\<forall>x. poly (offset_poly p a) x = poly p (a + x)"
- by (simp add: poly_offset_poly)
-qed
+ by (metis degree_offset_poly offset_poly_eq_0_iff poly_offset_poly psize_def)
text \<open>An alternative useful formulation of completeness of the reals\<close>
lemma real_sup_exists:
@@ -136,19 +131,17 @@
by (simp add: cmod_def)
have False if "cmod (z + 1) \<ge> 1" "cmod (z - 1) \<ge> 1" "cmod (z + \<i>) \<ge> 1" "cmod (z - \<i>) \<ge> 1"
proof -
- from that z xy have "2 * x \<le> 1" "2 * x \<ge> -1" "2 * y \<le> 1" "2 * y \<ge> -1"
+ from that z xy have *: "2 * x \<le> 1" "2 * x \<ge> -1" "2 * y \<le> 1" "2 * y \<ge> -1"
by (simp_all add: cmod_def power2_eq_square algebra_simps)
then have "\<bar>2 * x\<bar> \<le> 1" "\<bar>2 * y\<bar> \<le> 1"
by simp_all
then have "\<bar>2 * x\<bar>\<^sup>2 \<le> 1\<^sup>2" "\<bar>2 * y\<bar>\<^sup>2 \<le> 1\<^sup>2"
- by - (rule power_mono, simp, simp)+
- then have th0: "4 * x\<^sup>2 \<le> 1" "4 * y\<^sup>2 \<le> 1"
- by (simp_all add: power_mult_distrib)
- from add_mono[OF th0] xy show ?thesis
- by simp
+ by (metis abs_square_le_1 one_power2 power2_abs)+
+ with xy * show ?thesis
+ by (smt (verit, best) four_x_squared square_le_1)
qed
then show ?thesis
- unfolding linorder_not_le[symmetric] by blast
+ by force
qed
text \<open>Hence we can always reduce modulus of \<open>1 + b z^n\<close> if nonzero\<close>
@@ -164,14 +157,10 @@
let ?P = "\<lambda>z n. cmod (1 + b * z ^ n) < 1"
show "\<exists>z. ?P z n"
proof cases
- assume "even n"
- then have "\<exists>m. n = 2 * m"
- by presburger
- then obtain m where m: "n = 2 * m"
- by blast
- from n m have "m \<noteq> 0" "m < n"
- by presburger+
- with IH[rule_format, of m] obtain z where z: "?P z m"
+ assume "even n"
+ then obtain m where m: "n = 2 * m" and "m \<noteq> 0" "m < n"
+ using n by auto
+ with IH obtain z where z: "?P z m"
by blast
from z have "?P (csqrt z) n"
by (simp add: m power_mult)
@@ -182,49 +171,66 @@
by presburger+
then obtain m where m: "n = Suc (2 * m)"
by blast
- have th0: "cmod (complex_of_real (cmod b) / b) = 1"
+ have 0: "cmod (complex_of_real (cmod b) / b) = 1"
using b by (simp add: norm_divide)
- from unimodular_reduce_norm[OF th0] \<open>odd n\<close>
have "\<exists>v. cmod (complex_of_real (cmod b) / b + v^n) < 1"
- apply (cases "cmod (complex_of_real (cmod b) / b + 1) < 1")
- apply (rule_tac x="1" in exI)
- apply simp
- apply (cases "cmod (complex_of_real (cmod b) / b - 1) < 1")
- apply (rule_tac x="-1" in exI)
- apply simp
- apply (cases "cmod (complex_of_real (cmod b) / b + \<i>) < 1")
- apply (cases "even m")
- apply (rule_tac x="\<i>" in exI)
- apply (simp add: m power_mult)
- apply (rule_tac x="- \<i>" in exI)
- apply (simp add: m power_mult)
- apply (cases "even m")
- apply (rule_tac x="- \<i>" in exI)
- apply (simp add: m power_mult)
- apply (auto simp add: m power_mult)
- apply (rule_tac x="\<i>" in exI)
- apply (auto simp add: m power_mult)
- done
+ proof (cases "cmod (complex_of_real (cmod b) / b + 1) < 1")
+ case True
+ then show ?thesis
+ by (metis power_one)
+ next
+ case F1: False
+ show ?thesis
+ proof (cases "cmod (complex_of_real (cmod b) / b - 1) < 1")
+ case True
+ with \<open>odd n\<close> show ?thesis
+ by (metis add_uminus_conv_diff neg_one_odd_power)
+ next
+ case F2: False
+ show ?thesis
+ proof (cases "cmod (complex_of_real (cmod b) / b + \<i>) < 1")
+ case T1: True
+ show ?thesis
+ proof (cases "even m")
+ case True
+ with T1 show ?thesis
+ by (rule_tac x="\<i>" in exI) (simp add: m power_mult)
+ next
+ case False
+ with T1 show ?thesis
+ by (rule_tac x="- \<i>" in exI) (simp add: m power_mult)
+ qed
+ next
+ case False
+ with F1 F2 m unimodular_reduce_norm[OF 0] \<open>odd n\<close> show ?thesis
+ apply (cases "even m")
+ apply (rule_tac x="- \<i>" in exI)
+ apply (simp add: power_mult)
+ apply (rule_tac x="\<i>" in exI)
+ apply (auto simp add: m power_mult)
+ done
+ qed
+ qed
+ qed
then obtain v where v: "cmod (complex_of_real (cmod b) / b + v^n) < 1"
by blast
let ?w = "v / complex_of_real (root n (cmod b))"
from odd_real_root_pow[OF \<open>odd n\<close>, of "cmod b"]
- have th1: "?w ^ n = v^n / complex_of_real (cmod b)"
+ have 1: "?w ^ n = v^n / complex_of_real (cmod b)"
by (simp add: power_divide of_real_power[symmetric])
- have th2:"cmod (complex_of_real (cmod b) / b) = 1"
+ have 2:"cmod (complex_of_real (cmod b) / b) = 1"
using b by (simp add: norm_divide)
- then have th3: "cmod (complex_of_real (cmod b) / b) \<ge> 0"
+ then have 3: "cmod (complex_of_real (cmod b) / b) \<ge> 0"
by simp
- have th4: "cmod (complex_of_real (cmod b) / b) *
+ have 4: "cmod (complex_of_real (cmod b) / b) *
cmod (1 + b * (v ^ n / complex_of_real (cmod b))) <
cmod (complex_of_real (cmod b) / b) * 1"
apply (simp only: norm_mult[symmetric] distrib_left)
using b v
- apply (simp add: th2)
+ apply (simp add: 2)
done
- from mult_left_less_imp_less[OF th4 th3]
- have "?P ?w n" unfolding th1 .
- then show ?thesis ..
+ show ?thesis
+ by (metis 1 mult_left_less_imp_less[OF 4 3])
qed
qed
@@ -245,44 +251,21 @@
obtain g where g: "strict_mono g" "monoseq (\<lambda>n. Im (s (f (g n))))"
unfolding o_def by blast
let ?h = "f \<circ> g"
- from r[rule_format, of 0] have rp: "r \<ge> 0"
- using norm_ge_zero[of "s 0"] by arith
- have th: "\<forall>n. r + 1 \<ge> \<bar>Re (s n)\<bar>"
- proof
- fix n
- from abs_Re_le_cmod[of "s n"] r[rule_format, of n]
- show "\<bar>Re (s n)\<bar> \<le> r + 1" by arith
- qed
- have conv1: "convergent (\<lambda>n. Re (s (f n)))"
- apply (rule Bseq_monoseq_convergent)
- apply (simp add: Bseq_def)
- apply (metis gt_ex le_less_linear less_trans order.trans th)
- apply (rule f(2))
- done
- have th: "\<forall>n. r + 1 \<ge> \<bar>Im (s n)\<bar>"
- proof
- fix n
- from abs_Im_le_cmod[of "s n"] r[rule_format, of n]
- show "\<bar>Im (s n)\<bar> \<le> r + 1"
- by arith
- qed
+ have "r \<ge> 0"
+ by (meson norm_ge_zero order_trans r)
+ have "\<forall>n. r + 1 \<ge> \<bar>Re (s n)\<bar>"
+ by (smt (verit, ccfv_threshold) abs_Re_le_cmod r)
+ then have conv1: "convergent (\<lambda>n. Re (s (f n)))"
+ by (metis Bseq_monoseq_convergent f(2) BseqI' real_norm_def)
+ have "\<forall>n. r + 1 \<ge> \<bar>Im (s n)\<bar>"
+ by (smt (verit) abs_Im_le_cmod r)
+ then have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
+ by (metis Bseq_monoseq_convergent g(2) BseqI' real_norm_def)
- have conv2: "convergent (\<lambda>n. Im (s (f (g n))))"
- apply (rule Bseq_monoseq_convergent)
- apply (simp add: Bseq_def)
- apply (metis gt_ex le_less_linear less_trans order.trans th)
- apply (rule g(2))
- done
-
- from conv1[unfolded convergent_def] obtain x where "LIMSEQ (\<lambda>n. Re (s (f n))) x"
- by blast
- then have x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Re (s (f n)) - x\<bar> < r"
- unfolding LIMSEQ_iff real_norm_def .
-
- from conv2[unfolded convergent_def] obtain y where "LIMSEQ (\<lambda>n. Im (s (f (g n)))) y"
- by blast
- then have y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Im (s (f (g n))) - y\<bar> < r"
- unfolding LIMSEQ_iff real_norm_def .
+ obtain x where x: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Re (s (f n)) - x\<bar> < r"
+ using conv1[unfolded convergent_def] LIMSEQ_iff real_norm_def by metis
+ obtain y where y: "\<forall>r>0. \<exists>n0. \<forall>n\<ge>n0. \<bar>Im (s (f (g n))) - y\<bar> < r"
+ using conv2[unfolded convergent_def] LIMSEQ_iff real_norm_def by metis
let ?w = "Complex x y"
from f(1) g(1) have hs: "strict_mono ?h"
unfolding strict_mono_def by auto
@@ -290,7 +273,7 @@
proof -
from that have e2: "e/2 > 0"
by simp
- from x[rule_format, OF e2] y[rule_format, OF e2]
+ from x y e2
obtain N1 N2 where N1: "\<forall>n\<ge>N1. \<bar>Re (s (f n)) - x\<bar> < e / 2"
and N2: "\<forall>n\<ge>N2. \<bar>Im (s (f (g n))) - y\<bar> < e / 2"
by blast
@@ -298,9 +281,8 @@
proof -
from that have nN1: "g n \<ge> N1" and nN2: "n \<ge> N2"
using seq_suble[OF g(1), of n] by arith+
- from add_strict_mono[OF N1[rule_format, OF nN1] N2[rule_format, OF nN2]]
show ?thesis
- using metric_bound_lemma[of "s (f (g n))" ?w] by simp
+ using metric_bound_lemma[of "s (f (g n))" ?w] N1 N2 nN1 nN2 by fastforce
qed
then show ?thesis by blast
qed
@@ -315,12 +297,7 @@
shows "\<exists>d >0. \<forall>w. 0 < norm (w - z) \<and> norm (w - z) < d \<longrightarrow> norm (poly p w - poly p z) < e"
proof -
obtain q where q: "degree q = degree p" "poly q x = poly p (z + x)" for x
- proof
- show "degree (offset_poly p z) = degree p"
- by (rule degree_offset_poly)
- show "\<And>x. poly (offset_poly p z) x = poly p (z + x)"
- by (rule poly_offset_poly)
- qed
+ using degree_offset_poly poly_offset_poly by blast
have th: "\<And>w. poly q (w - z) = poly p w"
using q(2)[of "w - z" for w] by simp
show ?thesis unfolding th[symmetric]
@@ -333,7 +310,7 @@
from poly_bound_exists[of 1 "cs"]
obtain m where m: "m > 0" "norm z \<le> 1 \<Longrightarrow> norm (poly cs z) \<le> m" for z
by blast
- from ep m(1) have em0: "e/m > 0"
+ with ep have em0: "e/m > 0"
by (simp add: field_simps)
have one0: "1 > (0::real)"
by arith
@@ -352,9 +329,8 @@
by (simp add: field_simps)
from H have th: "norm (w - z) \<le> d"
by simp
- from mult_mono[OF th m(2)[OF d1(1)] d1(2) norm_ge_zero] dme
show "norm (w - z) * norm (poly cs (w - z)) < e"
- by simp
+ by (smt (verit, del_insts) d1(1) dme m(2) mult_mono' norm_ge_zero th)
qed
qed
qed
@@ -370,37 +346,26 @@
by (metis norm_ge_zero order.trans)
next
case True
- then have "cmod 0 \<le> r \<and> cmod (poly p 0) = - (- cmod (poly p 0))"
- by simp
then have mth1: "\<exists>x z. cmod z \<le> r \<and> cmod (poly p z) = - x"
- by blast
- have False if "cmod z \<le> r" "cmod (poly p z) = - x" "\<not> x < 1" for x z
- proof -
- from that have "- x < 0 "
- by arith
- with that(2) norm_ge_zero[of "poly p z"] show ?thesis
- by simp
- qed
- then have mth2: "\<exists>z. \<forall>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<longrightarrow> x < z"
- by blast
- from real_sup_exists[OF mth1 mth2] obtain s where
- s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow> y < s"
- by blast
+ by (metis add.inverse_inverse norm_zero)
+ have False if "cmod (poly p z) = - x" "\<not> x < 1" for x z
+ by (smt (verit, del_insts) norm_ge_zero that)
+ with real_sup_exists[OF mth1]
+ obtain s where s: "\<forall>y. (\<exists>x. (\<exists>z. cmod z \<le> r \<and> cmod (poly p z) = - x) \<and> y < x) \<longleftrightarrow> y < s"
+ by auto
+
let ?m = "- s"
- have s1[unfolded minus_minus]:
- "(\<exists>z x. cmod z \<le> r \<and> - (- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" for y
- using s[rule_format, of "-y"]
- unfolding minus_less_iff[of y] equation_minus_iff by blast
+ have s1: "(\<exists>z x. cmod z \<le> r \<and> - (- cmod (poly p z)) < y) \<longleftrightarrow> ?m < y" for y
+ by (metis add.inverse_inverse minus_less_iff s)
from s1[of ?m] have s1m: "\<And>z x. cmod z \<le> r \<Longrightarrow> cmod (poly p z) \<ge> ?m"
by auto
have "\<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" for n
using s1[rule_format, of "?m + 1/real (Suc n)"] by simp
- then have th: "\<forall>n. \<exists>z. cmod z \<le> r \<and> cmod (poly p z) < - s + 1 / real (Suc n)" ..
- from choice[OF th] obtain g where
- g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m + 1 /real(Suc n)"
- by blast
+ then obtain g where g: "\<forall>n. cmod (g n) \<le> r" "\<forall>n. cmod (poly p (g n)) <?m + 1 /real(Suc n)"
+ by metis
from Bolzano_Weierstrass_complex_disc[OF g(1)]
- obtain f z where fz: "strict_mono (f :: nat \<Rightarrow> nat)" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
+ obtain f::"nat \<Rightarrow> nat" and z
+ where fz: "strict_mono f" "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. cmod (g (f n) - z) < e"
by blast
{
fix w
@@ -413,54 +378,31 @@
from poly_cont[OF e2, of z p] obtain d where
d: "d > 0" "\<forall>w. 0<cmod (w - z)\<and> cmod(w - z) < d \<longrightarrow> cmod(poly p w - poly p z) < ?e/2"
by blast
- have th1: "cmod(poly p w - poly p z) < ?e / 2" if w: "cmod (w - z) < d" for w
+ have 1: "cmod(poly p w - poly p z) < ?e / 2" if w: "cmod (w - z) < d" for w
using d(2)[rule_format, of w] w e by (cases "w = z") simp_all
from fz(2) d(1) obtain N1 where N1: "\<forall>n\<ge>N1. cmod (g (f n) - z) < d"
by blast
- from reals_Archimedean2[of "2/?e"] obtain N2 :: nat where N2: "2/?e < real N2"
+ from reals_Archimedean2 obtain N2 :: nat where N2: "2/?e < real N2"
by blast
- have th2: "cmod (poly p (g (f (N1 + N2))) - poly p z) < ?e/2"
- using N1[rule_format, of "N1 + N2"] th1 by simp
- have th0: "a < e2 \<Longrightarrow> \<bar>b - m\<bar> < e2 \<Longrightarrow> 2 * e2 \<le> \<bar>b - m\<bar> + a \<Longrightarrow> False"
+ have 2: "cmod (poly p (g (f (N1 + N2))) - poly p z) < ?e/2"
+ using N1 1 by auto
+ have 0: "a < e2 \<Longrightarrow> \<bar>b - m\<bar> < e2 \<Longrightarrow> 2 * e2 \<le> \<bar>b - m\<bar> + a \<Longrightarrow> False"
for a b e2 m :: real
by arith
- have ath: "m \<le> x \<Longrightarrow> x < m + e \<Longrightarrow> \<bar>x - m\<bar> < e" for m x e :: real
- by arith
- from s1m[OF g(1)[rule_format]] have th31: "?m \<le> cmod(poly p (g (f (N1 + N2))))" .
from seq_suble[OF fz(1), of "N1 + N2"]
- have th00: "real (Suc (N1 + N2)) \<le> real (Suc (f (N1 + N2)))"
- by simp
- have th000: "0 \<le> (1::real)" "(1::real) \<le> 1" "real (Suc (N1 + N2)) > 0"
- using N2 by auto
- from frac_le[OF th000 th00]
- have th00: "?m + 1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))"
- by simp
- from g(2)[rule_format, of "f (N1 + N2)"]
- have th01:"cmod (poly p (g (f (N1 + N2)))) < - s + 1 / real (Suc (f (N1 + N2)))" .
- from order_less_le_trans[OF th01 th00]
- have th32: "cmod (poly p (g (f (N1 + N2)))) < ?m + (1/ real(Suc (N1 + N2)))" .
- from N2 have "2/?e < real (Suc (N1 + N2))"
- by arith
- with e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
+ have 00: "?m + 1 / real (Suc (f (N1 + N2))) \<le> ?m + 1 / real (Suc (N1 + N2))"
+ by (simp add: frac_le)
+ from N2 e2 less_imp_inverse_less[of "2/?e" "real (Suc (N1 + N2))"]
have "?e/2 > 1/ real (Suc (N1 + N2))"
by (simp add: inverse_eq_divide)
- with ath[OF th31 th32] have thc1: "\<bar>cmod (poly p (g (f (N1 + N2)))) - ?m\<bar> < ?e/2"
- by arith
- have ath2: "\<bar>a - b\<bar> \<le> c \<Longrightarrow> \<bar>b - m\<bar> \<le> \<bar>a - m\<bar> + c" for a b c m :: real
- by arith
- have th22: "\<bar>cmod (poly p (g (f (N1 + N2)))) - cmod (poly p z)\<bar> \<le>
- cmod (poly p (g (f (N1 + N2))) - poly p z)"
- by (simp add: norm_triangle_ineq3)
- from ath2[OF th22, of ?m]
- have thc2: "2 * (?e/2) \<le>
- \<bar>cmod(poly p (g (f (N1 + N2)))) - ?m\<bar> + cmod (poly p (g (f (N1 + N2))) - poly p z)"
- by simp
- from th0[OF th2 thc1 thc2] have False .
+ with order_less_le_trans[OF _ 00]
+ have 1: "\<bar>cmod (poly p (g (f (N1 + N2)))) - ?m\<bar> < ?e/2"
+ using g s1m by (smt (verit))
+ with 0[OF 2] have False
+ by (smt (verit) field_sum_of_halves norm_triangle_ineq3)
}
then have "?e = 0"
by auto
- then have "cmod (poly p z) = ?m"
- by simp
with s1m[OF wr] have "cmod (poly p z) \<le> cmod (poly p w)"
by simp
}
@@ -488,39 +430,24 @@
let ?r = "1 + \<bar>r\<bar>"
have "d \<le> norm (poly (pCons a (pCons c cs)) z)" if "1 + \<bar>r\<bar> \<le> norm z" for z
proof -
- have r0: "r \<le> norm z"
- using that by arith
- from r[rule_format, OF r0] have th0: "d + norm a \<le> 1 * norm(poly (pCons c cs) z)"
- by arith
- from that have z1: "norm z \<ge> 1"
- by arith
- from order_trans[OF th0 mult_right_mono[OF z1 norm_ge_zero[of "poly (pCons c cs) z"]]]
- have th1: "d \<le> norm(z * poly (pCons c cs) z) - norm a"
- unfolding norm_mult by (simp add: algebra_simps)
- from norm_diff_ineq[of "z * poly (pCons c cs) z" a]
- have th2: "norm (z * poly (pCons c cs) z) - norm a \<le> norm (poly (pCons a (pCons c cs)) z)"
- by (simp add: algebra_simps)
- from th1 th2 show ?thesis
- by arith
+ have "d \<le> norm(z * poly (pCons c cs) z) - norm a"
+ by (smt (verit, best) norm_ge_zero mult_less_cancel_right2 norm_mult r that)
+ with norm_diff_ineq add.commute
+ show ?thesis
+ by (metis order.trans poly_pCons)
qed
then show ?thesis by blast
next
case True
- with pCons.prems have c0: "c \<noteq> 0"
- by simp
have "d \<le> norm (poly (pCons a (pCons c cs)) z)"
- if h: "(\<bar>d\<bar> + norm a) / norm c \<le> norm z" for z :: 'a
+ if "(\<bar>d\<bar> + norm a) / norm c \<le> norm z" for z :: 'a
proof -
- from c0 have "norm c > 0"
- by simp
- from h c0 have th0: "\<bar>d\<bar> + norm a \<le> norm (z * c)"
- by (simp add: field_simps norm_mult)
- have ath: "\<And>mzh mazh ma. mzh \<le> mazh + ma \<Longrightarrow> \<bar>d\<bar> + ma \<le> mzh \<Longrightarrow> d \<le> mazh"
- by arith
- from norm_diff_ineq[of "z * c" a] have th1: "norm (z * c) \<le> norm (a + z * c) + norm a"
- by (simp add: algebra_simps)
- from ath[OF th1 th0] show ?thesis
- using True by simp
+ have "\<bar>d\<bar> + norm a \<le> norm (z * c)"
+ by (metis that True norm_mult pCons.hyps(1) pos_divide_le_eq zero_less_norm_iff)
+ also have "\<dots> \<le> norm (a + z * c) + norm a"
+ by (simp add: add.commute norm_add_leD)
+ finally show ?thesis
+ using True by auto
qed
then show ?thesis by blast
qed
@@ -540,21 +467,9 @@
obtain r where r: "cmod (poly (pCons c cs) 0) \<le> cmod (poly (pCons c cs) z)"
if "r \<le> cmod z" for z
by blast
- have ath: "\<And>z r. r \<le> cmod z \<or> cmod z \<le> \<bar>r\<bar>"
- by arith
- from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"]
- obtain v where v: "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) w)"
- if "cmod w \<le> \<bar>r\<bar>" for w
- by blast
- have "cmod (poly (pCons c cs) v) \<le> cmod (poly (pCons c cs) z)" if z: "r \<le> cmod z" for z
- using v[of 0] r[OF z] by simp
- with v ath[of r] show ?thesis
- by blast
- next
- case True
- with pCons.hyps show ?thesis
- by simp
- qed
+ from poly_minimum_modulus_disc[of "\<bar>r\<bar>" "pCons c cs"] show ?thesis
+ by (smt (verit, del_insts) order.trans linorder_linear r)
+ qed (use pCons.hyps in auto)
qed
text \<open>Constant function (non-syntactic characterization).\<close>
@@ -589,21 +504,15 @@
apply (rule_tac x="q" in exI)
apply auto
done
- next
- case False
- show ?thesis
- apply (rule exI[where x=0])
- apply (rule exI[where x=c])
- apply (auto simp: False)
- done
- qed
+ qed force
qed
lemma poly_decompose:
+ fixes p :: "'a::idom poly"
assumes nc: "\<not> constant (poly p)"
- shows "\<exists>k a q. a \<noteq> (0::'a::idom) \<and> k \<noteq> 0 \<and>
+ shows "\<exists>k a q. a \<noteq> 0 \<and> k \<noteq> 0 \<and>
psize q + k + 1 = psize p \<and>
- (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
+ (\<forall>z. poly p z = poly p 0 + z^k * poly (pCons a q) z)"
using nc
proof (induct p)
case 0
@@ -612,22 +521,15 @@
next
case (pCons c cs)
have "\<not> (\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0)"
- proof
- assume "\<forall>z. z \<noteq> 0 \<longrightarrow> poly cs z = 0"
- then have "poly (pCons c cs) x = poly (pCons c cs) y" for x y
- by (cases "x = 0") auto
- with pCons.prems show False
- by (auto simp add: constant_def)
- qed
+ by (smt (verit) constant_def mult_eq_0_iff pCons.prems poly_pCons)
from poly_decompose_lemma[OF this]
- show ?case
- apply clarsimp
- apply (rule_tac x="k+1" in exI)
- apply (rule_tac x="a" in exI)
- apply simp
- apply (rule_tac x="q" in exI)
- apply (auto simp add: psize_def split: if_splits)
- done
+ obtain k a q where *: "a \<noteq> 0 \<and>
+ Suc (psize q + k) = psize cs \<and> (\<forall>z. poly cs z = z ^ k * poly (pCons a q) z)"
+ by blast
+ then have "psize q + k + 2 = psize (pCons c cs)"
+ by (auto simp add: psize_def split: if_splits)
+ then show ?case
+ using "*" by force
qed
text \<open>Fundamental theorem of algebra\<close>
@@ -651,27 +553,10 @@
then show ?thesis by blast
next
case False
- from poly_offset[of p c] obtain q where q: "psize q = psize p" "\<forall>x. poly q x = ?p (c + x)"
- by blast
- have False if h: "constant (poly q)"
- proof -
- from q(2) have th: "\<forall>x. poly q (x - c) = ?p x"
- by auto
- have "?p x = ?p y" for x y
- proof -
- from th have "?p x = poly q (x - c)"
- by auto
- also have "\<dots> = poly q (y - c)"
- using h unfolding constant_def by blast
- also have "\<dots> = ?p y"
- using th by auto
- finally show ?thesis .
- qed
- with less(2) show ?thesis
- unfolding constant_def by blast
- qed
+ obtain q where q: "psize q = psize p" "\<forall>x. poly q x = ?p (c + x)"
+ using poly_offset[of p c] by blast
then have qnc: "\<not> constant (poly q)"
- by blast
+ by (metis (no_types, opaque_lifting) add.commute constant_def diff_add_cancel less.prems)
from q(2) have pqc0: "?p c = poly q 0"
by simp
from c pqc0 have cq0: "\<forall>w. cmod (poly q 0) \<le> cmod (?p w)"
@@ -683,36 +568,13 @@
by simp
let ?r = "smult (inverse ?a0) q"
have lgqr: "psize q = psize ?r"
- using a00
- unfolding psize_def degree_def
- by (simp add: poly_eq_iff)
- have False if h: "\<And>x y. poly ?r x = poly ?r y"
- proof -
- have "poly q x = poly q y" for x y
- proof -
- from qr[rule_format, of x] have "poly q x = poly ?r x * ?a0"
- by auto
- also have "\<dots> = poly ?r y * ?a0"
- using h by simp
- also have "\<dots> = poly q y"
- using qr[rule_format, of y] by simp
- finally show ?thesis .
- qed
- with qnc show ?thesis
- unfolding constant_def by blast
- qed
- then have rnc: "\<not> constant (poly ?r)"
- unfolding constant_def by blast
- from qr[rule_format, of 0] a00 have r01: "poly ?r 0 = 1"
- by auto
+ by (simp add: a00 psize_def)
+ have rnc: "\<not> constant (poly ?r)"
+ using constant_def qnc qr by fastforce
+ have r01: "poly ?r 0 = 1"
+ by (simp add: a00)
have mrmq_eq: "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w) < cmod ?a0" for w
- proof -
- have "cmod (poly ?r w) < 1 \<longleftrightarrow> cmod (poly q w / ?a0) < 1"
- using qr[rule_format, of w] a00 by (simp add: divide_inverse ac_simps)
- also have "\<dots> \<longleftrightarrow> cmod (poly q w) < cmod ?a0"
- using a00 unfolding norm_divide by (simp add: field_simps)
- finally show ?thesis .
- qed
+ by (smt (verit, del_insts) a00 mult_less_cancel_right2 norm_mult qr zero_less_norm_iff)
from poly_decompose[OF rnc] obtain k a s where
kas: "a \<noteq> 0" "k \<noteq> 0" "psize s + k + 1 = psize ?r"
"\<forall>z. poly ?r z = poly ?r 0 + z^k* poly (pCons a s) z" by blast
@@ -721,39 +583,28 @@
case True
with kas(3) lgqr[symmetric] q(1) have s0: "s = 0"
by auto
- have hth[symmetric]: "cmod (poly ?r w) = cmod (1 + a * w ^ k)" for w
- using kas(4)[rule_format, of w] s0 r01 by (simp add: algebra_simps)
- from reduce_poly_simple[OF kas(1,2)] show ?thesis
- unfolding hth by blast
+ with reduce_poly_simple kas show ?thesis
+ by (metis mult.commute mult.right_neutral poly_1 poly_smult r01 smult_one)
next
case False note kn = this
from kn kas(3) q(1) lgqr have k1n: "k + 1 < psize p"
by simp
- have th01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
+ have 01: "\<not> constant (poly (pCons 1 (monom a (k - 1))))"
unfolding constant_def poly_pCons poly_monom
- using kas(1)
- apply simp
- apply (rule exI[where x=0])
- apply (rule exI[where x=1])
- apply simp
- done
- from kas(1) kas(2) have th02: "k + 1 = psize (pCons 1 (monom a (k - 1)))"
+ by (metis add_cancel_left_right kas(1) mult.commute mult_cancel_right2 power_one)
+ from kas(1) kas(2) have 02: "k + 1 = psize (pCons 1 (monom a (k - 1)))"
by (simp add: psize_def degree_monom_eq)
- from less(1) [OF k1n [simplified th02] th01]
+ from less(1) [OF k1n [simplified 02] 01]
obtain w where w: "1 + w^k * a = 0"
- unfolding poly_pCons poly_monom
- using kas(2) by (cases k) (auto simp add: algebra_simps)
+ by (metis kas(2) mult.commute mult.left_commute poly_monom poly_pCons power_eq_if)
from poly_bound_exists[of "cmod w" s] obtain m where
m: "m > 0" "\<forall>z. cmod z \<le> cmod w \<longrightarrow> cmod (poly s z) \<le> m" by blast
- have w0: "w \<noteq> 0"
+ have "w \<noteq> 0"
using kas(2) w by (auto simp add: power_0_left)
- from w have "(1 + w ^ k * a) - 1 = 0 - 1"
- by simp
- then have wm1: "w^k * a = - 1"
- by simp
+ from w have wm1: "w^k * a = - 1"
+ by (simp add: add_eq_0_iff)
have inv0: "0 < inverse (cmod w ^ (k + 1) * m)"
- using norm_ge_zero[of w] w0 m(1)
- by (simp add: inverse_eq_divide zero_less_mult_iff)
+ by (simp add: \<open>w \<noteq> 0\<close> m(1))
with field_lbound_gt_zero[OF zero_less_one] obtain t where
t: "t > 0" "t < 1" "t < inverse (cmod w ^ (k + 1) * m)" by blast
let ?ct = "complex_of_real t"
@@ -766,42 +617,29 @@
cmod (complex_of_real (1 - t^k) + ?w^k * ?w * poly s ?w)"
by metis
with norm_triangle_ineq[of "complex_of_real (1 - t^k)" "?w^k * ?w * poly s ?w"]
- have th11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)"
+ have 11: "cmod (1 + ?w^k * (a + ?w * poly s ?w)) \<le> \<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w)"
unfolding norm_of_real by simp
have ath: "\<And>x t::real. 0 \<le> x \<Longrightarrow> x < t \<Longrightarrow> t \<le> 1 \<Longrightarrow> \<bar>1 - t\<bar> + x < 1"
by arith
- have "t * cmod w \<le> 1 * cmod w"
- apply (rule mult_mono)
- using t(1,2)
- apply auto
- done
- then have tw: "cmod ?w \<le> cmod w"
- using t(1) by (simp add: norm_mult)
- from t inv0 have "t * (cmod w ^ (k + 1) * m) < 1"
- by (simp add: field_simps)
- with zero_less_power[OF t(1), of k] have th30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
+ have tw: "cmod ?w \<le> cmod w"
+ by (smt (verit) mult_le_cancel_right2 norm_ge_zero norm_mult norm_of_real t)
+ have "t * (cmod w ^ (k + 1) * m) < 1"
+ by (smt (verit, best) inv0 inverse_positive_iff_positive left_inverse mult_strict_right_mono t(3))
+ with zero_less_power[OF t(1), of k] have 30: "t^k * (t* (cmod w ^ (k + 1) * m)) < t^k * 1"
by simp
have "cmod (?w^k * ?w * poly s ?w) = t^k * (t* (cmod w ^ (k + 1) * cmod (poly s ?w)))"
- using w0 t(1)
- by (simp add: algebra_simps power_mult_distrib norm_power norm_mult)
- then have "cmod (?w^k * ?w * poly s ?w) \<le> t^k * (t* (cmod w ^ (k + 1) * m))"
- using t(1,2) m(2)[rule_format, OF tw] w0
- by auto
- with th30 have th120: "cmod (?w^k * ?w * poly s ?w) < t^k"
- by simp
- from power_strict_mono[OF t(2), of k] t(1) kas(2) have th121: "t^k \<le> 1"
+ using \<open>w \<noteq> 0\<close> t(1) by (simp add: algebra_simps norm_power norm_mult)
+ with 30 have 120: "cmod (?w^k * ?w * poly s ?w) < t^k"
+ by (smt (verit, ccfv_SIG) m(2) mult_left_mono norm_ge_zero t(1) tw zero_le_power)
+ from power_strict_mono[OF t(2), of k] t(1) kas(2) have 121: "t^k \<le> 1"
by auto
- from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] th120 th121]
- have th12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
- from th11 th12 have "cmod (1 + ?w^k * (a + ?w * poly s ?w)) < 1"
- by arith
- then have "cmod (poly ?r ?w) < 1"
- unfolding kas(4)[rule_format, of ?w] r01 by simp
+ from ath[OF norm_ge_zero[of "?w^k * ?w * poly s ?w"] 120 121]
+ have 12: "\<bar>1 - t^k\<bar> + cmod (?w^k * ?w * poly s ?w) < 1" .
then show ?thesis
- by blast
+ by (smt (verit) "11" kas(4) poly_pCons r01)
qed
with cq0 q(2) show ?thesis
- unfolding mrmq_eq not_less[symmetric] by auto
+ by (smt (verit) mrmq_eq)
qed
qed
@@ -825,8 +663,8 @@
have "\<not> constant (poly (pCons c cs))"
proof
assume nc: "constant (poly (pCons c cs))"
- from nc[unfolded constant_def, rule_format, of 0]
- have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0" by auto
+ have "\<forall>w. w \<noteq> 0 \<longrightarrow> poly cs w = 0"
+ by (metis add_cancel_right_right constant_def mult_eq_0_iff nc poly_pCons)
then have "cs = 0"
proof (induct cs)
case 0
@@ -844,24 +682,19 @@
m: "m > 0" "\<forall>z. \<forall>z. cmod z \<le> 1 \<longrightarrow> cmod (poly ds z) \<le> m" by blast
have dm: "cmod d / m > 0"
using False m(1) by (simp add: field_simps)
- from field_lbound_gt_zero[OF dm zero_less_one]
- obtain x where x: "x > 0" "x < cmod d / m" "x < 1"
- by blast
+ then obtain x where x: "x > 0" "x < cmod d / m" "x < 1"
+ by (meson field_lbound_gt_zero zero_less_one)
let ?x = "complex_of_real x"
from x have cx: "?x \<noteq> 0" "cmod ?x \<le> 1"
by simp_all
- from pCons.prems[rule_format, OF cx(1)]
have cth: "cmod (?x*poly ds ?x) = cmod d"
- by (simp add: eq_diff_eq[symmetric])
- from m(2)[rule_format, OF cx(2)] x(1)
- have th0: "cmod (?x*poly ds ?x) \<le> x*m"
- by (simp add: norm_mult)
+ by (metis add_eq_0_iff cx(1) norm_minus_cancel pCons.prems poly_pCons)
+ have 0: "cmod (?x*poly ds ?x) \<le> x*m"
+ by (smt (verit) cx(2) m(2) mult_left_mono norm_mult norm_of_real x(1))
from x(2) m(1) have "x * m < cmod d"
by (simp add: field_simps)
- with th0 have "cmod (?x*poly ds ?x) \<noteq> cmod d"
- by auto
- with cth show ?thesis
- by blast
+ with 0 cth show ?thesis
+ by force
qed
qed
then show False
@@ -936,10 +769,7 @@
next
case False
with sne dpn s oa have dsn: "degree s < n"
- apply auto
- apply (erule ssubst)
- apply (simp add: degree_mult_eq degree_linear_power)
- done
+ by (metis degree_div_less degree_linear_power mult_eq_0_iff n0 nonzero_mult_div_cancel_left not_gr0 pne)
have "poly r x = 0" if h: "poly s x = 0" for x
proof -
have xa: "x \<noteq> a"
@@ -948,10 +778,7 @@
from h[unfolded this poly_eq_0_iff_dvd] obtain u where u: "s = [:- a, 1:] * u"
by (rule dvdE)
have "p = [:- a, 1:] ^ (Suc ?op) * u"
- apply (subst s)
- apply (subst u)
- apply (simp only: power_Suc ac_simps)
- done
+ by (metis (no_types, lifting) mult.assoc power_Suc power_commutes s u)
with ap(2)[unfolded dvd_def] show False
by blast
qed
@@ -962,9 +789,8 @@
with r xa show ?thesis
by auto
qed
- with IH[rule_format, OF dsn, of s r] False have "s dvd (r ^ (degree s))"
+ with False IH dsn obtain u where u: "r ^ (degree s) = s * u"
by blast
- then obtain u where u: "r ^ (degree s) = s * u" ..
then have u': "\<And>x. poly s x * poly u x = poly r x ^ degree s"
by (simp only: poly_mult[symmetric] poly_power[symmetric])
let ?w = "(u * ([:-a,1:] ^ (n - ?op))) * (r ^ (n - degree s))"
@@ -988,16 +814,9 @@
using a order_root pne by blast
next
case False
- with fundamental_theorem_of_algebra_alt[of p]
- obtain c where ccs: "c \<noteq> 0" "p = pCons c 0"
- by blast
- then have pp: "poly p x = c" for x
- by simp
- let ?w = "[:1/c:] * (q ^ n)"
- from ccs have "(q ^ n) = (p * ?w)"
- by simp
then show ?thesis
- unfolding dvd_def by blast
+ using dpn n0 fundamental_theorem_of_algebra_alt[of p]
+ by fastforce
qed
qed
@@ -1010,35 +829,18 @@
then show ?thesis
proof cases
case p: 1
- then have eq: "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0"
+ then have "(\<forall>x. poly p x = (0::complex) \<longrightarrow> poly q x = 0) \<longleftrightarrow> q = 0"
by (auto simp add: poly_all_0_iff_0)
- {
- assume "p dvd (q ^ (degree p))"
- then obtain r where r: "q ^ (degree p) = p * r" ..
- from r p have False by simp
- }
- with eq p show ?thesis by blast
+ with p show ?thesis
+ by force
next
case dp: 2
- then obtain k where k: "p = [:k:]" "k \<noteq> 0"
- by (cases p) (simp split: if_splits)
- then have th1: "\<forall>x. poly p x \<noteq> 0"
- by simp
- from k dp(2) have "q ^ (degree p) = p * [:1/k:]"
- by simp
- then have th2: "p dvd (q ^ (degree p))" ..
- from dp(1) th1 th2 show ?thesis
- by blast
+ then show ?thesis
+ by (meson dvd_trans is_unit_iff_degree poly_eq_0_iff_dvd unit_imp_dvd)
next
case dp: 3
- have False if dvd: "p dvd (q ^ (Suc n))" and h: "poly p x = 0" "poly q x \<noteq> 0" for x
- proof -
- from dvd obtain u where u: "q ^ (Suc n) = p * u" ..
- from h have "poly (q ^ (Suc n)) x \<noteq> 0"
- by simp
- with u h(1) show ?thesis
- by (simp only: poly_mult) simp
- qed
+ have False if "p dvd (q ^ (Suc n))" "poly p x = 0" "poly q x \<noteq> 0" for x
+ by (metis dvd_trans poly_eq_0_iff_dvd poly_power power_eq_0_iff that)
with dp nullstellensatz_lemma[of p q "degree p"] show ?thesis
by auto
qed
@@ -1052,22 +854,14 @@
show ?rhs if ?lhs
proof -
from that[unfolded constant_def, rule_format, of _ "0"]
- have th: "poly p = poly [:poly p 0:]"
+ have "poly p = poly [:poly p 0:]"
by auto
- then have "p = [:poly p 0:]"
- by (simp add: poly_eq_poly_eq_iff)
- then have "degree p = degree [:poly p 0:]"
- by simp
then show ?thesis
- by simp
+ by (metis degree_pCons_0 poly_eq_poly_eq_iff)
qed
show ?lhs if ?rhs
- proof -
- from that obtain k where "p = [:k:]"
- by (cases p) (simp split: if_splits)
- then show ?thesis
- unfolding constant_def by auto
- qed
+ unfolding constant_def
+ by (metis degree_eq_zeroE pcompose_const poly_0 poly_pcompose that)
qed
text \<open>Arithmetic operations on multivariate polynomials.\<close>
@@ -1099,59 +893,21 @@
fixes p:: "('a::comm_ring_1) poly"
assumes pq: "p dvd q"
shows "p dvd (pCons 0 q)"
-proof -
- have "pCons 0 q = q * [:0,1:]" by simp
- then have "q dvd (pCons 0 q)" ..
- with pq show ?thesis by (rule dvd_trans)
-qed
+ by (metis add_0 dvd_def mult_pCons_right pq smult_0_left)
lemma poly_divides_conv0:
fixes p:: "'a::field poly"
- assumes lgpq: "degree q < degree p"
- and lq: "p \<noteq> 0"
- shows "p dvd q \<longleftrightarrow> q = 0" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
- assume ?rhs
- then have "q = p * 0" by simp
- then show ?lhs ..
-next
- assume l: ?lhs
- show ?rhs
- proof (cases "q = 0")
- case True
- then show ?thesis by simp
- next
- assume q0: "q \<noteq> 0"
- from l q0 have "degree p \<le> degree q"
- by (rule dvd_imp_degree_le)
- with lgpq show ?thesis by simp
- qed
-qed
+ assumes lgpq: "degree q < degree p" and lq: "p \<noteq> 0"
+ shows "p dvd q \<longleftrightarrow> q = 0"
+ using lgpq mod_poly_less by fastforce
lemma poly_divides_conv1:
fixes p :: "'a::field poly"
assumes a0: "a \<noteq> 0"
and pp': "p dvd p'"
and qrp': "smult a q - p' = r"
- shows "p dvd q \<longleftrightarrow> p dvd r" (is "?lhs \<longleftrightarrow> ?rhs")
-proof
- from pp' obtain t where t: "p' = p * t" ..
- show ?rhs if ?lhs
- proof -
- from that obtain u where u: "q = p * u" ..
- have "r = p * (smult a u - t)"
- using u qrp' [symmetric] t by (simp add: algebra_simps)
- then show ?thesis ..
- qed
- show ?lhs if ?rhs
- proof -
- from that obtain u where u: "r = p * u" ..
- from u [symmetric] t qrp' [symmetric] a0
- have "q = p * smult (1/a) (u + t)"
- by (simp add: algebra_simps)
- then show ?thesis ..
- qed
-qed
+ shows "p dvd q \<longleftrightarrow> p dvd r"
+ by (metis a0 diff_add_cancel dvd_add_left_iff dvd_smult_iff pp' qrp')
lemma basic_cqe_conv1:
"(\<exists>x. poly p x = 0 \<and> poly 0 x \<noteq> 0) \<longleftrightarrow> False"
@@ -1164,14 +920,7 @@
lemma basic_cqe_conv2:
assumes l: "p \<noteq> 0"
shows "\<exists>x. poly (pCons a (pCons b p)) x = (0::complex)"
-proof -
- have False if "h \<noteq> 0" "t = 0" and "pCons a (pCons b p) = pCons h t" for h t
- using l that by simp
- then have th: "\<not> (\<exists> h t. h \<noteq> 0 \<and> t = 0 \<and> pCons a (pCons b p) = pCons h t)"
- by blast
- from fundamental_theorem_of_algebra_alt[OF th] show ?thesis
- by auto
-qed
+ by (meson fundamental_theorem_of_algebra_alt l pCons_eq_0_iff pCons_eq_iff)
lemma basic_cqe_conv_2b: "(\<exists>x. poly p x \<noteq> (0::complex)) \<longleftrightarrow> p \<noteq> 0"
by (metis poly_all_0_iff_0)
@@ -1180,13 +929,7 @@
fixes p q :: "complex poly"
assumes l: "p \<noteq> 0"
shows "(\<exists>x. poly (pCons a p) x = 0 \<and> poly q x \<noteq> 0) \<longleftrightarrow> \<not> (pCons a p) dvd (q ^ psize p)"
-proof -
- from l have dp: "degree (pCons a p) = psize p"
- by (simp add: psize_def)
- from nullstellensatz_univariate[of "pCons a p" q] l
- show ?thesis
- by (metis dp pCons_eq_0_iff)
-qed
+ by (metis degree_pCons_eq_if l nullstellensatz_univariate pCons_eq_0_iff psize_def)
lemma basic_cqe_conv4:
fixes p q :: "complex poly"
@@ -1195,10 +938,8 @@
proof -
from h have "poly (q ^ n) = poly r"
by auto
- then have "(q ^ n) = r"
+ then show "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
by (simp add: poly_eq_poly_eq_iff)
- then show "p dvd (q ^ n) \<longleftrightarrow> p dvd r"
- by simp
qed
lemma poly_const_conv:
--- a/src/HOL/Filter.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Filter.thy Mon Feb 20 13:59:42 2023 +0100
@@ -1511,7 +1511,21 @@
fixes f :: "'a \<Rightarrow> ('b::unbounded_dense_linorder)" and c :: "'b"
shows "(LIM x F. f x :> at_bot) \<longleftrightarrow> (\<forall>Z<c. eventually (\<lambda>x. Z \<ge> f x) F)"
by (metis filterlim_at_bot filterlim_at_bot_le lt_ex order_le_less_trans)
-
+
+lemma filterlim_at_top_div_const_nat:
+ assumes "c > 0"
+ shows "filterlim (\<lambda>x::nat. x div c) at_top at_top"
+ unfolding filterlim_at_top
+proof
+ fix C :: nat
+ have *: "n div c \<ge> C" if "n \<ge> C * c" for n
+ using assms that by (metis div_le_mono div_mult_self_is_m)
+ have "eventually (\<lambda>n. n \<ge> C * c) at_top"
+ by (rule eventually_ge_at_top)
+ thus "eventually (\<lambda>n. n div c \<ge> C) at_top"
+ by eventually_elim (use * in auto)
+qed
+
lemma filterlim_finite_subsets_at_top:
"filterlim f (finite_subsets_at_top A) F \<longleftrightarrow>
(\<forall>X. finite X \<and> X \<subseteq> A \<longrightarrow> eventually (\<lambda>y. finite (f y) \<and> X \<subseteq> f y \<and> f y \<subseteq> A) F)"
--- a/src/HOL/Library/Landau_Symbols.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Library/Landau_Symbols.thy Mon Feb 20 13:59:42 2023 +0100
@@ -1742,6 +1742,10 @@
by (rule Lim_transform_eventually)
qed (simp_all add: asymp_equiv_def)
+lemma tendsto_imp_asymp_equiv_const:
+ assumes "(f \<longlongrightarrow> c) F" "c \<noteq> 0"
+ shows "f \<sim>[F] (\<lambda>_. c)"
+ by (rule asymp_equivI' tendsto_eq_intros assms refl)+ (use assms in auto)
lemma asymp_equiv_cong:
assumes "eventually (\<lambda>x. f1 x = f2 x) F" "eventually (\<lambda>x. g1 x = g2 x) F"
--- a/src/HOL/Library/Multiset_Order.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Library/Multiset_Order.thy Mon Feb 20 13:59:42 2023 +0100
@@ -147,6 +147,48 @@
lemma irreflp_on_multp\<^sub>H\<^sub>O[simp]: "irreflp_on B (multp\<^sub>H\<^sub>O R)"
by (simp add: irreflp_onI multp\<^sub>H\<^sub>O_def)
+text \<open>The following lemma is a negative result stating that asymmetry of an arbitrary binary
+relation cannot be simply lifted to @{const multp\<^sub>H\<^sub>O}. It suffices to have four distinct values to
+build a counterexample.\<close>
+
+lemma asymp_not_liftable_to_multp\<^sub>H\<^sub>O:
+ fixes a b c d :: 'a
+ assumes "distinct [a, b, c, d]"
+ shows "\<not> (\<forall>(R :: 'a \<Rightarrow> 'a \<Rightarrow> bool). asymp R \<longrightarrow> asymp (multp\<^sub>H\<^sub>O R))"
+proof -
+ define R :: "'a \<Rightarrow> 'a \<Rightarrow> bool" where
+ "R = (\<lambda>x y. x = a \<and> y = c \<or> x = b \<and> y = d \<or> x = c \<and> y = b \<or> x = d \<and> y = a)"
+
+ from assms(1) have "{#a, b#} \<noteq> {#c, d#}"
+ by (metis add_mset_add_single distinct.simps(2) list.set(1) list.simps(15) multi_member_this
+ set_mset_add_mset_insert set_mset_single)
+
+ from assms(1) have "asymp R"
+ by (auto simp: R_def intro: asymp_onI)
+ moreover have "\<not> asymp (multp\<^sub>H\<^sub>O R)"
+ unfolding asymp_on_def Set.ball_simps not_all not_imp not_not
+ proof (intro exI conjI)
+ show "multp\<^sub>H\<^sub>O R {#a, b#} {#c, d#}"
+ unfolding multp\<^sub>H\<^sub>O_def
+ using \<open>{#a, b#} \<noteq> {#c, d#}\<close> R_def assms by auto
+ next
+ show "multp\<^sub>H\<^sub>O R {#c, d#} {#a, b#}"
+ unfolding multp\<^sub>H\<^sub>O_def
+ using \<open>{#a, b#} \<noteq> {#c, d#}\<close> R_def assms by auto
+ qed
+ ultimately show ?thesis
+ unfolding not_all not_imp by auto
+qed
+
+text \<open>However, if the binary relation is both asymmetric and transitive, then @{const multp\<^sub>H\<^sub>O} is
+also asymmetric.\<close>
+
+lemma asymp_multp\<^sub>H\<^sub>O:
+ assumes "asymp R" and "transp R"
+ shows "asymp (multp\<^sub>H\<^sub>O R)"
+ using assms
+ by (metis asymp_on_iff_irreflp_on_if_transp_on irreflp_multp multp_eq_multp\<^sub>H\<^sub>O transp_multp)
+
lemma totalp_on_multp\<^sub>D\<^sub>M:
"totalp_on A R \<Longrightarrow> (\<And>M. M \<in> B \<Longrightarrow> set_mset M \<subseteq> A) \<Longrightarrow> totalp_on B (multp\<^sub>D\<^sub>M R)"
by (smt (verit, ccfv_SIG) count_inI in_mono multp\<^sub>H\<^sub>O_def multp\<^sub>H\<^sub>O_imp_multp\<^sub>D\<^sub>M not_less_iff_gr_or_eq
--- a/src/HOL/Library/Periodic_Fun.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/Library/Periodic_Fun.thy Mon Feb 20 13:59:42 2023 +0100
@@ -147,6 +147,26 @@
lemma cos_eq_periodic_intro:
assumes "x - y = 2*(of_int k)*pi \<or> x + y = 2*(of_int k)*pi"
shows "cos x = cos y"
- by (smt (verit, ccfv_SIG) assms cos_eq_neg_periodic_intro cos_minus_pi cos_periodic_pi)
+ by (smt (verit, best) assms cos_eq_neg_periodic_intro cos_minus_pi cos_periodic_pi)
+
+lemma cos_eq_arccos_Ex:
+ "cos x = y \<longleftrightarrow> -1\<le>y \<and> y\<le>1 \<and> (\<exists>k::int. x = arccos y + 2*k*pi \<or> x = - arccos y + 2*k*pi)" (is "?L=?R")
+proof
+ assume ?R then show "cos x = y"
+ by (metis cos.plus_of_int cos_arccos cos_minus id_apply mult.assoc mult.left_commute of_real_eq_id)
+next
+ assume L: ?L
+ let ?goal = "(\<exists>k::int. x = arccos y + 2*k*pi \<or> x = - arccos y + 2*k*pi)"
+ obtain k::int where k: "-pi < x - k*(2*pi)" "x - k*(2*pi) \<le> pi"
+ using ceiling_divide_lower [of "2*pi" "x-pi"] ceiling_divide_upper [of "2*pi" "x-pi"]
+ by (simp add: divide_simps algebra_simps) (metis mult.commute)
+ have *: "cos (x - k * 2*pi) = y"
+ using cos.periodic_simps(3)[of x "-k"] L by (auto simp add:field_simps)
+ then have **: ?goal when "x-k*2*pi \<ge> 0"
+ using arccos_cos k that by force
+ then show "-1\<le>y \<and> y\<le>1 \<and> ?goal"
+ using "*" arccos_cos2 k(1) by force
+qed
+
end
--- a/src/HOL/MacLaurin.thy Mon Feb 20 13:59:16 2023 +0100
+++ b/src/HOL/MacLaurin.thy Mon Feb 20 13:59:42 2023 +0100
@@ -114,10 +114,8 @@
qed (simp add: differentiable_difg n)
next
case (Suc m')
- then have "\<exists>t. 0 < t \<and> t < h \<and> DERIV (difg m') t :> 0"
- by simp
then obtain t where t: "0 < t" "t < h" "DERIV (difg m') t :> 0"
- by fast
+ by force
have "\<exists>t'. 0 < t' \<and> t' < t \<and> DERIV (difg (Suc m')) t' :> 0"
proof (rule Rolle)
show "0 < t" by fact
@@ -131,14 +129,11 @@
with \<open>t < h\<close> show ?case
by auto
qed
- then obtain t where "0 < t" "t < h" "DERIV (difg m) t :> 0"
- by fast
- with \<open>m < n\<close> have "difg (Suc m) t = 0"
- by (simp add: difg_Suc_eq_0)
+ then obtain t where "0 < t" "t < h" "difg (Suc m) t = 0"
+ using \<open>m < n\<close> difg_Suc_eq_0 by force
show ?thesis
proof (intro exI conjI)
- show "0 < t" by fact
- show "t < h" by fact
+ show "0 < t" "t < h" by fact+
show "f h = (\<Sum>m<n. diff m 0 / (fact m) * h ^ m) + diff n t / (fact n) * h ^ n"
using \<open>difg (Suc m) t = 0\<close> by (simp add: m f_h difg_def)
qed
@@ -157,10 +152,8 @@
next
case Suc
then have "n > 0" by simp
- from INIT1 this INIT2 DERIV
- have "\<exists>t>0. t < h \<and> f h = (\<Sum>m<n. diff m 0 / fact m * h ^ m) + diff n t / fact n * h ^ n"
- by (rule Maclaurin)
- then show ?thesis by fastforce
+ from Maclaurin [OF INIT1 this INIT2 DERIV]
+ show ?thesis by fastforce
qed
lemma Maclaurin_minus:
@@ -216,9 +209,7 @@
then obtain t where "x < t" "t < 0"
"diff 0 x = (\<Sum>m<n. diff m 0 / fact m * x ^ m) + diff n t / fact n * x ^ n"
by blast
- with \<open>x < 0\<close> \<open>diff 0 = f\<close> have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t"
- by simp
- then show ?thesis ..
+ with \<open>x < 0\<close> \<open>diff 0 = f\<close> show ?thesis by force
next
assume "x > 0"
with \<open>n \<noteq> 0\<close> \<open>diff 0 = f\<close> DERIV have "\<exists>t>0. t < x \<and> diff 0 x = ?f x t"
@@ -247,20 +238,12 @@
assume "x < 0"
with assms have "\<exists>t>x. t < 0 \<and> f x = ?f x t"
by (intro Maclaurin_minus) auto
- then obtain t where "t > x" "t < 0" "f x = ?f x t"
- by blast
- with \<open>x < 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
- by simp
- then show ?thesis ..
+ then show ?thesis by force
next
assume "x > 0"
with assms have "\<exists>t>0. t < x \<and> f x = ?f x t"
by (intro Maclaurin) auto
- then obtain t where "t > 0" "t < x" "f x = ?f x t"
- by blast
- with \<open>x > 0\<close> have "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
- by simp
- then show ?thesis ..
+ then show ?thesis by force
qed
lemma Maclaurin_zero: "x = 0 \<Longrightarrow> n \<noteq> 0 \<Longrightarrow> (\<Sum>m<n. (diff m 0 / fact m) * x ^ m) = diff 0 0"
@@ -280,22 +263,7 @@
next
case False
show ?thesis
- proof (cases "x = 0")
- case True
- with \<open>n \<noteq> 0\<close> have "(\<Sum>m<n. diff m 0 / (fact m) * x ^ m) = diff 0 0"
- by (intro Maclaurin_zero) auto
- with INIT \<open>x = 0\<close> \<open>n \<noteq> 0\<close> have " \<bar>0\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x 0"
- by force
- then show ?thesis ..
- next
- case False
- with INIT \<open>n \<noteq> 0\<close> DERIV have "\<exists>t. 0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t"
- by (intro Maclaurin_all_lt) auto
- then obtain t where "0 < \<bar>t\<bar> \<and> \<bar>t\<bar> < \<bar>x\<bar> \<and> f x = ?f x t" ..
- then have "\<bar>t\<bar> \<le> \<bar>x\<bar> \<and> f x = ?f x t"
- by simp
- then show ?thesis ..
- qed
+ using DERIV INIT Maclaurin_bi_le by auto
qed
lemma Maclaurin_all_le_objl:
@@ -324,12 +292,7 @@
using Maclaurin_exp_le [of x 3] by (auto simp: numeral_3_eq_3 power2_eq_square)
corollary ln_2_less_1: "ln 2 < (1::real)"
-proof -
- have "2 < 5/(2::real)" by simp
- also have "5/2 \<le> exp (1::real)" using exp_lower_Taylor_quadratic[of 1, simplified] by simp
- finally have "exp (ln 2) < exp (1::real)" by simp
- thus "ln 2 < (1::real)" by (subst (asm) exp_less_cancel_iff) simp
-qed
+ by (smt (verit) ln_eq_minus_one ln_le_minus_one)
subsection \<open>Version for Sine Function\<close>
@@ -430,8 +393,8 @@
proof (rule Maclaurin_all_lt)
show "\<forall>m x. ((\<lambda>t. cos (t + 1/2 * real m * pi)) has_real_derivative
cos (x + 1/2 * real (Suc m) * pi)) (at x)"
- apply (rule allI derivative_eq_intros | simp)+
- using cos_expansion_lemma by force
+ using cos_expansion_lemma
+ by (intro allI derivative_eq_intros | simp)+
qed (use False in auto)
then show ?thesis
apply (rule ex_forward, simp)
@@ -511,10 +474,10 @@
using mod_exhaust_less_4 [of m]
by (auto simp: minus_one_power_iff even_even_mod_4_iff [of m] dest: even_mod_4_div_2 odd_mod_4_div_2)
show ?thesis
- unfolding sin_coeff_def
apply (subst t2)
apply (rule sin_bound_lemma)
apply (rule sum.cong[OF refl])
+ unfolding sin_coeff_def
apply (subst diff_m_0, simp)
using est
apply (auto intro: mult_right_mono [where b=1, simplified] mult_right_mono
@@ -559,10 +522,8 @@
f (b - c + c) =
(\<Sum>m<n. diff m (0 + c) / fact m * (b - c) ^ m) + diff n (x + c) / fact n * (b - c) ^ n"
by (rule Maclaurin [THEN exE])
- then have "c < x + c \<and> x + c < b \<and> f b =
- (\<Sum>m<n. diff m c / fact m * (b - c) ^ m) + diff n (x + c) / fact n * (b - c) ^ n"
- by fastforce
- then show ?thesis by fastforce
+ then show ?thesis
+ by (smt (verit) sum.cong)
qed
lemma Taylor_down:
@@ -585,10 +546,8 @@
by auto
moreover from DERIV_ident and DERIV_const have "DERIV (\<lambda>x. x + c) t :> 1 + 0"
by (rule DERIV_add)
- ultimately have "DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c) * (1 + 0)"
- by (rule DERIV_chain2)
- then show "DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c)"
- by simp
+ ultimately show "DERIV (\<lambda>x. diff m (x + c)) t :> diff (Suc m) (t + c)"
+ using DERIV_chain2 DERIV_shift by blast
qed
ultimately obtain x where
"a - c < x \<and> x < 0 \<and>
@@ -614,26 +573,15 @@
note INIT
moreover have "\<forall>m t. m < n \<and> x \<le> t \<and> t \<le> b \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
using DERIV and INTERV by fastforce
- moreover note True
- moreover from INTERV have "c \<le> b"
- by simp
- ultimately have "\<exists>t>x. t < c \<and> f x =
- (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n"
- by (rule Taylor_down)
- with True show ?thesis by simp
+ ultimately show ?thesis
+ using True INTERV Taylor_down by simp
next
case False
note INIT
moreover have "\<forall>m t. m < n \<and> a \<le> t \<and> t \<le> x \<longrightarrow> DERIV (diff m) t :> diff (Suc m) t"
using DERIV and INTERV by fastforce
- moreover from INTERV have "a \<le> c"
- by arith
- moreover from False and INTERV have "c < x"
- by arith
- ultimately have "\<exists>t>c. t < x \<and> f x =
- (\<Sum>m<n. diff m c / (fact m) * (x - c) ^ m) + diff n t / (fact n) * (x - c) ^ n"
- by (rule Taylor_up)
- with False show ?thesis by simp
+ ultimately show ?thesis
+ using Taylor_up INTERV False by simp
qed
end
--- a/src/Pure/General/logger.scala Mon Feb 20 13:59:16 2023 +0100
+++ b/src/Pure/General/logger.scala Mon Feb 20 13:59:42 2023 +0100
@@ -35,7 +35,8 @@
}
class System_Logger extends Logger {
- def apply(msg: => String): Unit =
+ def apply(msg: => String): Unit = synchronized {
if (Platform.is_windows) System.out.println(msg)
else System.console.writer.println(msg)
+ }
}
--- a/src/Pure/PIDE/document.scala Mon Feb 20 13:59:16 2023 +0100
+++ b/src/Pure/PIDE/document.scala Mon Feb 20 13:59:42 2023 +0100
@@ -371,14 +371,16 @@
object Nodes {
val empty: Nodes = new Nodes(Graph.empty(Node.Name.Ordering))
+
+ private def init(graph: Graph[Node.Name, Node], name: Node.Name): Graph[Node.Name, Node] =
+ graph.default_node(name, Node.empty)
}
final class Nodes private(graph: Graph[Node.Name, Node]) {
- def apply(name: Node.Name): Node =
- graph.default_node(name, Node.empty).get_node(name)
+ def apply(name: Node.Name): Node = Nodes.init(graph, name).get_node(name)
def is_suppressed(name: Node.Name): Boolean = {
- val graph1 = graph.default_node(name, Node.empty)
+ val graph1 = Nodes.init(graph, name)
graph1.is_maximal(name) && graph1.get_node(name).is_empty
}
@@ -391,10 +393,7 @@
def + (entry: (Node.Name, Node)): Nodes = {
val (name, node) = entry
val imports = node.header.imports
- val graph1 =
- imports.foldLeft(graph.default_node(name, Node.empty)) {
- case (g, p) => g.default_node(p, Node.empty)
- }
+ val graph1 = (name :: imports).foldLeft(graph)(Nodes.init)
val graph2 =
graph1.imm_preds(name).foldLeft(graph1) { case (g, dep) => g.del_edge(dep, name) }
val graph3 = imports.foldLeft(graph2) { case (g, dep) => g.add_edge(dep, name) }
@@ -417,8 +416,8 @@
if name == file_name
} yield cmd).toList
- def descendants(names: List[Node.Name]): List[Node.Name] = graph.all_succs(names)
- def requirements(names: List[Node.Name]): List[Node.Name] = graph.all_preds_rev(names)
+ def descendants(names: List[Node.Name]): List[Node.Name] =
+ names.foldLeft(graph)(Nodes.init).all_succs(names)
def topological_order: List[Node.Name] = graph.topological_order
override def toString: String = topological_order.mkString("Nodes(", ",", ")")
--- a/src/Pure/Thy/sessions.scala Mon Feb 20 13:59:16 2023 +0100
+++ b/src/Pure/Thy/sessions.scala Mon Feb 20 13:59:42 2023 +0100
@@ -1331,6 +1331,18 @@
/** persistent store **/
+ /** auxiliary **/
+
+ sealed case class Build_Info(
+ sources: SHA1.Shasum,
+ input_heaps: SHA1.Shasum,
+ output_heap: SHA1.Shasum,
+ return_code: Int,
+ uuid: String
+ ) {
+ def ok: Boolean = return_code == 0
+ }
+
object Session_Info {
val session_name = SQL.Column.string("session_name").make_primary_key
@@ -1345,7 +1357,7 @@
List(session_name, session_timing, command_timings, theory_timings,
ml_statistics, task_statistics, errors)
- // Build.Session_Info
+ // Build_Info
val sources = SQL.Column.string("sources")
val input_heaps = SQL.Column.string("input_heaps")
val output_heap = SQL.Column.string("output_heap")
@@ -1553,7 +1565,7 @@
session_name: String,
sources: Sources,
build_log: Build_Log.Session_Info,
- build: Build.Session_Info
+ build: Build_Info
): Unit = {
db.transaction {
write_sources(db, session_name, sources)
@@ -1596,7 +1608,7 @@
def read_errors(db: SQL.Database, name: String): List[String] =
Build_Log.uncompress_errors(read_bytes(db, name, Session_Info.errors), cache = cache)
- def read_build(db: SQL.Database, name: String): Option[Build.Session_Info] = {
+ def read_build(db: SQL.Database, name: String): Option[Build_Info] = {
if (db.tables.contains(Session_Info.table.name)) {
db.using_statement(Session_Info.table.select(Nil,
Session_Info.session_name.where_equal(name))) { stmt =>
@@ -1607,7 +1619,7 @@
try { Option(res.string(Session_Info.uuid)).getOrElse("") }
catch { case _: SQLException => "" }
Some(
- Build.Session_Info(
+ Build_Info(
SHA1.fake_shasum(res.string(Session_Info.sources)),
SHA1.fake_shasum(res.string(Session_Info.input_heaps)),
SHA1.fake_shasum(res.string(Session_Info.output_heap)),
--- a/src/Pure/Tools/build.scala Mon Feb 20 13:59:16 2023 +0100
+++ b/src/Pure/Tools/build.scala Mon Feb 20 13:59:42 2023 +0100
@@ -9,22 +9,6 @@
object Build {
- /** auxiliary **/
-
- /* persistent build info */
-
- sealed case class Session_Info(
- sources: SHA1.Shasum,
- input_heaps: SHA1.Shasum,
- output_heap: SHA1.Shasum,
- return_code: Int,
- uuid: String
- ) {
- def ok: Boolean = return_code == 0
- }
-
-
-
/** build with results **/
class Results private[Build](
--- a/src/Pure/Tools/build_job.scala Mon Feb 20 13:59:16 2023 +0100
+++ b/src/Pure/Tools/build_job.scala Mon Feb 20 13:59:42 2023 +0100
@@ -11,7 +11,373 @@
import scala.util.matching.Regex
+trait Build_Job {
+ def session_name: String
+ def numa_node: Option[Int] = None
+ def start(): Unit = ()
+ def terminate(): Unit = ()
+ def is_finished: Boolean = false
+ def join: Process_Result = Process_Result.undefined
+}
+
object Build_Job {
+ class Build_Session(progress: Progress,
+ session_background: Sessions.Background,
+ store: Sessions.Store,
+ val do_store: Boolean,
+ resources: Resources,
+ session_setup: (String, Session) => Unit,
+ val input_heaps: SHA1.Shasum,
+ override val numa_node: Option[Int]
+ ) extends Build_Job {
+ def session_name: String = session_background.session_name
+ val info: Sessions.Info = session_background.sessions_structure(session_name)
+ val options: Options = NUMA.policy_options(info.options, numa_node)
+
+ val session_sources: Sessions.Sources =
+ Sessions.Sources.load(session_background.base, cache = store.cache.compress)
+
+ private lazy val future_result: Future[Process_Result] =
+ Future.thread("build", uninterruptible = true) {
+ val parent = info.parent.getOrElse("")
+
+ val env =
+ Isabelle_System.settings(
+ List("ISABELLE_ML_DEBUGGER" -> options.bool("ML_debugger").toString))
+
+ val is_pure = Sessions.is_pure(session_name)
+
+ val use_prelude = if (is_pure) Thy_Header.ml_roots.map(_._1) else Nil
+
+ val eval_store =
+ if (do_store) {
+ (if (info.theories.nonEmpty) List("ML_Heap.share_common_data ()") else Nil) :::
+ List("ML_Heap.save_child " +
+ ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(session_name))))
+ }
+ else Nil
+
+ def session_blobs(node_name: Document.Node.Name): List[(Command.Blob, Document.Blobs.Item)] =
+ session_background.base.theory_load_commands.get(node_name.theory) match {
+ case None => Nil
+ case Some(spans) =>
+ val syntax = session_background.base.theory_syntax(node_name)
+ val master_dir = Path.explode(node_name.master_dir)
+ for (span <- spans; file <- span.loaded_files(syntax).files)
+ yield {
+ val src_path = Path.explode(file)
+ val blob_name = Document.Node.Name(File.symbolic_path(master_dir + src_path))
+
+ val bytes = session_sources(blob_name.node).bytes
+ val text = bytes.text
+ val chunk = Symbol.Text_Chunk(text)
+
+ Command.Blob(blob_name, src_path, Some((SHA1.digest(bytes), chunk))) ->
+ Document.Blobs.Item(bytes, text, chunk, changed = false)
+ }
+ }
+
+ val session =
+ new Session(options, resources) {
+ override val cache: Term.Cache = store.cache
+
+ override def build_blobs_info(node_name: Document.Node.Name): Command.Blobs_Info =
+ Command.Blobs_Info.make(session_blobs(node_name))
+
+ override def build_blobs(node_name: Document.Node.Name): Document.Blobs =
+ Document.Blobs.make(session_blobs(node_name))
+ }
+
+ object Build_Session_Errors {
+ private val promise: Promise[List[String]] = Future.promise
+
+ def result: Exn.Result[List[String]] = promise.join_result
+ def cancel(): Unit = promise.cancel()
+ def apply(errs: List[String]): Unit = {
+ try { promise.fulfill(errs) }
+ catch { case _: IllegalStateException => }
+ }
+ }
+
+ val export_consumer =
+ Export.consumer(store.open_database(session_name, output = true), store.cache,
+ progress = progress)
+
+ val stdout = new StringBuilder(1000)
+ val stderr = new StringBuilder(1000)
+ val command_timings = new mutable.ListBuffer[Properties.T]
+ val theory_timings = new mutable.ListBuffer[Properties.T]
+ val session_timings = new mutable.ListBuffer[Properties.T]
+ val runtime_statistics = new mutable.ListBuffer[Properties.T]
+ val task_statistics = new mutable.ListBuffer[Properties.T]
+
+ def fun(
+ name: String,
+ acc: mutable.ListBuffer[Properties.T],
+ unapply: Properties.T => Option[Properties.T]
+ ): (String, Session.Protocol_Function) = {
+ name -> ((msg: Prover.Protocol_Output) =>
+ unapply(msg.properties) match {
+ case Some(props) => acc += props; true
+ case _ => false
+ })
+ }
+
+ session.init_protocol_handler(new Session.Protocol_Handler {
+ override def exit(): Unit = Build_Session_Errors.cancel()
+
+ private def build_session_finished(msg: Prover.Protocol_Output): Boolean = {
+ val (rc, errors) =
+ try {
+ val (rc, errs) = {
+ import XML.Decode._
+ pair(int, list(x => x))(Symbol.decode_yxml(msg.text))
+ }
+ val errors =
+ for (err <- errs) yield {
+ val prt = Protocol_Message.expose_no_reports(err)
+ Pretty.string_of(prt, metric = Symbol.Metric)
+ }
+ (rc, errors)
+ }
+ catch { case ERROR(err) => (Process_Result.RC.failure, List(err)) }
+
+ session.protocol_command("Prover.stop", rc.toString)
+ Build_Session_Errors(errors)
+ true
+ }
+
+ private def loading_theory(msg: Prover.Protocol_Output): Boolean =
+ msg.properties match {
+ case Markup.Loading_Theory(Markup.Name(name)) =>
+ progress.theory(Progress.Theory(name, session = session_name))
+ false
+ case _ => false
+ }
+
+ private def export_(msg: Prover.Protocol_Output): Boolean =
+ msg.properties match {
+ case Protocol.Export(args) =>
+ export_consumer.make_entry(session_name, args, msg.chunk)
+ true
+ case _ => false
+ }
+
+ override val functions: Session.Protocol_Functions =
+ List(
+ Markup.Build_Session_Finished.name -> build_session_finished,
+ Markup.Loading_Theory.name -> loading_theory,
+ Markup.EXPORT -> export_,
+ fun(Markup.Theory_Timing.name, theory_timings, Markup.Theory_Timing.unapply),
+ fun(Markup.Session_Timing.name, session_timings, Markup.Session_Timing.unapply),
+ fun(Markup.Task_Statistics.name, task_statistics, Markup.Task_Statistics.unapply))
+ })
+
+ session.command_timings += Session.Consumer("command_timings") {
+ case Session.Command_Timing(props) =>
+ for {
+ elapsed <- Markup.Elapsed.unapply(props)
+ elapsed_time = Time.seconds(elapsed)
+ if elapsed_time.is_relevant && elapsed_time >= options.seconds("command_timing_threshold")
+ } command_timings += props.filter(Markup.command_timing_property)
+ }
+
+ session.runtime_statistics += Session.Consumer("ML_statistics") {
+ case Session.Runtime_Statistics(props) => runtime_statistics += props
+ }
+
+ session.finished_theories += Session.Consumer[Document.Snapshot]("finished_theories") {
+ case snapshot =>
+ if (!progress.stopped) {
+ def export_(name: String, xml: XML.Body, compress: Boolean = true): Unit = {
+ if (!progress.stopped) {
+ val theory_name = snapshot.node_name.theory
+ val args =
+ Protocol.Export.Args(theory_name = theory_name, name = name, compress = compress)
+ val body = Bytes(Symbol.encode(YXML.string_of_body(xml)))
+ export_consumer.make_entry(session_name, args, body)
+ }
+ }
+ def export_text(name: String, text: String, compress: Boolean = true): Unit =
+ export_(name, List(XML.Text(text)), compress = compress)
+
+ for (command <- snapshot.snippet_command) {
+ export_text(Export.DOCUMENT_ID, command.id.toString, compress = false)
+ }
+
+ export_text(Export.FILES,
+ cat_lines(snapshot.node_files.map(name => File.symbolic_path(name.path))),
+ compress = false)
+
+ for ((blob_name, i) <- snapshot.node_files.tail.zipWithIndex) {
+ val xml = snapshot.switch(blob_name).xml_markup()
+ export_(Export.MARKUP + (i + 1), xml)
+ }
+ export_(Export.MARKUP, snapshot.xml_markup())
+ export_(Export.MESSAGES, snapshot.messages.map(_._1))
+ }
+ }
+
+ session.all_messages += Session.Consumer[Any]("build_session_output") {
+ case msg: Prover.Output =>
+ val message = msg.message
+ if (msg.is_system) resources.log(Protocol.message_text(message))
+
+ if (msg.is_stdout) {
+ stdout ++= Symbol.encode(XML.content(message))
+ }
+ else if (msg.is_stderr) {
+ stderr ++= Symbol.encode(XML.content(message))
+ }
+ else if (msg.is_exit) {
+ val err =
+ "Prover terminated" +
+ (msg.properties match {
+ case Markup.Process_Result(result) => ": " + result.print_rc
+ case _ => ""
+ })
+ Build_Session_Errors(List(err))
+ }
+ case _ =>
+ }
+
+ session_setup(session_name, session)
+
+ val eval_main = Command_Line.ML_tool("Isabelle_Process.init_build ()" :: eval_store)
+
+ val process =
+ Isabelle_Process.start(store, options, session, session_background,
+ logic = parent, raw_ml_system = is_pure,
+ use_prelude = use_prelude, eval_main = eval_main,
+ cwd = info.dir.file, env = env)
+
+ val build_errors =
+ Isabelle_Thread.interrupt_handler(_ => process.terminate()) {
+ Exn.capture { process.await_startup() } match {
+ case Exn.Res(_) =>
+ val resources_yxml = resources.init_session_yxml
+ val encode_options: XML.Encode.T[Options] =
+ options => session.prover_options(options).encode
+ val args_yxml =
+ YXML.string_of_body(
+ {
+ import XML.Encode._
+ pair(string, list(pair(encode_options, list(pair(string, properties)))))(
+ (session_name, info.theories))
+ })
+ session.protocol_command("build_session", resources_yxml, args_yxml)
+ Build_Session_Errors.result
+ case Exn.Exn(exn) => Exn.Res(List(Exn.message(exn)))
+ }
+ }
+
+ val process_result =
+ Isabelle_Thread.interrupt_handler(_ => process.terminate()) { process.await_shutdown() }
+
+ session.stop()
+
+ val export_errors =
+ export_consumer.shutdown(close = true).map(Output.error_message_text)
+
+ val (document_output, document_errors) =
+ try {
+ if (build_errors.isInstanceOf[Exn.Res[_]] && process_result.ok && info.documents.nonEmpty) {
+ using(Export.open_database_context(store)) { database_context =>
+ val documents =
+ using(database_context.open_session(session_background)) {
+ session_context =>
+ Document_Build.build_documents(
+ Document_Build.context(session_context, progress = progress),
+ output_sources = info.document_output,
+ output_pdf = info.document_output)
+ }
+ using(database_context.open_database(session_name, output = true))(session_database =>
+ documents.foreach(_.write(session_database.db, session_name)))
+ (documents.flatMap(_.log_lines), Nil)
+ }
+ }
+ else (Nil, Nil)
+ }
+ catch {
+ case exn: Document_Build.Build_Error => (exn.log_lines, exn.log_errors)
+ case Exn.Interrupt.ERROR(msg) => (Nil, List(msg))
+ }
+
+ val result = {
+ val theory_timing =
+ theory_timings.iterator.flatMap(
+ {
+ case props @ Markup.Name(name) => Some(name -> props)
+ case _ => None
+ }).toMap
+ val used_theory_timings =
+ for { (name, _) <- session_background.base.used_theories }
+ yield theory_timing.getOrElse(name.theory, Markup.Name(name.theory))
+
+ val more_output =
+ Library.trim_line(stdout.toString) ::
+ command_timings.toList.map(Protocol.Command_Timing_Marker.apply) :::
+ used_theory_timings.map(Protocol.Theory_Timing_Marker.apply) :::
+ session_timings.toList.map(Protocol.Session_Timing_Marker.apply) :::
+ runtime_statistics.toList.map(Protocol.ML_Statistics_Marker.apply) :::
+ task_statistics.toList.map(Protocol.Task_Statistics_Marker.apply) :::
+ document_output
+
+ process_result.output(more_output)
+ .error(Library.trim_line(stderr.toString))
+ .errors_rc(export_errors ::: document_errors)
+ }
+
+ build_errors match {
+ case Exn.Res(build_errs) =>
+ val errs = build_errs ::: document_errors
+ if (errs.nonEmpty) {
+ result.error_rc.output(
+ errs.flatMap(s => split_lines(Output.error_message_text(s))) :::
+ errs.map(Protocol.Error_Message_Marker.apply))
+ }
+ else if (progress.stopped && result.ok) result.copy(rc = Process_Result.RC.interrupt)
+ else result
+ case Exn.Exn(Exn.Interrupt()) =>
+ if (result.ok) result.copy(rc = Process_Result.RC.interrupt)
+ else result
+ case Exn.Exn(exn) => throw exn
+ }
+ }
+
+ override def start(): Unit = future_result
+ override def terminate(): Unit = future_result.cancel()
+ override def is_finished: Boolean = future_result.is_finished
+
+ private val timeout_request: Option[Event_Timer.Request] = {
+ if (info.timeout_ignored) None
+ else Some(Event_Timer.request(Time.now() + info.timeout) { terminate() })
+ }
+
+ override def join: Process_Result = {
+ val result = future_result.join
+
+ val was_timeout =
+ timeout_request match {
+ case None => false
+ case Some(request) => !request.cancel()
+ }
+
+ if (result.ok) result
+ else if (was_timeout) result.error(Output.error_message_text("Timeout")).timeout_rc
+ else if (result.interrupted) result.error(Output.error_message_text("Interrupt"))
+ else result
+ }
+
+ lazy val finish: SHA1.Shasum = {
+ require(is_finished, "Build job not finished: " + quote(session_name))
+ if (join.ok && do_store && store.output_heap(session_name).is_file) {
+ SHA1.shasum(ML_Heap.write_digest(store.output_heap(session_name)), session_name)
+ }
+ else SHA1.no_shasum
+ }
+ }
+
/* theory markup/messages from session database */
def read_theory(
@@ -234,350 +600,3 @@
}
})
}
-
-class Build_Job(progress: Progress,
- session_background: Sessions.Background,
- store: Sessions.Store,
- val do_store: Boolean,
- resources: Resources,
- session_setup: (String, Session) => Unit,
- val numa_node: Option[Int]
-) {
- def session_name: String = session_background.session_name
- val info: Sessions.Info = session_background.sessions_structure(session_name)
- val options: Options = NUMA.policy_options(info.options, numa_node)
-
- val session_sources: Sessions.Sources =
- Sessions.Sources.load(session_background.base, cache = store.cache.compress)
-
- private val future_result: Future[Process_Result] =
- Future.thread("build", uninterruptible = true) {
- val parent = info.parent.getOrElse("")
-
- val env =
- Isabelle_System.settings(
- List("ISABELLE_ML_DEBUGGER" -> options.bool("ML_debugger").toString))
-
- val is_pure = Sessions.is_pure(session_name)
-
- val use_prelude = if (is_pure) Thy_Header.ml_roots.map(_._1) else Nil
-
- val eval_store =
- if (do_store) {
- (if (info.theories.nonEmpty) List("ML_Heap.share_common_data ()") else Nil) :::
- List("ML_Heap.save_child " +
- ML_Syntax.print_string_bytes(File.platform_path(store.output_heap(session_name))))
- }
- else Nil
-
- def session_blobs(node_name: Document.Node.Name): List[(Command.Blob, Document.Blobs.Item)] =
- session_background.base.theory_load_commands.get(node_name.theory) match {
- case None => Nil
- case Some(spans) =>
- val syntax = session_background.base.theory_syntax(node_name)
- val master_dir = Path.explode(node_name.master_dir)
- for (span <- spans; file <- span.loaded_files(syntax).files)
- yield {
- val src_path = Path.explode(file)
- val blob_name = Document.Node.Name(File.symbolic_path(master_dir + src_path))
-
- val bytes = session_sources(blob_name.node).bytes
- val text = bytes.text
- val chunk = Symbol.Text_Chunk(text)
-
- Command.Blob(blob_name, src_path, Some((SHA1.digest(bytes), chunk))) ->
- Document.Blobs.Item(bytes, text, chunk, changed = false)
- }
- }
-
- val session =
- new Session(options, resources) {
- override val cache: Term.Cache = store.cache
-
- override def build_blobs_info(node_name: Document.Node.Name): Command.Blobs_Info =
- Command.Blobs_Info.make(session_blobs(node_name))
-
- override def build_blobs(node_name: Document.Node.Name): Document.Blobs =
- Document.Blobs.make(session_blobs(node_name))
- }
-
- object Build_Session_Errors {
- private val promise: Promise[List[String]] = Future.promise
-
- def result: Exn.Result[List[String]] = promise.join_result
- def cancel(): Unit = promise.cancel()
- def apply(errs: List[String]): Unit = {
- try { promise.fulfill(errs) }
- catch { case _: IllegalStateException => }
- }
- }
-
- val export_consumer =
- Export.consumer(store.open_database(session_name, output = true), store.cache,
- progress = progress)
-
- val stdout = new StringBuilder(1000)
- val stderr = new StringBuilder(1000)
- val command_timings = new mutable.ListBuffer[Properties.T]
- val theory_timings = new mutable.ListBuffer[Properties.T]
- val session_timings = new mutable.ListBuffer[Properties.T]
- val runtime_statistics = new mutable.ListBuffer[Properties.T]
- val task_statistics = new mutable.ListBuffer[Properties.T]
-
- def fun(
- name: String,
- acc: mutable.ListBuffer[Properties.T],
- unapply: Properties.T => Option[Properties.T]
- ): (String, Session.Protocol_Function) = {
- name -> ((msg: Prover.Protocol_Output) =>
- unapply(msg.properties) match {
- case Some(props) => acc += props; true
- case _ => false
- })
- }
-
- session.init_protocol_handler(new Session.Protocol_Handler {
- override def exit(): Unit = Build_Session_Errors.cancel()
-
- private def build_session_finished(msg: Prover.Protocol_Output): Boolean = {
- val (rc, errors) =
- try {
- val (rc, errs) = {
- import XML.Decode._
- pair(int, list(x => x))(Symbol.decode_yxml(msg.text))
- }
- val errors =
- for (err <- errs) yield {
- val prt = Protocol_Message.expose_no_reports(err)
- Pretty.string_of(prt, metric = Symbol.Metric)
- }
- (rc, errors)
- }
- catch { case ERROR(err) => (Process_Result.RC.failure, List(err)) }
-
- session.protocol_command("Prover.stop", rc.toString)
- Build_Session_Errors(errors)
- true
- }
-
- private def loading_theory(msg: Prover.Protocol_Output): Boolean =
- msg.properties match {
- case Markup.Loading_Theory(Markup.Name(name)) =>
- progress.theory(Progress.Theory(name, session = session_name))
- false
- case _ => false
- }
-
- private def export_(msg: Prover.Protocol_Output): Boolean =
- msg.properties match {
- case Protocol.Export(args) =>
- export_consumer.make_entry(session_name, args, msg.chunk)
- true
- case _ => false
- }
-
- override val functions: Session.Protocol_Functions =
- List(
- Markup.Build_Session_Finished.name -> build_session_finished,
- Markup.Loading_Theory.name -> loading_theory,
- Markup.EXPORT -> export_,
- fun(Markup.Theory_Timing.name, theory_timings, Markup.Theory_Timing.unapply),
- fun(Markup.Session_Timing.name, session_timings, Markup.Session_Timing.unapply),
- fun(Markup.Task_Statistics.name, task_statistics, Markup.Task_Statistics.unapply))
- })
-
- session.command_timings += Session.Consumer("command_timings") {
- case Session.Command_Timing(props) =>
- for {
- elapsed <- Markup.Elapsed.unapply(props)
- elapsed_time = Time.seconds(elapsed)
- if elapsed_time.is_relevant && elapsed_time >= options.seconds("command_timing_threshold")
- } command_timings += props.filter(Markup.command_timing_property)
- }
-
- session.runtime_statistics += Session.Consumer("ML_statistics") {
- case Session.Runtime_Statistics(props) => runtime_statistics += props
- }
-
- session.finished_theories += Session.Consumer[Document.Snapshot]("finished_theories") {
- case snapshot =>
- if (!progress.stopped) {
- def export_(name: String, xml: XML.Body, compress: Boolean = true): Unit = {
- if (!progress.stopped) {
- val theory_name = snapshot.node_name.theory
- val args =
- Protocol.Export.Args(theory_name = theory_name, name = name, compress = compress)
- val body = Bytes(Symbol.encode(YXML.string_of_body(xml)))
- export_consumer.make_entry(session_name, args, body)
- }
- }
- def export_text(name: String, text: String, compress: Boolean = true): Unit =
- export_(name, List(XML.Text(text)), compress = compress)
-
- for (command <- snapshot.snippet_command) {
- export_text(Export.DOCUMENT_ID, command.id.toString, compress = false)
- }
-
- export_text(Export.FILES,
- cat_lines(snapshot.node_files.map(name => File.symbolic_path(name.path))),
- compress = false)
-
- for ((blob_name, i) <- snapshot.node_files.tail.zipWithIndex) {
- val xml = snapshot.switch(blob_name).xml_markup()
- export_(Export.MARKUP + (i + 1), xml)
- }
- export_(Export.MARKUP, snapshot.xml_markup())
- export_(Export.MESSAGES, snapshot.messages.map(_._1))
- }
- }
-
- session.all_messages += Session.Consumer[Any]("build_session_output") {
- case msg: Prover.Output =>
- val message = msg.message
- if (msg.is_system) resources.log(Protocol.message_text(message))
-
- if (msg.is_stdout) {
- stdout ++= Symbol.encode(XML.content(message))
- }
- else if (msg.is_stderr) {
- stderr ++= Symbol.encode(XML.content(message))
- }
- else if (msg.is_exit) {
- val err =
- "Prover terminated" +
- (msg.properties match {
- case Markup.Process_Result(result) => ": " + result.print_rc
- case _ => ""
- })
- Build_Session_Errors(List(err))
- }
- case _ =>
- }
-
- session_setup(session_name, session)
-
- val eval_main = Command_Line.ML_tool("Isabelle_Process.init_build ()" :: eval_store)
-
- val process =
- Isabelle_Process.start(store, options, session, session_background,
- logic = parent, raw_ml_system = is_pure,
- use_prelude = use_prelude, eval_main = eval_main,
- cwd = info.dir.file, env = env)
-
- val build_errors =
- Isabelle_Thread.interrupt_handler(_ => process.terminate()) {
- Exn.capture { process.await_startup() } match {
- case Exn.Res(_) =>
- val resources_yxml = resources.init_session_yxml
- val encode_options: XML.Encode.T[Options] =
- options => session.prover_options(options).encode
- val args_yxml =
- YXML.string_of_body(
- {
- import XML.Encode._
- pair(string, list(pair(encode_options, list(pair(string, properties)))))(
- (session_name, info.theories))
- })
- session.protocol_command("build_session", resources_yxml, args_yxml)
- Build_Session_Errors.result
- case Exn.Exn(exn) => Exn.Res(List(Exn.message(exn)))
- }
- }
-
- val process_result =
- Isabelle_Thread.interrupt_handler(_ => process.terminate()) { process.await_shutdown() }
-
- session.stop()
-
- val export_errors =
- export_consumer.shutdown(close = true).map(Output.error_message_text)
-
- val (document_output, document_errors) =
- try {
- if (build_errors.isInstanceOf[Exn.Res[_]] && process_result.ok && info.documents.nonEmpty) {
- using(Export.open_database_context(store)) { database_context =>
- val documents =
- using(database_context.open_session(session_background)) {
- session_context =>
- Document_Build.build_documents(
- Document_Build.context(session_context, progress = progress),
- output_sources = info.document_output,
- output_pdf = info.document_output)
- }
- using(database_context.open_database(session_name, output = true))(session_database =>
- documents.foreach(_.write(session_database.db, session_name)))
- (documents.flatMap(_.log_lines), Nil)
- }
- }
- else (Nil, Nil)
- }
- catch {
- case exn: Document_Build.Build_Error => (exn.log_lines, exn.log_errors)
- case Exn.Interrupt.ERROR(msg) => (Nil, List(msg))
- }
-
- val result = {
- val theory_timing =
- theory_timings.iterator.flatMap(
- {
- case props @ Markup.Name(name) => Some(name -> props)
- case _ => None
- }).toMap
- val used_theory_timings =
- for { (name, _) <- session_background.base.used_theories }
- yield theory_timing.getOrElse(name.theory, Markup.Name(name.theory))
-
- val more_output =
- Library.trim_line(stdout.toString) ::
- command_timings.toList.map(Protocol.Command_Timing_Marker.apply) :::
- used_theory_timings.map(Protocol.Theory_Timing_Marker.apply) :::
- session_timings.toList.map(Protocol.Session_Timing_Marker.apply) :::
- runtime_statistics.toList.map(Protocol.ML_Statistics_Marker.apply) :::
- task_statistics.toList.map(Protocol.Task_Statistics_Marker.apply) :::
- document_output
-
- process_result.output(more_output)
- .error(Library.trim_line(stderr.toString))
- .errors_rc(export_errors ::: document_errors)
- }
-
- build_errors match {
- case Exn.Res(build_errs) =>
- val errs = build_errs ::: document_errors
- if (errs.nonEmpty) {
- result.error_rc.output(
- errs.flatMap(s => split_lines(Output.error_message_text(s))) :::
- errs.map(Protocol.Error_Message_Marker.apply))
- }
- else if (progress.stopped && result.ok) result.copy(rc = Process_Result.RC.interrupt)
- else result
- case Exn.Exn(Exn.Interrupt()) =>
- if (result.ok) result.copy(rc = Process_Result.RC.interrupt)
- else result
- case Exn.Exn(exn) => throw exn
- }
- }
-
- def terminate(): Unit = future_result.cancel()
- def is_finished: Boolean = future_result.is_finished
-
- private val timeout_request: Option[Event_Timer.Request] = {
- if (info.timeout_ignored) None
- else Some(Event_Timer.request(Time.now() + info.timeout) { terminate() })
- }
-
- def join: Process_Result = {
- val result = future_result.join
-
- val was_timeout =
- timeout_request match {
- case None => false
- case Some(request) => !request.cancel()
- }
-
- if (result.ok) result
- else if (was_timeout) result.error(Output.error_message_text("Timeout")).timeout_rc
- else if (result.interrupted) result.error(Output.error_message_text("Interrupt"))
- else result
- }
-}
--- a/src/Pure/Tools/build_process.scala Mon Feb 20 13:59:16 2023 +0100
+++ b/src/Pure/Tools/build_process.scala Mon Feb 20 13:59:42 2023 +0100
@@ -160,13 +160,6 @@
private val build_deps = build_context.deps
private val progress = build_context.progress
- // global state
- private val numa_nodes = new NUMA.Nodes(numa_shuffling)
- private var build_graph = build_context.sessions_structure.build_graph
- private var build_order = SortedSet.from(build_graph.keys)(build_context.ordering)
- private var running = Map.empty[String, (SHA1.Shasum, Build_Job)]
- private var results = Map.empty[String, Build_Process.Result]
-
private val log =
build_options.string("system_log") match {
case "" => No_Logger
@@ -174,19 +167,65 @@
case log_file => Logger.make(Some(Path.explode(log_file)))
}
- private def remove_pending(name: String): Unit = {
- build_graph = build_graph.del_node(name)
- build_order = build_order - name
+ // global state
+ private val _numa_nodes = new NUMA.Nodes(numa_shuffling)
+ private var _build_graph = build_context.sessions_structure.build_graph
+ private var _build_order = SortedSet.from(_build_graph.keys)(build_context.ordering)
+ private var _running = Map.empty[String, Build_Job]
+ private var _results = Map.empty[String, Build_Process.Result]
+
+ private def remove_pending(name: String): Unit = synchronized {
+ _build_graph = _build_graph.del_node(name)
+ _build_order = _build_order - name
+ }
+
+ private def next_pending(): Option[String] = synchronized {
+ if (_running.size < (max_jobs max 1)) {
+ _build_order.iterator
+ .dropWhile(name => _running.isDefinedAt(name) || !_build_graph.is_minimal(name))
+ .nextOption()
+ }
+ else None
+ }
+
+ private def next_numa_node(): Option[Int] = synchronized {
+ _numa_nodes.next(used =
+ Set.from(for { job <- _running.valuesIterator; i <- job.numa_node } yield i))
}
- private def next_pending(): Option[String] =
- build_order.iterator
- .dropWhile(name => running.isDefinedAt(name) || !build_graph.is_minimal(name))
- .nextOption()
+ private def test_running(): Boolean = synchronized { !_build_graph.is_empty }
+
+ private def stop_running(): Unit = synchronized { _running.valuesIterator.foreach(_.terminate()) }
+
+ private def finished_running(): List[Build_Job.Build_Session] = synchronized {
+ List.from(
+ _running.valuesIterator.flatMap {
+ case job: Build_Job.Build_Session if job.is_finished => Some(job)
+ case _ => None
+ })
+ }
+
+ private def job_running(name: String, job: Build_Job): Build_Job = synchronized {
+ _running += (name -> job)
+ job
+ }
- private def used_node(i: Int): Boolean =
- running.iterator.exists(
- { case (_, (_, job)) => job.numa_node.isDefined && job.numa_node.get == i })
+ private def remove_running(name: String): Unit = synchronized {
+ _running -= name
+ }
+
+ private def add_result(
+ name: String,
+ current: Boolean,
+ output_heap: SHA1.Shasum,
+ process_result: Process_Result
+ ): Unit = synchronized {
+ _results += (name -> Build_Process.Result(current, output_heap, process_result))
+ }
+
+ private def get_results(names: List[String]): List[Build_Process.Result] = synchronized {
+ names.map(_results.apply)
+ }
private def session_finished(session_name: String, process_result: Process_Result): String =
"Finished " + session_name + " (" + process_result.timing.message_resources + ")"
@@ -198,14 +237,10 @@
"Timing " + session_name + " (" + threads + " threads, " + timing.message_factor + ")"
}
- private def finish_job(session_name: String, input_heaps: SHA1.Shasum, job: Build_Job): Unit = {
+ private def finish_job(job: Build_Job.Build_Session): Unit = {
+ val session_name = job.session_name
val process_result = job.join
-
- val output_heap =
- if (process_result.ok && job.do_store && store.output_heap(session_name).is_file) {
- SHA1.shasum(ML_Heap.write_digest(store.output_heap(session_name)), session_name)
- }
- else SHA1.no_shasum
+ val output_heap = job.finish
val log_lines = process_result.out_lines.filterNot(Protocol_Message.Marker.test)
val process_result_tail = {
@@ -236,7 +271,7 @@
build_log =
if (process_result.timeout) build_log.error("Timeout") else build_log,
build =
- Build.Session_Info(build_deps.sources_shasum(session_name), input_heaps,
+ Sessions.Build_Info(build_deps.sources_shasum(session_name), job.input_heaps,
output_heap, process_result.rc, UUID.random().toString)))
// messages
@@ -251,15 +286,18 @@
if (!process_result.interrupted) progress.echo(process_result_tail.out)
}
- remove_pending(session_name)
- running -= session_name
- results += (session_name -> Build_Process.Result(false, output_heap, process_result_tail))
+ synchronized {
+ remove_pending(session_name)
+ remove_running(session_name)
+ add_result(session_name, false, output_heap, process_result_tail)
+ }
}
private def start_job(session_name: String): Unit = {
val ancestor_results =
- build_deps.sessions_structure.build_requirements(List(session_name)).
- filterNot(_ == session_name).map(results(_))
+ get_results(
+ build_deps.sessions_structure.build_requirements(List(session_name)).
+ filterNot(_ == session_name))
val input_heaps =
if (ancestor_results.isEmpty) {
SHA1.shasum_meta_info(SHA1.digest(Path.explode("$POLYML_EXE")))
@@ -289,13 +327,17 @@
val all_current = current && ancestor_results.forall(_.current)
if (all_current) {
- remove_pending(session_name)
- results += (session_name -> Build_Process.Result(true, output_heap, Process_Result.ok))
+ synchronized {
+ remove_pending(session_name)
+ add_result(session_name, true, output_heap, Process_Result.ok)
+ }
}
else if (no_build) {
progress.echo_if(verbose, "Skipping " + session_name + " ...")
- remove_pending(session_name)
- results += (session_name -> Build_Process.Result(false, output_heap, Process_Result.error))
+ synchronized {
+ remove_pending(session_name)
+ add_result(session_name, false, output_heap, Process_Result.error)
+ }
}
else if (ancestor_results.forall(_.ok) && !progress.stopped) {
progress.echo((if (do_store) "Building " else "Running ") + session_name + " ...")
@@ -309,16 +351,21 @@
new Resources(session_background, log = log,
command_timings = build_context(session_name).old_command_timings)
- val numa_node = numa_nodes.next(used_node)
val job =
- new Build_Job(progress, session_background, store, do_store,
- resources, session_setup, numa_node)
- running += (session_name -> (input_heaps, job))
+ synchronized {
+ val numa_node = next_numa_node()
+ job_running(session_name,
+ new Build_Job.Build_Session(progress, session_background, store, do_store,
+ resources, session_setup, input_heaps, numa_node))
+ }
+ job.start()
}
else {
progress.echo(session_name + " CANCELLED")
- remove_pending(session_name)
- results += (session_name -> Build_Process.Result(false, output_heap, Process_Result.undefined))
+ synchronized {
+ remove_pending(session_name)
+ add_result(session_name, false, output_heap, Process_Result.undefined)
+ }
}
}
@@ -328,23 +375,17 @@
}
def run(): Map[String, Build_Process.Result] = {
- while (!build_graph.is_empty) {
- if (progress.stopped) {
- for ((_, (_, job)) <- running) job.terminate()
- }
+ while (test_running()) {
+ if (progress.stopped) stop_running()
- running.find({ case (_, (_, job)) => job.is_finished }) match {
- case Some((session_name, (input_heaps, job))) =>
- finish_job(session_name, input_heaps, job)
- case None if running.size < (max_jobs max 1) =>
- next_pending() match {
- case Some(session_name) => start_job(session_name)
- case None => sleep()
- }
+ for (job <- finished_running()) finish_job(job)
+
+ next_pending() match {
+ case Some(session_name) => start_job(session_name)
case None => sleep()
}
}
- results
+ synchronized { _results }
}
}