# HG changeset patch # User Manuel Eberl # Date 1609789378 -3600 # Node ID 7ad9f197ca7e961e5993f72a937763bfcc1e361f # Parent ab9e27da0e85c365eefdeb45a1e711ddb923affd HOL-Complex_Analysis: coefficient asymptotics for meromorphic functions diff -r ab9e27da0e85 -r 7ad9f197ca7e src/HOL/Complex_Analysis/Residue_Theorem.thy --- a/src/HOL/Complex_Analysis/Residue_Theorem.thy Mon Jan 04 19:41:38 2021 +0100 +++ b/src/HOL/Complex_Analysis/Residue_Theorem.thy Mon Jan 04 20:42:58 2021 +0100 @@ -1,6 +1,6 @@ section \The Residue Theorem, the Argument Principle and Rouch\'{e}'s Theorem\ theory Residue_Theorem - imports Complex_Residues + imports Complex_Residues "HOL-Library.Landau_Symbols" begin subsection \Cauchy's residue theorem\ @@ -666,6 +666,130 @@ then show ?thesis unfolding ff_def c_def w_def by simp qed + +subsection \Coefficient asymptotics for generating functions\ + +text \ + For a formal power series that has a meromorphic continuation on some disc in the + context plane, we can use the Residue Theorem to extract precise asymptotic information + from the residues at the poles. This can be used to derive the asymptotic behaviour + of the coefficients (\a\<^sub>n \ \\). If the function is meromorphic on the entire + complex plane, one can even derive a full asymptotic expansion. + + We will first show the relationship between the coefficients and the sum over the residues + with an explicit remainder term (the contour integral along the circle used in the + Residue theorem). +\ +theorem + fixes f :: "complex \ complex" and n :: nat and r :: real + defines "g \ (\w. f w / w ^ Suc n)" and "\ \ circlepath 0 r" + assumes "open A" "connected A" "cball 0 r \ A" "r > 0" + assumes "f holomorphic_on A - S" "S \ ball 0 r" "finite S" "0 \ S" + shows fps_coeff_conv_residues: + "(deriv ^^ n) f 0 / fact n = + contour_integral \ g / (2 * pi * \) - (\z\S. residue g z)" (is ?thesis1) + and fps_coeff_residues_bound: + "(\z. norm z = r \ z \ k \ norm (f z) \ C) \ C \ 0 \ finite k \ + norm ((deriv ^^ n) f 0 / fact n + (\z\S. residue g z)) \ C / r ^ n" +proof - + have holo: "g holomorphic_on A - insert 0 S" + unfolding g_def using assms by (auto intro!: holomorphic_intros) + have "contour_integral \ g = 2 * pi * \ * (\z\insert 0 S. winding_number \ z * residue g z)" + proof (rule Residue_theorem) + show "g holomorphic_on A - insert 0 S" by fact + from assms show "\z. z \ A \ winding_number \ z = 0" + unfolding \_def by (intro allI impI winding_number_zero_outside[of _ "cball 0 r"]) auto + qed (insert assms, auto simp: \_def) + also have "winding_number \ z = 1" if "z \ insert 0 S" for z + unfolding \_def using assms that by (intro winding_number_circlepath) auto + hence "(\z\insert 0 S. winding_number \ z * residue g z) = (\z\insert 0 S. residue g z)" + by (intro sum.cong) simp_all + also have "\ = residue g 0 + (\z\S. residue g z)" + using \0 \ S\ and \finite S\ by (subst sum.insert) auto + also from \r > 0\ have "0 \ cball 0 r" by simp + with assms have "0 \ A - S" by blast + with assms have "residue g 0 = (deriv ^^ n) f 0 / fact n" + unfolding g_def by (subst residue_holomorphic_over_power'[of "A - S"]) + (auto simp: finite_imp_closed) + finally show ?thesis1 + by (simp add: field_simps) + + assume C: "\z. norm z = r \ z \ k \ norm (f z) \ C" "C \ 0" and k: "finite k" + have "(deriv ^^ n) f 0 / fact n + (\z\S. residue g z) = contour_integral \ g / (2 * pi * \)" + using \?thesis1\ by (simp add: algebra_simps) + also have "norm \ = norm (contour_integral \ g) / (2 * pi)" + by (simp add: norm_divide norm_mult) + also have "norm (contour_integral \ g) \ C / r ^ Suc n * (2 * pi * r)" + proof (rule has_contour_integral_bound_circlepath_strong) + from \open A\ and \finite S\ have "open (A - insert 0 S)" + by (blast intro: finite_imp_closed) + with assms show "(g has_contour_integral contour_integral \ g) (circlepath 0 r)" + unfolding \_def + by (intro has_contour_integral_integral contour_integrable_holomorphic_simple [OF holo]) auto + next + fix z assume z: "norm (z - 0) = r" "z \ k" + hence "norm (g z) = norm (f z) / r ^ Suc n" + by (simp add: norm_divide g_def norm_mult norm_power) + also have "\ \ C / r ^ Suc n" + using k and \r > 0\ and z by (intro divide_right_mono C zero_le_power) auto + finally show "norm (g z) \ C / r ^ Suc n" . + qed (insert C(2) k \r > 0\, auto) + also from \r > 0\ have "C / r ^ Suc n * (2 * pi * r) / (2 * pi) = C / r ^ n" + by simp + finally show "norm ((deriv ^^ n) f 0 / fact n + (\z\S. residue g z)) \ \" + by - (simp_all add: divide_right_mono) +qed + +text \ + Since the circle is fixed, we can get an upper bound on the values of the generating + function on the circle and therefore show that the integral is $O(r^{-n})$. +\ +corollary fps_coeff_residues_bigo: + fixes f :: "complex \ complex" and r :: real + assumes "open A" "connected A" "cball 0 r \ A" "r > 0" + assumes "f holomorphic_on A - S" "S \ ball 0 r" "finite S" "0 \ S" + assumes g: "eventually (\n. g n = -(\z\S. residue (\z. f z / z ^ Suc n) z)) sequentially" + (is "eventually (\n. _ = -?g' n) _") + shows "(\n. (deriv ^^ n) f 0 / fact n - g n) \ O(\n. 1 / r ^ n)" (is "(\n. ?c n - _) \ O(_)") +proof - + from assms have "compact (f ` sphere 0 r)" + by (intro compact_continuous_image holomorphic_on_imp_continuous_on + holomorphic_on_subset[OF \f holomorphic_on A - S\]) auto + hence "bounded (f ` sphere 0 r)" by (rule compact_imp_bounded) + then obtain C where C: "\z. z \ sphere 0 r \ norm (f z) \ C" + by (auto simp: bounded_iff sphere_def) + have "0 \ norm (f (of_real r))" by simp + also from C[of "of_real r"] and \r > 0\ have "\ \ C" by simp + finally have C_nonneg: "C \ 0" . + + have "(\n. ?c n + ?g' n) \ O(\n. of_real (1 / r ^ n))" + proof (intro bigoI[of _ C] always_eventually allI ) + fix n :: nat + from assms and C and C_nonneg have "norm (?c n + ?g' n) \ C / r ^ n" + by (intro fps_coeff_residues_bound[where A = A and k = "{}"]) auto + also have "\ = C * norm (complex_of_real (1 / r ^ n))" + using \r > 0\ by (simp add: norm_divide norm_power) + finally show "norm (?c n + ?g' n) \ \" . + qed + also have "?this \ (\n. ?c n - g n) \ O(\n. of_real (1 / r ^ n))" + by (intro landau_o.big.in_cong eventually_mono[OF g]) simp_all + finally show ?thesis . +qed + +corollary fps_coeff_residues_bigo': + fixes f :: "complex \ complex" and r :: real + assumes exp: "f has_fps_expansion F" + assumes "open A" "connected A" "cball 0 r \ A" "r > 0" + assumes "f holomorphic_on A - S" "S \ ball 0 r" "finite S" "0 \ S" + assumes "eventually (\n. g n = -(\z\S. residue (\z. f z / z ^ Suc n) z)) sequentially" + (is "eventually (\n. _ = -?g' n) _") + shows "(\n. fps_nth F n - g n) \ O(\n. 1 / r ^ n)" (is "(\n. ?c n - _) \ O(_)") +proof - + have "fps_nth F = (\n. (deriv ^^ n) f 0 / fact n)" + using fps_nth_fps_expansion[OF exp] by (intro ext) simp_all + with fps_coeff_residues_bigo[OF assms(2-)] show ?thesis by simp +qed + subsection \Rouche's theorem \ theorem Rouche_theorem: