diff --git a/src/HOL/Analysis/Complex_Transcendental.thy b/src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy +++ b/src/HOL/Analysis/Complex_Transcendental.thy @@ -1,4063 +1,4061 @@ section \Complex Transcendental Functions\ text\By John Harrison et al. Ported from HOL Light by L C Paulson (2015)\ theory Complex_Transcendental imports Complex_Analysis_Basics Summation_Tests "HOL-Library.Periodic_Fun" begin subsection\Möbius transformations\ (* TODO: Figure out what to do with Möbius transformations *) definition\<^marker>\tag important\ "moebius a b c d \ (\z. (a*z+b) / (c*z+d :: 'a :: field))" theorem moebius_inverse: assumes "a * d \ b * c" "c * z + d \ 0" shows "moebius d (-b) (-c) a (moebius a b c d z) = z" proof - from assms have "(-c) * moebius a b c d z + a \ 0" unfolding moebius_def by (simp add: field_simps) with assms show ?thesis unfolding moebius_def by (simp add: moebius_def divide_simps) (simp add: algebra_simps)? qed lemma moebius_inverse': assumes "a * d \ b * c" "c * z - a \ 0" shows "moebius a b c d (moebius d (-b) (-c) a z) = z" using assms moebius_inverse[of d a "-b" "-c" z] by (auto simp: algebra_simps) lemma cmod_add_real_less: assumes "Im z \ 0" "r\0" shows "cmod (z + r) < cmod z + \r\" proof (cases z) case (Complex x y) then have "0 < y * y" using assms mult_neg_neg by force with assms have "r * x / \r\ < sqrt (x*x + y*y)" by (simp add: real_less_rsqrt power2_eq_square) then show ?thesis using assms Complex apply (simp add: cmod_def) apply (rule power2_less_imp_less, auto) apply (simp add: power2_eq_square field_simps) done qed lemma cmod_diff_real_less: "Im z \ 0 \ x\0 \ cmod (z - x) < cmod z + \x\" using cmod_add_real_less [of z "-x"] by simp lemma cmod_square_less_1_plus: assumes "Im z = 0 \ \Re z\ < 1" shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)" proof (cases "Im z = 0 \ Re z = 0") case True with assms abs_square_less_1 show ?thesis by (force simp add: Re_power2 Im_power2 cmod_def) next case False with cmod_diff_real_less [of "1 - z\<^sup>2" "1"] show ?thesis by (simp add: norm_power Im_power2) qed subsection\<^marker>\tag unimportant\\The Exponential Function\ lemma norm_exp_i_times [simp]: "norm (exp(\ * of_real y)) = 1" by simp lemma norm_exp_imaginary: "norm(exp z) = 1 \ Re z = 0" by simp lemma field_differentiable_within_exp: "exp field_differentiable (at z within s)" using DERIV_exp field_differentiable_at_within field_differentiable_def by blast lemma continuous_within_exp: fixes z::"'a::{real_normed_field,banach}" shows "continuous (at z within s) exp" by (simp add: continuous_at_imp_continuous_within) lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s" by (simp add: field_differentiable_within_exp holomorphic_on_def) lemma holomorphic_on_exp' [holomorphic_intros]: "f holomorphic_on s \ (\x. exp (f x)) holomorphic_on s" using holomorphic_on_compose[OF _ holomorphic_on_exp] by (simp add: o_def) subsection\Euler and de Moivre formulas\ text\The sine series times \<^term>\i\\ lemma sin_i_eq: "(\n. (\ * sin_coeff n) * z^n) sums (\ * sin z)" proof - have "(\n. \ * sin_coeff n *\<^sub>R z^n) sums (\ * sin z)" using sin_converges sums_mult by blast then show ?thesis by (simp add: scaleR_conv_of_real field_simps) qed theorem exp_Euler: "exp(\ * z) = cos(z) + \ * sin(z)" proof - have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) = (\n. (\ * z) ^ n /\<^sub>R (fact n))" proof fix n show "(cos_coeff n + \ * sin_coeff n) * z^n = (\ * 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 also have "... sums (exp (\ * z))" by (rule exp_converges) finally have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (exp (\ * z))" . moreover have "(\n. (cos_coeff n + \ * sin_coeff n) * z^n) sums (cos z + \ * sin z)" using sums_add [OF cos_converges [of z] sin_i_eq [of z]] by (simp add: field_simps scaleR_conv_of_real) ultimately show ?thesis using sums_unique2 by blast qed corollary\<^marker>\tag unimportant\ exp_minus_Euler: "exp(-(\ * z)) = cos(z) - \ * sin(z)" using exp_Euler [of "-z"] by simp lemma sin_exp_eq: "sin z = (exp(\ * z) - exp(-(\ * z))) / (2*\)" by (simp add: exp_Euler exp_minus_Euler) lemma sin_exp_eq': "sin z = \ * (exp(-(\ * z)) - exp(\ * z)) / 2" by (simp add: exp_Euler exp_minus_Euler) lemma cos_exp_eq: "cos z = (exp(\ * z) + exp(-(\ * z))) / 2" by (simp add: exp_Euler exp_minus_Euler) theorem Euler: "exp(z) = of_real(exp(Re z)) * (of_real(cos(Im z)) + \ * of_real(sin(Im z)))" by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real Complex_eq) lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Re_divide Im_exp) lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2" by (simp add: sin_exp_eq field_simps Im_divide Re_exp) lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2" by (simp add: cos_exp_eq field_simps Re_divide Re_exp) lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2" by (simp add: cos_exp_eq field_simps Im_divide Im_exp) lemma Re_sin_pos: "0 < Re z \ Re z < pi \ Re (sin z) > 0" by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero) lemma Im_sin_nonneg: "Re z = 0 \ 0 \ Im z \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) lemma Im_sin_nonneg2: "Re z = pi \ Im z \ 0 \ 0 \ Im (sin z)" by (simp add: Re_sin Im_sin algebra_simps) subsection\<^marker>\tag unimportant\\Relationships between real and complex trigonometric and hyperbolic functions\ lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x" by (simp add: sin_of_real) lemma real_cos_eq [simp]: "Re(cos(of_real x)) = cos x" by (simp add: cos_of_real) lemma DeMoivre: "(cos z + \ * sin z) ^ n = cos(n * z) + \ * sin(n * z)" by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute) lemma exp_cnj: "cnj (exp z) = exp (cnj z)" proof - have "(\n. cnj (z ^ n /\<^sub>R (fact n))) = (\n. (cnj z)^n /\<^sub>R (fact n))" by auto also have "... sums (exp (cnj z))" by (rule exp_converges) finally have "(\n. cnj (z ^ n /\<^sub>R (fact n))) sums (exp (cnj z))" . moreover have "(\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 lemma cnj_sin: "cnj(sin z) = sin(cnj z)" by (simp add: sin_exp_eq exp_cnj field_simps) lemma cnj_cos: "cnj(cos z) = cos(cnj z)" by (simp add: cos_exp_eq exp_cnj field_simps) lemma field_differentiable_at_sin: "sin field_differentiable at z" using DERIV_sin field_differentiable_def by blast lemma field_differentiable_within_sin: "sin field_differentiable (at z within S)" by (simp add: field_differentiable_at_sin field_differentiable_at_within) lemma field_differentiable_at_cos: "cos field_differentiable at z" using DERIV_cos field_differentiable_def by blast lemma field_differentiable_within_cos: "cos field_differentiable (at z within S)" by (simp add: field_differentiable_at_cos field_differentiable_at_within) lemma holomorphic_on_sin: "sin holomorphic_on S" by (simp add: field_differentiable_within_sin holomorphic_on_def) lemma holomorphic_on_cos: "cos holomorphic_on S" by (simp add: field_differentiable_within_cos holomorphic_on_def) lemma holomorphic_on_sin' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. sin (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def) lemma holomorphic_on_cos' [holomorphic_intros]: assumes "f holomorphic_on A" shows "(\x. cos (f x)) holomorphic_on A" using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def) subsection\<^marker>\tag unimportant\\More on the Polar Representation of Complex Numbers\ 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) lemma exp_eq_1: "exp z = 1 \ Re(z) = 0 \ (\n::int. Im(z) = of_int (2 * n) * pi)" (is "?lhs = ?rhs") proof assume "exp z = 1" then have "Re z = 0" by (metis exp_eq_one_iff norm_exp_eq_Re norm_one) with \?lhs\ 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) next assume ?rhs then show ?lhs using Im_exp Re_exp complex_eq_iff by (simp add: cos_one_2pi_int cos_one_sin_zero mult.commute) qed lemma exp_eq: "exp w = exp z \ (\n::int. w = z + (of_int (2 * n) * pi) * \)" (is "?lhs = ?rhs") proof - have "exp w = exp z \ exp (w-z) = 1" by (simp add: exp_diff) also have "... \ (Re w = Re z \ (\n::int. Im w - Im z = of_int (2 * n) * pi))" by (simp add: exp_eq_1) also have "... \ ?rhs" by (auto simp: algebra_simps intro!: complex_eqI) finally show ?thesis . qed lemma exp_complex_eqI: "\Im w - Im z\ < 2*pi \ exp w = exp z \ w = z" by (auto simp: exp_eq abs_mult) lemma exp_integer_2pi: assumes "n \ \" shows "exp((2 * n * pi) * \) = 1" proof - have "exp((2 * n * pi) * \) = exp 0" using assms unfolding Ints_def exp_eq by auto also have "... = 1" by simp finally show ?thesis . qed lemma exp_plus_2pin [simp]: "exp (z + \ * (of_int n * (of_real pi * 2))) = exp z" by (simp add: exp_eq) lemma exp_integer_2pi_plus1: assumes "n \ \" shows "exp(((2 * n + 1) * pi) * \) = - 1" proof - from assms obtain n' where [simp]: "n = of_int n'" by (auto simp: Ints_def) have "exp(((2 * n + 1) * pi) * \) = exp (pi * \)" using assms by (subst exp_eq) (auto intro!: exI[of _ n'] simp: algebra_simps) also have "... = - 1" by simp finally show ?thesis . qed lemma inj_on_exp_pi: fixes z::complex shows "inj_on exp (ball z pi)" proof (clarsimp simp: inj_on_def exp_eq) fix y n assume "dist z (y + 2 * of_int n * of_real pi * \) < pi" "dist z y < pi" then have "dist y (y + 2 * of_int n * of_real pi * \) < pi+pi" using dist_commute_lessI dist_triangle_less_add by blast then have "norm (2 * of_int n * of_real pi * \) < 2*pi" by (simp add: dist_norm) then show "n = 0" by (auto simp: norm_mult) qed lemma cmod_add_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) + r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\1 - \2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs") proof - have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)" by (rule complex_norm_square) also have "\ = (?z1 * cnj ?z1 + ?z2 * cnj ?z2) + (?z1 * cnj ?z2 + cnj ?z1 * ?z2)" by (simp add: algebra_simps) also have "\ = (norm ?z1)\<^sup>2 + (norm ?z2)\<^sup>2 + 2 * Re (?z1 * cnj ?z2)" unfolding complex_norm_square [symmetric] cnj_add_mult_eq_Re by simp also have "\ = ?rhs" by (simp add: norm_mult) (simp add: exp_Euler complex_is_Real_iff [THEN iffD1] cos_diff algebra_simps) finally show ?thesis using of_real_eq_iff by blast qed lemma cmod_diff_squared: fixes r1 r2::real assumes "r1 \ 0" "r2 \ 0" shows "(cmod (r1 * exp (\ * \1) - r2 * exp (\ * \2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\1 - \2)" (is "(cmod (?z1 - ?z2))\<^sup>2 = ?rhs") proof - have "exp (\ * (\2 + pi)) = - exp (\ * \2)" by (simp add: exp_Euler cos_plus_pi sin_plus_pi) then have "(cmod (?z1 - ?z2))\<^sup>2 = cmod (?z1 + r2 * exp (\ * (\2 + pi))) ^2" by simp also have "\ = r1\<^sup>2 + r2\<^sup>2 + 2*r1*r2*cos (\1 - (\2 + pi))" using assms cmod_add_squared by blast also have "\ = ?rhs" by (simp add: add.commute diff_add_eq_diff_diff_swap) finally show ?thesis . qed lemma polar_convergence: fixes R::real assumes "\j. r j > 0" "R > 0" shows "((\j. r j * exp (\ * \ j)) \ (R * exp (\ * \))) \ (r \ R) \ (\k. (\j. \ j - of_int (k j) * (2 * pi)) \ \)" (is "(?z \ ?Z) = ?rhs") proof assume L: "?z \ ?Z" have rR: "r \ R" using tendsto_norm [OF L] assms by (auto simp: norm_mult abs_of_pos) moreover obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - have "cos (\ j - \) = ((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) moreover have "(\j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \ ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))" by (intro L rR tendsto_intros) (use \R > 0\ in force) moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1" using \R > 0\ by (simp add: power2_eq_square field_split_simps) ultimately have "(\j. cos (\ j - \)) \ 1" by auto then show ?thesis using that cos_diff_limit_1 by blast qed ultimately show ?rhs by metis next assume R: ?rhs show "?z \ ?Z" proof (rule tendsto_mult) show "(\x. complex_of_real (r x)) \ of_real R" using R by (auto simp: tendsto_of_real_iff) obtain k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" using R by metis then have "(\j. complex_of_real (\ j - of_int (k j) * (2 * pi))) \ of_real \" using tendsto_of_real_iff by force then have "(\j. exp (\ * of_real (\ j - of_int (k j) * (2 * pi)))) \ exp (\ * \)" using tendsto_mult [OF tendsto_const] isCont_exp isCont_tendsto_compose by blast moreover have "exp (\ * of_real (\ j - of_int (k j) * (2 * pi))) = exp (\ * \ j)" for j unfolding exp_eq by (rule_tac x="- k j" in exI) (auto simp: algebra_simps) ultimately show "(\j. exp (\ * \ j)) \ exp (\ * \)" by auto qed qed lemma sin_cos_eq_iff: "sin y = sin x \ cos y = cos x \ (\n::int. y = x + 2 * pi * n)" proof - { assume "sin y = sin x" "cos y = cos x" then have "cos (y-x) = 1" using cos_add [of y "-x"] by simp then have "\n::int. y-x = 2 * pi * n" using cos_one_2pi_int by auto } then show ?thesis apply (auto simp: sin_add cos_add) apply (metis add.commute diff_add_cancel) done qed lemma exp_i_ne_1: assumes "0 < x" "x < 2*pi" shows "exp(\ * of_real x) \ 1" proof assume "exp (\ * of_real x) = 1" then have "exp (\ * of_real x) = exp 0" by simp then obtain n where "\ * of_real x = (of_int (2 * n) * pi) * \" 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 lemma sin_eq_0: fixes z::complex shows "sin z = 0 \ (\n::int. z = of_real(n * pi))" by (simp add: sin_exp_eq exp_eq) lemma cos_eq_0: fixes z::complex shows "cos z = 0 \ (\n::int. z = of_real(n * pi) + of_real pi/2)" using sin_eq_0 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps) lemma cos_eq_1: fixes z::complex shows "cos z = 1 \ (\n::int. z = of_real(2 * n * pi))" proof - have "cos z = cos (2*(z/2))" by simp also have "... = 1 - 2 * sin (z/2) ^ 2" by (simp only: cos_double_sin) finally have [simp]: "cos z = 1 \ sin (z/2) = 0" by simp show ?thesis by (auto simp: sin_eq_0) qed lemma csin_eq_1: fixes z::complex shows "sin z = 1 \ (\n::int. z = of_real(2 * n * pi) + of_real pi/2)" using cos_eq_1 [of "z - of_real pi/2"] by (simp add: cos_diff algebra_simps) lemma csin_eq_minus1: fixes z::complex shows "sin z = -1 \ (\n::int. z = of_real(2 * n * pi) + 3/2*pi)" (is "_ = ?rhs") proof - have "sin z = -1 \ sin (-z) = 1" by (simp add: equation_minus_iff) also have "... \ (\n::int. -z = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\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) also have "... = ?rhs" apply safe apply (rule_tac [2] x="-(x+1)" in exI) apply (rule_tac x="-(x+1)" in exI) apply (simp_all add: algebra_simps) done finally show ?thesis . qed lemma ccos_eq_minus1: fixes z::complex shows "cos z = -1 \ (\n::int. z = of_real(2 * n * pi) + pi)" using csin_eq_1 [of "z - of_real pi/2"] by (simp add: sin_diff algebra_simps equation_minus_iff) lemma sin_eq_1: "sin x = 1 \ (\n::int. x = (2 * n + 1 / 2) * pi)" (is "_ = ?rhs") proof - have "sin x = 1 \ sin (complex_of_real x) = 1" by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)" by (simp only: csin_eq_1) also have "... \ (\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]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma sin_eq_minus1: "sin x = -1 \ (\n::int. x = (2*n + 3/2) * pi)" (is "_ = ?rhs") proof - have "sin x = -1 \ sin (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)" by (simp only: csin_eq_minus1) also have "... \ (\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]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma cos_eq_minus1: "cos x = -1 \ (\n::int. x = (2*n + 1) * pi)" (is "_ = ?rhs") proof - have "cos x = -1 \ cos (complex_of_real x) = -1" by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real) also have "... \ (\n::int. complex_of_real x = of_real(2 * n * pi) + pi)" by (simp only: ccos_eq_minus1) also have "... \ (\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]]) also have "... = ?rhs" by (auto simp: algebra_simps) finally show ?thesis . qed lemma dist_exp_i_1: "norm(exp(\ * of_real t) - 1) = 2 * \sin(t / 2)\" proof - have "sqrt (2 - cos t * 2) = 2 * \sin (t / 2)\" using cos_double_sin [of "t/2"] by (simp add: real_sqrt_mult) then show ?thesis by (simp add: exp_Euler cmod_def power2_diff sin_of_real cos_of_real algebra_simps) qed lemma sin_cx_2pi [simp]: "\z = of_int m; even m\ \ sin (z * complex_of_real pi) = 0" by (simp add: sin_eq_0) lemma cos_cx_2pi [simp]: "\z = of_int m; even m\ \ cos (z * complex_of_real pi) = 1" using cos_eq_1 by auto lemma complex_sin_eq: fixes w :: complex shows "sin w = sin z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real((2*n + 1)*pi))" (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 then show ?rhs proof cases case 1 then show ?thesis by (simp add: sin_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) next case 2 then show ?thesis by (simp add: cos_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq) qed next assume ?rhs then consider n::int where "w = z + of_real (2 * of_int n * pi)" | n::int where " w = -z + of_real ((2 * of_int n + 1) * pi)" using Ints_cases by blast then show ?lhs proof cases case 1 then show ?thesis using Periodic_Fun.sin.plus_of_int [of z n] by (auto simp: algebra_simps) next case 2 then show ?thesis using Periodic_Fun.sin.plus_of_int [of "-z" "n"] apply (simp add: algebra_simps) by (metis add.commute add.inverse_inverse add_diff_cancel_left diff_add_cancel sin_plus_pi) qed qed lemma complex_cos_eq: fixes w :: complex shows "cos w = cos z \ (\n \ \. w = z + of_real(2*n*pi) \ w = -z + of_real(2*n*pi))" (is "?lhs = ?rhs") 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 then show ?rhs proof cases case 1 then obtain n where "w + z = of_int n * (complex_of_real pi * 2)" by (auto simp: sin_eq_0 algebra_simps) then have "w = -z + of_real(2 * of_int n * pi)" by (auto simp: algebra_simps) then show ?thesis using Ints_of_int by blast next case 2 then obtain n where "z = w + of_int n * (complex_of_real pi * 2)" by (auto simp: sin_eq_0 algebra_simps) then have "w = z + complex_of_real (2 * of_int(-n) * pi)" by (auto simp: algebra_simps) then show ?thesis using Ints_of_int by blast qed next assume ?rhs then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \ w = -z + of_real(2*n*pi)" using Ints_cases by (metis of_int_mult of_int_numeral) then show ?lhs using Periodic_Fun.cos.plus_of_int [of z n] apply (simp add: algebra_simps) by (metis cos.plus_of_int cos_minus minus_add_cancel mult.commute) qed lemma sin_eq: "sin x = sin y \ (\n \ \. x = y + 2*n*pi \ x = -y + (2*n + 1)*pi)" using complex_sin_eq [of x y] by (simp only: sin_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) lemma cos_eq: "cos x = cos y \ (\n \ \. x = y + 2*n*pi \ x = -y + 2*n*pi)" using complex_cos_eq [of x y] by (simp only: cos_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff) lemma sinh_complex: fixes z :: complex shows "(exp z - inverse (exp z)) / 2 = -\ * sin(\ * z)" by (simp add: sin_exp_eq field_split_simps exp_minus) lemma sin_i_times: fixes z :: complex shows "sin(\ * z) = \ * ((exp z - inverse (exp z)) / 2)" using sinh_complex by auto lemma sinh_real: fixes x :: real shows "of_real((exp x - inverse (exp x)) / 2) = -\ * sin(\ * of_real x)" by (simp add: exp_of_real sin_i_times) lemma cosh_complex: fixes z :: complex shows "(exp z + inverse (exp z)) / 2 = cos(\ * z)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemma cosh_real: fixes x :: real shows "of_real((exp x + inverse (exp x)) / 2) = cos(\ * of_real x)" by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real) lemmas cos_i_times = cosh_complex [symmetric] lemma norm_cos_squared: "norm(cos z) ^ 2 = cos(Re z) ^ 2 + (exp(Im z) - inverse(exp(Im z))) ^ 2 / 4" proof (cases z) case (Complex x1 x2) then show ?thesis apply (simp only: cos_add cmod_power2 cos_of_real sin_of_real Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib sin_squared_eq) apply (simp add: power2_eq_square field_split_simps) done qed lemma norm_sin_squared: "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4" proof (cases z) case (Complex x1 x2) then show ?thesis apply (simp only: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double Complex_eq) apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide) apply (simp only: left_diff_distrib [symmetric] power_mult_distrib cos_squared_eq) apply (simp add: power2_eq_square field_split_simps) done qed lemma exp_uminus_Im: "exp (- Im z) \ exp (cmod z)" using abs_Im_le_cmod linear order_trans by fastforce lemma norm_cos_le: fixes z::complex shows "norm(cos z) \ exp(norm z)" proof - have "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto then have "exp (- Im z) + exp (Im z) \ exp (cmod z) * 2" by (metis exp_uminus_Im add_mono exp_le_cancel_iff mult_2_right) then show ?thesis by (force simp add: cos_exp_eq norm_divide intro: order_trans [OF norm_triangle_ineq]) qed lemma norm_cos_plus1_le: fixes z::complex shows "norm(1 + cos z) \ 2 * exp(norm z)" proof - have mono: "\u w z::real. (1 \ w | 1 \ z) \ (w \ u & z \ u) \ 2 + w + z \ 4 * u" by arith have *: "Im z \ cmod z" using abs_Im_le_cmod abs_le_D1 by auto have triangle3: "\x y z. norm(x + y + z) \ norm(x) + norm(y) + norm(z)" by (simp add: norm_add_rule_thm) have "norm(1 + cos z) = cmod (1 + (exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: cos_exp_eq) also have "... = cmod ((2 + exp (\ * z) + exp (- (\ * z))) / 2)" by (simp add: field_simps) also have "... = cmod (2 + exp (\ * z) + exp (- (\ * 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 lemma sinh_conv_sin: "sinh z = -\ * sin (\*z)" by (simp add: sinh_field_def sin_i_times exp_minus) lemma cosh_conv_cos: "cosh z = cos (\*z)" by (simp add: cosh_field_def cos_i_times exp_minus) lemma tanh_conv_tan: "tanh z = -\ * tan (\*z)" by (simp add: tanh_def sinh_conv_sin cosh_conv_cos tan_def) lemma sin_conv_sinh: "sin z = -\ * sinh (\*z)" by (simp add: sinh_conv_sin) lemma cos_conv_cosh: "cos z = cosh (\*z)" by (simp add: cosh_conv_cos) lemma tan_conv_tanh: "tan z = -\ * tanh (\*z)" by (simp add: tan_def sin_conv_sinh cos_conv_cosh tanh_def) lemma sinh_complex_eq_iff: "sinh (z :: complex) = sinh w \ (\n\\. z = w - 2 * \ * of_real n * of_real pi \ z = -(2 * complex_of_real n + 1) * \ * complex_of_real pi - w)" (is "_ = ?rhs") proof - have "sinh z = sinh w \ sin (\ * z) = sin (\ * w)" by (simp add: sinh_conv_sin) also have "\ \ ?rhs" by (subst complex_sin_eq) (force simp: field_simps complex_eq_iff) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Taylor series for complex exponential, sine and cosine\ declare power_Suc [simp del] lemma Taylor_exp_field: fixes z::"'a::{banach,real_normed_field}" shows "norm (exp z - (\i\n. z ^ i / fact i)) \ exp (norm z) * (norm z ^ Suc n) / fact n" proof (rule field_Taylor[of _ n "\k. exp" "exp (norm z)" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume x: "x \ closed_segment 0 z" have "norm (exp x) \ exp (norm x)" by (rule norm_exp) also have "norm x \ norm z" using x by (auto simp: closed_segment_def intro!: mult_left_le_one_le) finally show "norm (exp x) \ exp (norm z)" by simp qed auto lemma Taylor_exp: "norm(exp z - (\k\n. z ^ k / (fact k))) \ exp\Re z\ * (norm z) ^ (Suc n) / (fact n)" proof (rule complex_Taylor [of _ n "\k. exp" "exp\Re z\" 0 z, simplified]) show "convex (closed_segment 0 z)" by (rule convex_closed_segment [of 0 z]) next fix k x assume "x \ closed_segment 0 z" "k \ n" show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)" using DERIV_exp DERIV_subset by blast next fix x assume "x \ closed_segment 0 z" then obtain u where u: "x = complex_of_real u * z" "0 \ u" "u \ 1" by (auto simp: closed_segment_def scaleR_conv_of_real) then have "u * Re z \ \Re z\" by (metis abs_ge_self abs_ge_zero mult.commute mult.right_neutral mult_mono) then show "Re x \ \Re z\" by (simp add: u) qed auto lemma assumes "0 \ u" "u \ 1" shows cmod_sin_le_exp: "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" and cmod_cos_le_exp: "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" proof - have mono: "\u w z::real. w \ u \ z \ u \ (w + z)/2 \ u" by simp have *: "(cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2 \ exp \Im z\" proof (rule mono) show "cmod (exp (\ * (u * z))) \ exp \Im z\" using assms by (auto simp: abs_if mult_left_le_one_le not_less intro: order_trans [of _ 0]) show "cmod (exp (- (\ * (u * z)))) \ exp \Im z\" using assms by (auto simp: abs_if mult_left_le_one_le mult_nonneg_nonpos intro: order_trans [of _ 0]) qed have "cmod (sin (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) - exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq4) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (sin (u *\<^sub>R z)) \ exp \Im z\" . have "cmod (cos (u *\<^sub>R z)) = cmod (exp (\ * (u * z)) + exp (- (\ * (u * z)))) / 2" by (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide) also have "... \ (cmod (exp (\ * (u * z))) + cmod (exp (- (\ * (u * z)))) ) / 2" by (intro divide_right_mono norm_triangle_ineq) simp also have "... \ exp \Im z\" by (rule *) finally show "cmod (cos (u *\<^sub>R z)) \ exp \Im z\" . qed lemma Taylor_sin: "norm(sin z - (\k\n. complex_of_real (sin_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ (Suc n) / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (sin z - (\i\n. (-1) ^ (i div 2) * (if even i then sin 0 else cos 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(k div 2) * (if even k then sin x else cos x)" "exp\Im z\" 0 z, simplified]) fix k x show "((\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 next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ (Suc n div 2) * (if odd n then sin x else cos x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (sin_coeff k) * z ^ k = (-1)^(k div 2) * (if even k then sin 0 else cos 0) * z^k / of_nat (fact k)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis by (simp add: ** order_trans [OF _ *]) qed lemma Taylor_cos: "norm(cos z - (\k\n. complex_of_real (cos_coeff k) * z ^ k)) \ exp\Im z\ * (norm z) ^ Suc n / (fact n)" proof - have mono: "\u w z::real. w \ u \ z \ u \ w + z \ u*2" by arith have *: "cmod (cos z - (\i\n. (-1) ^ (Suc i div 2) * (if even i then cos 0 else sin 0) * z ^ i / (fact i))) \ exp \Im z\ * cmod z ^ Suc n / (fact n)" proof (rule complex_Taylor [of "closed_segment 0 z" n "\k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" "exp\Im z\" 0 z, simplified]) fix k x assume "x \ closed_segment 0 z" "k \ n" show "((\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 next fix x assume "x \ closed_segment 0 z" then show "cmod ((- 1) ^ Suc (n div 2) * (if odd n then cos x else sin x)) \ exp \Im z\" by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp) qed have **: "\k. complex_of_real (cos_coeff k) * z ^ k = (-1)^(Suc k div 2) * (if even k then cos 0 else sin 0) * z^k / of_nat (fact k)" by (auto simp: cos_coeff_def elim!: evenE) show ?thesis by (simp add: ** order_trans [OF _ *]) qed declare power_Suc [simp] text\32-bit Approximation to e\ lemma e_approx_32: "\exp(1) - 5837465777 / 2147483648\ \ (inverse(2 ^ 32)::real)" using Taylor_exp [of 1 14] exp_le apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral) apply (simp only: pos_le_divide_eq [symmetric]) done lemma e_less_272: "exp 1 < (272/100::real)" using e_approx_32 by (simp add: abs_if split: if_split_asm) lemma ln_272_gt_1: "ln (272/100) > (1::real)" by (metis e_less_272 exp_less_cancel_iff exp_ln_iff less_trans ln_exp) text\Apparently redundant. But many arguments involve integers.\ lemma ln3_gt_1: "ln 3 > (1::real)" by (simp add: less_trans [OF ln_272_gt_1]) subsection\The argument of a complex number (HOL Light version)\ definition\<^marker>\tag important\ is_Arg :: "[complex,real] \ bool" where "is_Arg z r \ z = of_real(norm z) * exp(\ * of_real r)" definition\<^marker>\tag important\ Arg2pi :: "complex \ real" where "Arg2pi z \ if z = 0 then 0 else THE t. 0 \ t \ t < 2*pi \ is_Arg z t" lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \ is_Arg z r" 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 \ 0" shows "r = s" proof - have zr: "z = (cmod z) * exp (\ * r)" and zs: "z = (cmod z) * exp (\ * s)" using r s by (auto simp: is_Arg_def) with \z \ 0\ have eq: "exp (\ * r) = exp (\ * s)" by (metis mult_eq_0_iff mult_left_cancel) have "\ * r = \ * s" by (rule exp_complex_eqI) (use rs in \auto simp: eq exp_complex_eqI\) then show ?thesis by simp qed text\This function returns the angle of a complex number from its representation in polar coordinates. Due to periodicity, its range is arbitrary. \<^term>\Arg2pi\ follows HOL Light in adopting the interval \[0,2\)\. But we have the same periodicity issue with logarithms, and it is usual to adopt the same interval for the complex logarithm and argument functions. Further on down, we shall define both functions for the interval \(-\,\]\. The present version is provided for compatibility.\ lemma Arg2pi_0 [simp]: "Arg2pi(0) = 0" by (simp add: Arg2pi_def) lemma Arg2pi_unique_lemma: assumes z: "is_Arg z t" and z': "is_Arg z t'" and t: "0 \ t" "t < 2*pi" and t': "0 \ t'" "t' < 2*pi" and nz: "z \ 0" shows "t' = t" proof - have [dest]: "\x y z::real. x\0 \ x+y < z \ y * of_real t') = of_real (cmod z) * exp (\ * of_real t)" by (metis z z' is_Arg_def) then have "exp (\ * of_real t') = exp (\ * of_real t)" by (metis nz mult_left_cancel mult_zero_left z is_Arg_def) then have "sin t' = sin t \ 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 \auto simp: mult_less_0_iff algebra_simps\) then show "t' = t" by (simp add: n) qed lemma Arg2pi: "0 \ Arg2pi z \ Arg2pi z < 2*pi \ is_Arg z (Arg2pi z)" proof (cases "z=0") case True then show ?thesis by (simp add: Arg2pi_def is_Arg_def) next case False obtain t where t: "0 \ t" "t < 2*pi" 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" 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) show ?thesis apply (simp add: Arg2pi_def False) apply (rule theI [where a=t]) using t z False apply (auto intro: Arg2pi_unique_lemma) done qed corollary\<^marker>\tag unimportant\ shows Arg2pi_ge_0: "0 \ Arg2pi z" and Arg2pi_lt_2pi: "Arg2pi z < 2*pi" and Arg2pi_eq: "z = of_real(norm z) * exp(\ * of_real(Arg2pi z))" using Arg2pi is_Arg_def by auto lemma complex_norm_eq_1_exp: "norm z = 1 \ exp(\ * of_real (Arg2pi z)) = z" by (metis Arg2pi_eq cis_conv_exp mult.left_neutral norm_cis of_real_1) lemma Arg2pi_unique: "\of_real r * exp(\ * of_real a) = z; 0 < r; 0 \ a; a < 2*pi\ \ Arg2pi z = a" by (rule Arg2pi_unique_lemma [unfolded is_Arg_def, OF _ Arg2pi_eq]) (use Arg2pi [of z] in \auto simp: norm_mult\) lemma cos_Arg2pi: "cmod z * cos (Arg2pi z) = Re z" and sin_Arg2pi: "cmod z * sin (Arg2pi z) = Im z" using Arg2pi_eq [of z] cis_conv_exp Re_rcis Im_rcis unfolding rcis_def by metis+ lemma Arg2pi_minus: assumes "z \ 0" shows "Arg2pi (-z) = (if Arg2pi z < pi then Arg2pi z + pi else Arg2pi z - pi)" apply (rule Arg2pi_unique [of "norm z", OF complex_eqI]) using cos_Arg2pi sin_Arg2pi Arg2pi_ge_0 Arg2pi_lt_2pi [of z] assms by (auto simp: Re_exp Im_exp) lemma Arg2pi_times_of_real [simp]: assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z" proof (cases "z=0") case False show ?thesis by (rule Arg2pi_unique [of "r * norm z"]) (use Arg2pi False assms is_Arg_def in auto) qed auto lemma Arg2pi_times_of_real2 [simp]: "0 < r \ Arg2pi (z * of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real mult.commute) lemma Arg2pi_divide_of_real [simp]: "0 < r \ Arg2pi (z / of_real r) = Arg2pi z" by (metis Arg2pi_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg2pi_le_pi: "Arg2pi z \ pi \ 0 \ Im z" proof (cases "z=0") case False have "0 \ Im z \ 0 \ Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 \ Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z \ pi" by (simp add: Im_exp) (metis Arg2pi_ge_0 Arg2pi_lt_2pi sin_lt_zero sin_ge_zero not_le) finally show ?thesis by blast qed auto lemma Arg2pi_lt_pi: "0 < Arg2pi z \ Arg2pi z < pi \ 0 < Im z" proof (cases "z=0") case False have "0 < Im z \ 0 < Im (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... = (0 < Im (exp (\ * complex_of_real (Arg2pi z))))" using False by (simp add: zero_less_mult_iff) also have "... \ 0 < Arg2pi z \ Arg2pi z < pi" (is "_ = ?rhs") proof - have "0 < sin (Arg2pi z) \ ?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 lemma Arg2pi_eq_0: "Arg2pi z = 0 \ z \ \ \ 0 \ Re z" proof (cases "z=0") case False have "z \ \ \ 0 \ Re z \ z \ \ \ 0 \ Re (of_real (cmod z) * exp (\ * complex_of_real (Arg2pi z)))" by (metis Arg2pi_eq) also have "... \ z \ \ \ 0 \ Re (exp (\ * complex_of_real (Arg2pi z)))" using False by (simp add: zero_le_mult_iff) also have "... \ Arg2pi z = 0" proof - have [simp]: "Arg2pi z = 0 \ z \ \" using Arg2pi_eq [of z] by (auto simp: Reals_def) moreover have "\z \ \; 0 \ cos (Arg2pi z)\ \ 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) ultimately show ?thesis by (auto simp: Re_exp) qed finally show ?thesis by blast qed auto corollary\<^marker>\tag unimportant\ Arg2pi_gt_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg2pi z > 0" using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order unfolding nonneg_Reals_def by fastforce lemma Arg2pi_eq_pi: "Arg2pi z = pi \ z \ \ \ Re z < 0" using Arg2pi_le_pi [of z] Arg2pi_lt_pi [of z] Arg2pi_eq_0 [of z] Arg2pi_ge_0 [of z] by (fastforce simp: complex_is_Real_iff) lemma Arg2pi_eq_0_pi: "Arg2pi z = 0 \ Arg2pi z = pi \ z \ \" using Arg2pi_eq_0 Arg2pi_eq_pi not_le by auto lemma Arg2pi_of_real: "Arg2pi (of_real r) = (if r<0 then pi else 0)" using Arg2pi_eq_0_pi Arg2pi_eq_pi by fastforce lemma Arg2pi_real: "z \ \ \ Arg2pi z = (if 0 \ Re z then 0 else pi)" using Arg2pi_eq_0 Arg2pi_eq_0_pi by auto lemma Arg2pi_inverse: "Arg2pi(inverse z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False show ?thesis apply (rule Arg2pi_unique [of "inverse (norm z)"]) using Arg2pi_eq False Arg2pi_ge_0 [of z] Arg2pi_lt_2pi [of z] Arg2pi_eq_0 [of z] by (auto simp: Arg2pi_real in_Reals_norm exp_diff field_simps) qed auto lemma Arg2pi_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg2pi w = Arg2pi z \ (\x. 0 < x & w = of_real x * z)" using assms Arg2pi_eq [of z] Arg2pi_eq [of w] apply auto apply (rule_tac x="norm w / norm z" in exI) apply (simp add: field_split_simps) by (metis mult.commute mult.left_commute) lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \ Arg2pi z = 0" by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq) lemma Arg2pi_divide: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi(z / w) = Arg2pi z - Arg2pi w" apply (rule Arg2pi_unique [of "norm(z / w)"]) using assms Arg2pi_eq Arg2pi_ge_0 [of w] Arg2pi_lt_2pi [of z] apply (auto simp: exp_diff norm_divide field_simps) done lemma Arg2pi_le_div_sum: assumes "w \ 0" "z \ 0" "Arg2pi w \ Arg2pi z" shows "Arg2pi z = Arg2pi w + Arg2pi(z / w)" by (simp add: Arg2pi_divide assms) lemma Arg2pi_le_div_sum_eq: assumes "w \ 0" "z \ 0" shows "Arg2pi w \ Arg2pi z \ Arg2pi z = Arg2pi w + Arg2pi(z / w)" using assms by (auto simp: Arg2pi_ge_0 intro: Arg2pi_le_div_sum) lemma Arg2pi_diff: assumes "w \ 0" "z \ 0" shows "Arg2pi w - Arg2pi z = (if Arg2pi z \ Arg2pi w then Arg2pi(w / z) else Arg2pi(w/z) - 2*pi)" using assms Arg2pi_divide Arg2pi_inverse [of "w/z"] Arg2pi_eq_0_pi by (force simp add: Arg2pi_ge_0 Arg2pi_divide not_le split: if_split_asm) lemma Arg2pi_add: assumes "w \ 0" "z \ 0" shows "Arg2pi w + Arg2pi z = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi(w * z) else Arg2pi(w * z) + 2*pi)" using assms Arg2pi_diff [of "w*z" z] Arg2pi_le_div_sum_eq [of z "w*z"] apply (auto simp: Arg2pi_ge_0 Arg2pi_divide not_le) apply (metis Arg2pi_lt_2pi add.commute) apply (metis (no_types) Arg2pi add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less) done lemma Arg2pi_times: assumes "w \ 0" "z \ 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 lemma Arg2pi_cnj_eq_inverse: "z\0 \ Arg2pi (cnj z) = Arg2pi (inverse z)" apply (simp add: Arg2pi_eq_iff field_split_simps complex_norm_square [symmetric]) by (metis of_real_power zero_less_norm_iff zero_less_power) lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" proof (cases "z=0") case False then show ?thesis by (simp add: Arg2pi_cnj_eq_inverse Arg2pi_inverse) qed auto lemma Arg2pi_exp: "0 \ Im z \ Im z < 2*pi \ Arg2pi(exp z) = Im z" by (rule Arg2pi_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma complex_split_polar: obtains r a::real where "z = complex_of_real r * (cos a + \ * sin a)" "0 \ r" "0 \ a" "a < 2*pi" using Arg2pi cis.ctr cis_conv_exp unfolding Complex_eq is_Arg_def by fastforce lemma Re_Im_le_cmod: "Im w * sin \ + Re w * cos \ \ cmod w" proof (cases w rule: complex_split_polar) case (1 r a) with sin_cos_le1 [of a \] show ?thesis apply (simp add: norm_mult cmod_unit_one) by (metis (no_types, hide_lams) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le) qed subsection\<^marker>\tag unimportant\\Analytic properties of tangent function\ lemma cnj_tan: "cnj(tan z) = tan(cnj z)" by (simp add: cnj_cos cnj_sin tan_def) lemma field_differentiable_at_tan: "cos z \ 0 \ tan field_differentiable at z" unfolding field_differentiable_def using DERIV_tan by blast lemma field_differentiable_within_tan: "cos z \ 0 \ tan field_differentiable (at z within s)" using field_differentiable_at_tan field_differentiable_at_within by blast lemma continuous_within_tan: "cos z \ 0 \ continuous (at z within s) tan" using continuous_at_imp_continuous_within isCont_tan by blast lemma continuous_on_tan [continuous_intros]: "(\z. z \ s \ cos z \ 0) \ continuous_on s tan" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_tan: "(\z. z \ s \ cos z \ 0) \ tan holomorphic_on s" by (simp add: field_differentiable_within_tan holomorphic_on_def) subsection\The principal branch of the Complex logarithm\ instantiation complex :: ln begin definition\<^marker>\tag important\ ln_complex :: "complex \ complex" where "ln_complex \ \z. THE w. exp w = z & -pi < Im(w) & Im(w) \ pi" text\NOTE: within this scope, the constant Ln is not yet available!\ lemma assumes "z \ 0" shows exp_Ln [simp]: "exp(ln z) = z" and mpi_less_Im_Ln: "-pi < Im(ln z)" and Im_Ln_le_pi: "Im(ln z) \ pi" proof - obtain \ where z: "z / (cmod z) = Complex (cos \) (sin \)" using complex_unimodular_polar [of "z / (norm z)"] assms by (auto simp: norm_divide field_split_simps) obtain \ where \: "- pi < \" "\ \ pi" "sin \ = sin \" "cos \ = cos \" using sincos_principal_value [of "\"] assms by (auto simp: norm_divide field_split_simps) have "exp(ln z) = z & -pi < Im(ln z) & Im(ln z) \ pi" unfolding ln_complex_def apply (rule theI [where a = "Complex (ln(norm z)) \"]) using z assms \ apply (auto simp: field_simps exp_complex_eqI exp_eq_polar cis.code) done then show "exp(ln z) = z" "-pi < Im(ln z)" "Im(ln z) \ pi" by auto qed lemma Ln_exp [simp]: assumes "-pi < Im(z)" "Im(z) \ pi" shows "ln(exp z) = z" proof (rule exp_complex_eqI) show "\Im (ln (exp z)) - Im z\ < 2 * pi" using assms mpi_less_Im_Ln [of "exp z"] Im_Ln_le_pi [of "exp z"] by auto qed auto subsection\<^marker>\tag unimportant\\Relation to Real Logarithm\ 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 "... = of_real(ln z)" using assms by (subst Ln_exp) auto finally show ?thesis using assms by simp qed corollary\<^marker>\tag unimportant\ Ln_in_Reals [simp]: "z \ \ \ Re z > 0 \ ln z \ \" by (auto simp: Ln_of_real elim: Reals_cases) corollary\<^marker>\tag unimportant\ Im_Ln_of_real [simp]: "r > 0 \ Im (ln (of_real r)) = 0" by (simp add: Ln_of_real) lemma cmod_Ln_Reals [simp]: "z \ \ \ 0 < Re z \ cmod (ln z) = norm (ln (Re z))" using Ln_of_real by force lemma Ln_Reals_eq: "\x \ \; Re x > 0\ \ ln x = of_real (ln (Re x))" 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 lemma Ln_eq_zero_iff [simp]: "x \ \\<^sub>\\<^sub>0 \ ln x = 0 \ x = 1" for x::complex by auto (metis exp_Ln exp_zero nonpos_Reals_zero_I) instance by intro_classes (rule ln_complex_def Ln_1) end abbreviation Ln :: "complex \ complex" where "Ln \ ln" lemma Ln_eq_iff: "w \ 0 \ z \ 0 \ (Ln w = Ln z \ w = z)" by (metis exp_Ln) lemma Ln_unique: "exp(z) = w \ -pi < Im(z) \ Im(z) \ pi \ Ln w = z" using Ln_exp by blast lemma Re_Ln [simp]: "z \ 0 \ Re(Ln z) = ln(norm z)" by (metis exp_Ln ln_exp norm_exp_eq_Re) corollary\<^marker>\tag unimportant\ ln_cmod_le: assumes z: "z \ 0" shows "ln (cmod z) \ 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>\tag unimportant\ exists_complex_root: fixes z :: complex assumes "n \ 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) corollary\<^marker>\tag unimportant\ exists_complex_root_nonzero: fixes z::complex assumes "z \ 0" "n \ 0" obtains w where "w \ 0" "z = w ^ n" by (metis exists_complex_root [of n z] assms power_0_left) subsection\<^marker>\tag unimportant\\Derivative of Ln away from the branch cut\ lemma assumes "z \ \\<^sub>\\<^sub>0" shows has_field_derivative_Ln: "(Ln has_field_derivative inverse(z)) (at z)" and Im_Ln_less_pi: "Im (Ln z) < pi" proof - have znz [simp]: "z \ 0" using assms by auto then have "Im (Ln z) \ 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) let ?U = "{w. -pi < Im(w) \ Im(w) < pi}" have 1: "open ?U" by (simp add: open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt) have 2: "\x. x \ ?U \ (exp has_derivative blinfun_apply(Blinfun ((*) (exp x)))) (at x)" by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right has_field_derivative_imp_has_derivative) have 3: "continuous_on ?U (\x. Blinfun ((*) (exp x)))" unfolding blinfun_mult_right.abs_eq [symmetric] by (intro continuous_intros) have 4: "Ln z \ ?U" by (auto simp: 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' \ ?U" and "Ln z \ U'" "open V" "z \ V" and hom: "homeomorphism U' V exp g" and g: "\y. y \ V \ (g has_derivative (g' y)) (at y)" and g': "\y. y \ V \ g' y = inv ((*) (exp (g y)))" and bij: "\y. y \ V \ bij ((*) (exp (g y)))" using inverse_function_theorem [OF 1 2 3 4 5] by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) blast show "(Ln has_field_derivative inverse(z)) (at z)" unfolding has_field_derivative_def proof (rule has_derivative_transform_within_open) show g_eq_Ln: "g y = Ln y" if "y \ V" for y proof - obtain x where "y = exp x" "x \ U'" using hom homeomorphism_image1 that \y \ V\ by blast then show ?thesis using sub hom homeomorphism_apply1 by fastforce qed have "0 \ V" by (meson exp_not_eq_zero hom homeomorphism_def) then have "\y. y \ V \ g' y = inv ((*) y)" by (metis exp_Ln g' g_eq_Ln) then have g': "g' z = (\x. x/z)" by (metis (no_types, hide_lams) bij \z \ V\ 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 \z \ V\] g' by (simp add: \z \ V\ field_class.field_divide_inverse has_derivative_imp_has_field_derivative has_field_derivative_imp_has_derivative) qed (auto simp: \z \ V\ \open V\) qed declare has_field_derivative_Ln [derivative_intros] declare has_field_derivative_Ln [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable at z" using field_differentiable_def has_field_derivative_Ln by blast lemma field_differentiable_within_Ln: "z \ \\<^sub>\\<^sub>0 \ Ln field_differentiable (at z within S)" using field_differentiable_at_Ln field_differentiable_within_subset by blast lemma continuous_at_Ln: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) Ln" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Ln) lemma isCont_Ln' [simp,continuous_intros]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. Ln (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_Ln]) lemma continuous_within_Ln [continuous_intros]: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Ln" using continuous_at_Ln continuous_at_imp_continuous_within by blast lemma continuous_on_Ln [continuous_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ continuous_on S Ln" by (simp add: continuous_at_imp_continuous_on continuous_within_Ln) lemma continuous_on_Ln' [continuous_intros]: "continuous_on S f \ (\z. z \ S \ f z \ \\<^sub>\\<^sub>0) \ continuous_on S (\x. Ln (f x))" by (rule continuous_on_compose2[OF continuous_on_Ln, of "UNIV - nonpos_Reals" S f]) auto lemma holomorphic_on_Ln [holomorphic_intros]: "(\z. z \ S \ z \ \\<^sub>\\<^sub>0) \ Ln holomorphic_on S" by (simp add: field_differentiable_within_Ln holomorphic_on_def) lemma holomorphic_on_Ln' [holomorphic_intros]: "(\z. z \ A \ f z \ \\<^sub>\\<^sub>0) \ f holomorphic_on A \ (\z. Ln (f z)) holomorphic_on A" using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \\<^sub>\\<^sub>0"] by (auto simp: o_def) lemma tendsto_Ln [tendsto_intros]: fixes L F assumes "(f \ L) F" "L \ \\<^sub>\\<^sub>0" shows "((\x. Ln (f x)) \ Ln L) F" proof - have "nhds L \ filtermap f F" using assms(1) by (simp add: filterlim_def) moreover have "\\<^sub>F y in nhds L. y \ - \\<^sub>\\<^sub>0" using eventually_nhds_in_open[of "- \\<^sub>\\<^sub>0" L] assms by (auto simp: open_Compl) ultimately have "\\<^sub>F y in filtermap f F. y \ - \\<^sub>\\<^sub>0" by (rule filter_leD) moreover have "continuous_on (-\\<^sub>\\<^sub>0) Ln" by (rule continuous_on_Ln) auto ultimately show ?thesis using continuous_on_tendsto_compose[of "- \\<^sub>\\<^sub>0" Ln f L F] assms by (simp add: eventually_filtermap) qed lemma divide_ln_mono: fixes x y::real assumes "3 \ x" "x \ y" shows "x / ln x \ y / ln y" proof (rule exE [OF complex_mvt_line [of x y "\z. z / Ln z" "\z. 1/(Ln z) - 1/(Ln z)^2"]]; clarsimp simp add: closed_segment_Reals closed_segment_eq_real_ivl assms) show "\u. \x \ u; u \ y\ \ ((\z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)" using \3 \ x\ by (force intro!: derivative_eq_intros simp: field_simps power_eq_if) show "x / ln x \ 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 \ u" "u \ y" for u proof - have eq: "y / ln y = (1 / ln u - 1 / (ln u)\<^sup>2) * (y - x) + x / ln x" using that \3 \ x\ by (auto simp: Ln_Reals_eq in_Reals_norm group_add_class.diff_eq_eq) show ?thesis using exp_le \3 \ x\ x by (simp add: eq) (simp add: power_eq_if divide_simps ln_ge_iff) qed qed theorem Ln_series: fixes z :: complex assumes "norm z < 1" shows "(\n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\n. ?f n * z^n) sums _") proof - let ?F = "\z. \n. ?f n * z^n" and ?F' = "\z. \n. diffs ?f n * z^n" have r: "conv_radius ?f = 1" by (intro conv_radius_ratio_limit_nonzero[of _ 1]) (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc) have "\c. \z\ball 0 1. ln (1 + z) - ?F z = c" proof (rule has_field_derivative_zero_constant) fix z :: complex assume z': "z \ ball 0 1" hence z: "norm z < 1" by simp define t :: complex where "t = of_real (1 + norm z) / 2" from z have t: "norm z < norm t" "norm t < 1" unfolding t_def by (simp_all add: field_simps norm_divide del: of_real_add) have "Re (-z) \ norm (-z)" by (rule complex_Re_le_cmod) also from z have "... < 1" by simp finally have "((\z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)" by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff) moreover have "(?F has_field_derivative ?F' z) (at z)" using t r by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all ultimately have "((\z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z)) (at z within ball 0 1)" by (intro derivative_intros) (simp_all add: at_within_open[OF z']) also have "(\n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all from sums_split_initial_segment[OF this, of 1] have "(\i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc) hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse) also have "inverse (1 + z) - inverse (1 + z) = 0" by simp finally show "((\z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" . qed simp_all then obtain c where c: "\z. z \ ball 0 1 \ ln (1 + z) - ?F z = c" by blast from c[of 0] have "c = 0" by (simp only: powser_zero) simp with c[of z] assms have "ln (1 + z) = ?F z" by simp moreover have "summable (\n. ?f n * z^n)" using assms r by (intro summable_in_conv_radius) simp_all ultimately show ?thesis by (simp add: sums_iff) qed lemma Ln_series': "cmod z < 1 \ (\n. - ((-z)^n) / of_nat n) sums ln (1 + z)" by (drule Ln_series) (simp add: power_minus') lemma ln_series': assumes "abs (x::real) < 1" shows "(\n. - ((-x)^n) / of_nat n) sums ln (1 + x)" proof - from assms have "(\n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)" by (intro Ln_series') simp_all also have "(\n. - ((-of_real x)^n) / of_nat n) = (\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 finally show ?thesis by (subst (asm) sums_of_real_iff) qed lemma Ln_approx_linear: fixes z :: complex assumes "norm z < 1" shows "norm (ln (1 + z) - z) \ norm z^2 / (1 - norm z)" proof - let ?f = "\n. (-1)^Suc n / of_nat n" from assms have "(\n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp moreover have "(\n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp ultimately have "(\n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)" by (subst left_diff_distrib, intro sums_diff) simp_all from sums_split_initial_segment[OF this, of "Suc 1"] have "(\i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)" by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse) hence "(Ln (1 + z) - z) = (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)" by (simp add: sums_iff) also have A: "summable (\n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))" by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]]) (auto simp: assms field_simps intro!: always_eventually) hence "norm (\i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \ (\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))" by (intro summable_norm) (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc) also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \ norm ((-z)^2 * (-z)^i) * 1" for i by (intro mult_left_mono) (simp_all add: field_split_simps) hence "(\i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i)) \ (\i. norm (-(z^2) * (-z)^i))" using A assms unfolding norm_power norm_inverse norm_divide norm_mult apply (intro suminf_le summable_mult summable_geometric) apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc) done also have "... = norm z^2 * (\i. norm z^i)" using assms by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power) also have "(\i. norm z^i) = inverse (1 - norm z)" using assms by (subst suminf_geometric) (simp_all add: divide_inverse) also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse) finally show ?thesis . qed subsection\<^marker>\tag unimportant\\Quadrant-type results for Ln\ lemma cos_lt_zero_pi: "pi/2 < x \ x < 3*pi/2 \ cos x < 0" using cos_minus_pi cos_gt_zero_pi [of "x-pi"] by simp lemma Re_Ln_pos_lt: assumes "z \ 0" shows "\Im(Ln z)\ < pi/2 \ 0 < Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto have "\Im w\ < pi/2 \ 0 < Re(exp w)" proof assume "\Im w\ < 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 "\Im w\ \ pi/2" by (metis cos_minus cos_pi_half mult_eq_0_iff Re_exp abs_if order_less_irrefl) then show "\Im w\ < 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 \ 0" shows "\Im(Ln z)\ \ pi/2 \ 0 \ Re(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ pi" "- pi < Im w" using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms by auto then have "\Im w\ \ pi/2 \ 0 \ Re(exp w)" using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le by (auto simp: Re_exp zero_le_mult_iff abs_if intro: cos_ge_zero) } then show ?thesis using assms by auto qed lemma Im_Ln_pos_lt: assumes "z \ 0" shows "0 < Im(Ln z) \ Im(Ln z) < pi \ 0 < Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ 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 \ Im w < pi \ 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 lemma Im_Ln_pos_le: assumes "z \ 0" shows "0 \ Im(Ln z) \ Im(Ln z) \ pi \ 0 \ Im(z)" proof - { fix w assume "w = Ln z" then have w: "Im w \ 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 \ Im w \ pi \ 0 \ Im(exp w)" using sin_ge_zero [of "- (Im w)"] sin_ge_zero [of "abs(Im w)"] sin_zero_pi_iff [of "Im w"] by (force simp: Im_exp zero_le_mult_iff sin_ge_zero) } then show ?thesis using assms by auto qed lemma Re_Ln_pos_lt_imp: "0 < Re(z) \ \Im(Ln z)\ < pi/2" by (metis Re_Ln_pos_lt less_irrefl zero_complex.simps(1)) lemma Im_Ln_pos_lt_imp: "0 < Im(z) \ 0 < Im(Ln z) \ Im(Ln z) < pi" by (metis Im_Ln_pos_lt not_le order_refl zero_complex.simps(2)) text\A reference to the set of positive real numbers\ lemma Im_Ln_eq_0: "z \ 0 \ (Im(Ln z) = 0 \ 0 < Re(z) \ Im(z) = 0)" by (metis Im_complex_of_real Im_exp Ln_in_Reals Re_Ln_pos_lt Re_Ln_pos_lt_imp Re_complex_of_real complex_is_Real_iff exp_Ln exp_of_real pi_gt_zero) lemma Im_Ln_eq_pi: "z \ 0 \ (Im(Ln z) = pi \ Re(z) < 0 \ Im(z) = 0)" by (metis Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt add.left_neutral complex_eq less_eq_real_def mult_zero_right not_less_iff_gr_or_eq pi_ge_zero pi_neq_zero rcis_zero_arg rcis_zero_mod) subsection\<^marker>\tag unimportant\\More Properties of Ln\ lemma cnj_Ln: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ \ \Im (cnj (Ln z))\ + \Im (Ln (cnj z))\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (cnj (Ln z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (Ln (cnj z))\ \ 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 "\Im (cnj (Ln z)) - Im (Ln (cnj z))\ < 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 qed (use assms in auto) lemma Ln_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "Ln(inverse z) = -(Ln z)" proof (cases "z=0") case False show ?thesis proof (rule exp_complex_eqI) have "\Im (Ln (inverse z)) - Im (- Ln z)\ \ \Im (Ln (inverse z))\ + \Im (- Ln z)\" by (rule abs_triangle_ineq4) also have "... < pi + pi" proof - have "\Im (Ln (inverse z))\ < pi" by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln) moreover have "\Im (- Ln z)\ \ pi" using False Im_Ln_le_pi mpi_less_Im_Ln by fastforce ultimately show ?thesis by simp qed finally show "\Im (Ln (inverse z)) - Im (- Ln z)\ < 2 * pi" by simp show "exp (Ln (inverse z)) = exp (- Ln z)" by (simp add: False exp_minus) qed qed (use assms in auto) lemma Ln_minus1 [simp]: "Ln(-1) = \ * pi" proof (rule exp_complex_eqI) show "\Im (Ln (- 1)) - Im (\ * pi)\ < 2 * pi" using Im_Ln_le_pi [of "-1"] mpi_less_Im_Ln [of "-1"] by auto qed auto lemma Ln_ii [simp]: "Ln \ = \ * of_real pi/2" using Ln_exp [of "\ * (of_real pi/2)"] unfolding exp_Euler by simp lemma Ln_minus_ii [simp]: "Ln(-\) = - (\ * pi/2)" proof - have "Ln(-\) = Ln(inverse \)" by simp also have "... = - (Ln \)" using Ln_inverse by blast also have "... = - (\ * pi/2)" by simp finally show ?thesis . qed lemma Ln_times: assumes "w \ 0" "z \ 0" shows "Ln(w * z) = (if Im(Ln w + Ln z) \ -pi then (Ln(w) + Ln(z)) + \ * of_real(2*pi) else if Im(Ln w + Ln z) > pi then (Ln(w) + Ln(z)) - \ * of_real(2*pi) else Ln(w) + Ln(z))" using pi_ge_zero Im_Ln_le_pi [of w] Im_Ln_le_pi [of z] using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z] by (auto simp: exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique) corollary\<^marker>\tag unimportant\ Ln_times_simple: "\w \ 0; z \ 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \ pi\ \ Ln(w * z) = Ln(w) + Ln(z)" by (simp add: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_of_real: "\r > 0; z \ 0\ \ Ln(of_real r * z) = ln r + Ln(z)" using mpi_less_Im_Ln Im_Ln_le_pi by (force simp: Ln_times) corollary\<^marker>\tag unimportant\ Ln_times_Reals: "\r \ Reals; Re r > 0; z \ 0\ \ Ln(r * z) = ln (Re r) + Ln(z)" using Ln_Reals_eq Ln_times_of_real by fastforce corollary\<^marker>\tag unimportant\ Ln_divide_of_real: "\r > 0; z \ 0\ \ 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) corollary\<^marker>\tag unimportant\ Ln_prod: fixes f :: "'a \ complex" assumes "finite A" "\x. x \ A \ f x \ 0" shows "\n. Ln (prod f A) = (\x \ A. Ln (f x) + (of_int (n x) * (2 * pi)) * \)" using assms proof (induction A) case (insert x A) then obtain n where n: "Ln (prod f A) = (\x\A. Ln (f x) + of_real (of_int (n x) * (2 * pi)) * \)" by auto define D where "D \ Im (Ln (f x)) + Im (Ln (prod f A))" define q::int where "q \ (if D \ -pi then 1 else if D > pi then -1 else 0)" have "prod f A \ 0" "f x \ 0" by (auto simp: insert.hyps insert.prems) with insert.hyps pi_ge_zero show ?case by (rule_tac x="n(x:=q)" in exI) (force simp: Ln_times q_def D_def n intro!: sum.cong) qed auto lemma Ln_minus: assumes "z \ 0" shows "Ln(-z) = (if Im(z) \ 0 \ \(Re(z) < 0 \ Im(z) = 0) then Ln(z) + \ * pi else Ln(z) - \ * pi)" (is "_ = ?rhs") 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) lemma Ln_inverse_if: assumes "z \ 0" shows "Ln (inverse z) = (if z \ \\<^sub>\\<^sub>0 then -(Ln z) + \ * 2 * complex_of_real pi else -(Ln z))" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis by (simp add: Ln_inverse) next case True then have z: "Im z = 0" "Re z < 0" "- z \ \\<^sub>\\<^sub>0" using assms complex_eq_iff complex_nonpos_Reals_iff by auto have "Ln(inverse z) = Ln(- (inverse (-z)))" by simp also have "... = Ln (inverse (-z)) + \ * complex_of_real pi" using assms z by (simp add: Ln_minus divide_less_0_iff) also have "... = - Ln (- z) + \ * complex_of_real pi" using z Ln_inverse by presburger also have "... = - (Ln z) + \ * 2 * complex_of_real pi" using Ln_minus assms z by auto finally show ?thesis by (simp add: True) qed lemma Ln_times_ii: assumes "z \ 0" shows "Ln(\ * z) = (if 0 \ Re(z) | Im(z) < 0 then Ln(z) + \ * of_real pi/2 else Ln(z) - \ * of_real(3 * pi/2))" 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] Re_Ln_pos_le [of z] by (simp add: Ln_times) auto lemma Ln_of_nat [simp]: "0 < n \ 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 lemma Ln_of_nat_over_of_nat: assumes "m > 0" "n > 0" shows "Ln (of_nat m / of_nat n) = of_real (ln (of_nat m) - ln (of_nat n))" proof - have "of_nat m / of_nat n = (of_real (of_nat m / of_nat n) :: complex)" by simp also from assms have "Ln ... = of_real (ln (of_nat m / of_nat n))" by (simp add: Ln_of_real[symmetric]) also from assms have "... = of_real (ln (of_nat m) - ln (of_nat n))" by (simp add: ln_div) finally show ?thesis . qed subsection\The Argument of a Complex Number\ text\Unlike in HOL Light, it's defined for the same interval as the complex logarithm: \(-\,\]\.\ lemma Arg_eq_Im_Ln: - assumes "z \ 0" shows "arg z = Im (Ln z)" + assumes "z \ 0" shows "Arg z = Im (Ln z)" proof (rule arg_unique) show "sgn z = cis (Im (Ln z))" by (metis assms exp_Ln exp_eq_polar nonzero_mult_div_cancel_left norm_eq_zero norm_exp_eq_Re of_real_eq_0_iff sgn_eq) show "- pi < Im (Ln z)" by (simp add: assms mpi_less_Im_Ln) show "Im (Ln z) \ pi" by (simp add: Im_Ln_le_pi assms) qed text \The 1990s definition of argument coincides with the more recent one\ -lemma Arg_definition_equivalence: - shows "arg z = (if z = 0 then 0 else Im (Ln z))" - by (simp add: Arg_eq_Im_Ln arg_zero) - -definition\<^marker>\tag important\ Arg :: "complex \ real" where "Arg z \ (if z = 0 then 0 else Im (Ln z))" +lemma\<^marker>\tag important\ Arg_def: + shows "Arg z = (if z = 0 then 0 else Im (Ln z))" + by (simp add: Arg_eq_Im_Ln Arg_zero) lemma Arg_of_real: "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 \ pi" by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi) lemma assumes "z \ 0" shows Arg_eq: "z = of_real(norm z) * exp(\ * Arg z)" using assms exp_Ln exp_eq_polar by (auto simp: Arg_def) lemma is_Arg_Arg: "z \ 0 \ is_Arg z (Arg z)" by (simp add: Arg_eq is_Arg_def) lemma Argument_exists: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" proof - let ?rp = "r - Arg z + pi" define k where "k \ \?rp / (2 * pi)\" have "(Arg z + of_int k * (2 * pi)) \ R" using floor_divide_lower [of "2*pi" ?rp] floor_divide_upper [of "2*pi" ?rp] by (auto simp: k_def algebra_simps R) then show ?thesis using Arg_eq \z \ 0\ is_Arg_2pi_iff is_Arg_def that by blast qed lemma Argument_exists_unique: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" obtains s where "is_Arg z s" "s\R" "\t. \is_Arg z t; t\R\ \ s=t" proof - obtain s where s: "is_Arg z s" "s\R" using Argument_exists [OF assms] . moreover have "\t. \is_Arg z t; t\R\ \ s=t" using assms s by (auto simp: is_Arg_eqI) ultimately show thesis using that by blast qed lemma Argument_Ex1: assumes "z \ 0" and R: "R = {r-pi<..r+pi}" shows "\!s. is_Arg z s \ s \ R" using Argument_exists_unique [OF assms] by metis lemma Arg_divide: assumes "w \ 0" "z \ 0" shows "is_Arg (z / w) (Arg z - Arg w)" using Arg_eq [of z] Arg_eq [of w] Arg_eq [of "norm(z / w)"] assms 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 \ pi" and t': "- pi < t'" "t' \ pi" and nz: "z \ 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(\ * t)))" by (metis is_Arg_def z) also have "... = (cmod (- inverse z)) * exp (\ * (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(\ * t')))" by (metis is_Arg_def z') also have "... = (cmod (- inverse z)) * exp (\ * (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 lemma complex_norm_eq_1_exp_eq: "norm z = 1 \ exp(\ * (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) lemma Arg_unique: "\of_real r * exp(\ * a) = z; 0 < r; -pi < a; a \ pi\ \ Arg z = a" by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq]) (use mpi_less_Arg Arg_le_pi in \auto simp: norm_mult\) lemma Arg_minus: assumes "z \ 0" shows "Arg (-z) = (if Arg z \ 0 then Arg z + pi else Arg z - pi)" proof - have [simp]: "cmod z * cos (Arg z) = Re z" using assms Arg_eq [of z] by (metis Re_exp exp_Ln norm_exp_eq_Re Arg_def) 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) qed lemma Arg_times_of_real [simp]: assumes "0 < r" shows "Arg (of_real r * z) = Arg z" proof (cases "z=0") case True then show ?thesis by (simp add: Arg_def) next case False with Arg_eq assms show ?thesis by (auto simp: mpi_less_Arg Arg_le_pi intro!: Arg_unique [of "r * norm z"]) qed lemma Arg_times_of_real2 [simp]: "0 < r \ Arg (z * of_real r) = Arg z" by (metis Arg_times_of_real mult.commute) lemma Arg_divide_of_real [simp]: "0 < r \ Arg (z / of_real r) = Arg z" by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff) lemma Arg_less_0: "0 \ Arg z \ 0 \ Im z" using Im_Ln_le_pi Im_Ln_pos_le by (simp add: Arg_def) lemma Arg_eq_pi: "Arg z = pi \ Re z < 0 \ Im z = 0" by (auto simp: Arg_def Im_Ln_eq_pi) lemma Arg_lt_pi: "0 < Arg z \ Arg z < pi \ 0 < Im z" using Arg_less_0 [of z] Im_Ln_pos_lt by (auto simp: order.order_iff_strict Arg_def) lemma Arg_eq_0: "Arg z = 0 \ z \ \ \ 0 \ Re z" using complex_is_Real_iff by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left) corollary\<^marker>\tag unimportant\ Arg_ne_0: assumes "z \ \\<^sub>\\<^sub>0" shows "Arg z \ 0" using assms by (auto simp: nonneg_Reals_def Arg_eq_0) lemma Arg_eq_pi_iff: "Arg z = pi \ z \ \ \ Re z < 0" proof (cases "z=0") case False then show ?thesis using Arg_eq_0 [of "-z"] Arg_eq_pi complex_is_Real_iff by blast qed (simp add: Arg_def) lemma Arg_eq_0_pi: "Arg z = 0 \ Arg z = pi \ z \ \" using Arg_eq_pi_iff Arg_eq_0 by force lemma Arg_real: "z \ \ \ Arg z = (if 0 \ Re z then 0 else pi)" using Arg_eq_0 Arg_eq_0_pi by auto lemma Arg_inverse: "Arg(inverse z) = (if z \ \ then Arg z else - Arg z)" proof (cases "z \ \") case True then show ?thesis by simp (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse) next case False then have z: "Arg z < pi" "z \ 0" using Arg_eq_0_pi Arg_le_pi by (auto simp: less_eq_real_def) show ?thesis apply (rule Arg_unique [of "inverse (norm z)"]) using False z mpi_less_Arg [of z] Arg_eq [of z] by (auto simp: exp_minus field_simps) qed lemma Arg_eq_iff: assumes "w \ 0" "z \ 0" shows "Arg w = Arg z \ (\x. 0 < x \ w = of_real x * z)" (is "?lhs = ?rhs") proof assume ?lhs then have "w = complex_of_real (cmod w / cmod z) * z" by (metis Arg_eq assms divide_divide_eq_right eq_divide_eq exp_not_eq_zero of_real_divide) then show ?rhs using assms divide_pos_pos zero_less_norm_iff by blast qed auto lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \ Arg z = 0" by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq) lemma Arg_cnj_eq_inverse: "z\0 \ Arg (cnj z) = Arg (inverse z)" using Arg2pi_cnj_eq_inverse Arg2pi_eq_iff Arg_eq_iff by auto lemma Arg_cnj: "Arg(cnj z) = (if z \ \ then Arg z else - Arg z)" by (metis Arg_cnj_eq_inverse Arg_inverse Reals_0 complex_cnj_zero) lemma Arg_exp: "-pi < Im z \ Im z \ pi \ Arg(exp z) = Im z" by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar) lemma Ln_Arg: "z\0 \ Ln(z) = ln(norm z) + \ * Arg(z)" by (metis Arg_def Re_Ln complex_eq) lemma continuous_at_Arg: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg" proof - have [simp]: "(\z. Im (Ln z)) \z\ Arg z" using Arg_def assms continuous_at by fastforce show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\w. \w \ - \\<^sub>\\<^sub>0; w \ z\ \ Im (Ln w) = Arg w" by (metis Arg_def Compl_iff nonpos_Reals_zero_I) qed (use assms in auto) qed lemma continuous_within_Arg: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within S) Arg" using continuous_at_Arg continuous_at_imp_continuous_within by blast subsection\The Unwinding Number and the Ln product Formula\ text\Note that in this special case the unwinding number is -1, 0 or 1. But it's always an integer.\ lemma is_Arg_exp_Im: "is_Arg (exp z) (Im z)" using exp_eq_polar is_Arg_def norm_exp_eq_Re by auto lemma is_Arg_exp_diff_2pi: assumes "is_Arg (exp z) \" shows "\k. Im z - of_int k * (2 * pi) = \" proof (intro exI is_Arg_eqI) let ?k = "\(Im z - \) / (2 * pi)\" show "is_Arg (exp z) (Im z - real_of_int ?k * (2 * pi))" by (metis diff_add_cancel is_Arg_2pi_iff is_Arg_exp_Im) show "\Im z - real_of_int ?k * (2 * pi) - \\ < 2 * pi" using floor_divide_upper [of "2*pi" "Im z - \"] floor_divide_lower [of "2*pi" "Im z - \"] by (auto simp: algebra_simps abs_if) qed (auto simp: is_Arg_exp_Im assms) lemma Arg_exp_diff_2pi: "\k. Im z - of_int k * (2 * pi) = Arg (exp z)" using is_Arg_exp_diff_2pi [OF is_Arg_Arg] by auto lemma unwinding_in_Ints: "(z - Ln(exp z)) / (of_real(2*pi) * \) \ \" using Arg_exp_diff_2pi [of z] by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI) definition\<^marker>\tag important\ unwinding :: "complex \ int" where "unwinding z \ THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \)" lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \)" using unwinding_in_Ints [of z] unfolding unwinding_def Ints_def by force lemma unwinding_2pi: "(2*pi) * \ * unwinding(z) = z - Ln(exp z)" by (simp add: unwinding) lemma Ln_times_unwinding: "w \ 0 \ z \ 0 \ Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \ * unwinding(Ln w + Ln z)" using unwinding_2pi by (simp add: exp_add) subsection\<^marker>\tag unimportant\\Relation between Ln and Arg2pi, and hence continuity of Arg2pi\ lemma Arg2pi_Ln: assumes "0 < Arg2pi z" shows "Arg2pi z = Im(Ln(-z)) + pi" proof (cases "z = 0") case True with assms show ?thesis by simp next case False then have "z / of_real(norm z) = exp(\ * of_real(Arg2pi z))" using Arg2pi [of z] by (metis is_Arg_def abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff) then have "- z / of_real(norm z) = exp (\ * (of_real (Arg2pi z) - pi))" using cis_conv_exp cis_pi by (auto simp: exp_diff algebra_simps) then have "ln (- z / of_real(norm z)) = ln (exp (\ * (of_real (Arg2pi z) - pi)))" by simp also have "... = \ * (of_real(Arg2pi z) - pi)" using Arg2pi [of z] assms pi_not_less_zero by auto finally have "Arg2pi z = Im (Ln (- z / of_real (cmod z))) + pi" by simp also have "... = Im (Ln (-z) - ln (cmod z)) + pi" by (metis diff_0_right minus_diff_eq zero_less_norm_iff Ln_divide_of_real False) also have "... = Im (Ln (-z)) + pi" by simp finally show ?thesis . qed lemma continuous_at_Arg2pi: assumes "z \ \\<^sub>\\<^sub>0" shows "continuous (at z) Arg2pi" proof - have *: "isCont (\z. Im (Ln (- z)) + pi) z" by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+ have [simp]: "Im x \ 0 \ Im (Ln (- x)) + pi = Arg2pi x" for x using Arg2pi_Ln by (simp add: Arg2pi_gt_0 complex_nonneg_Reals_iff) consider "Re z < 0" | "Im z \ 0" using assms using complex_nonneg_Reals_iff not_le by blast then have [simp]: "(\z. Im (Ln (- z)) + pi) \z\ Arg2pi z" using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within) show ?thesis unfolding continuous_at proof (rule Lim_transform_within_open) show "\x. \x \ - \\<^sub>\\<^sub>0; x \ z\ \ Im (Ln (- x)) + pi = Arg2pi x" by (auto simp add: Arg2pi_Ln [OF Arg2pi_gt_0] complex_nonneg_Reals_iff) qed (use assms in auto) qed text\Relation between Arg2pi and arctangent in upper halfplane\ lemma Arg2pi_arctan_upperhalf: assumes "0 < Im z" shows "Arg2pi z = pi/2 - arctan(Re z / Im z)" proof (cases "z = 0") case False show ?thesis proof (rule Arg2pi_unique [of "norm z"]) show "(cmod z) * exp (\ * (pi / 2 - arctan (Re z / Im z))) = z" apply (rule complex_eqI) using assms norm_complex_def [of z, symmetric] unfolding exp_Euler cos_diff sin_diff sin_of_real cos_of_real by (simp_all add: field_simps real_sqrt_divide sin_arctan cos_arctan) qed (use False arctan [of "Re z / Im z"] in auto) qed (use assms in auto) lemma Arg2pi_eq_Im_Ln: assumes "0 \ Im z" "0 < Re z" shows "Arg2pi z = Im (Ln z)" proof (cases "Im z = 0") case True then show ?thesis using Arg2pi_eq_0 Ln_in_Reals assms(2) complex_is_Real_iff by auto next case False then have *: "Arg2pi z > 0" using Arg2pi_gt_0 complex_is_Real_iff by blast then have "z \ 0" by auto with * assms False show ?thesis by (subst Arg2pi_Ln) (auto simp: Ln_minus) qed lemma continuous_within_upperhalf_Arg2pi: assumes "z \ 0" shows "continuous (at z within {z. 0 \ Im z}) Arg2pi" proof (cases "z \ \\<^sub>\\<^sub>0") case False then show ?thesis using continuous_at_Arg2pi continuous_at_imp_continuous_within by auto next case True then have z: "z \ \" "0 < Re z" using assms by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0) then have [simp]: "Arg2pi z = 0" "Im (Ln z) = 0" by (auto simp: Arg2pi_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff) show ?thesis proof (clarsimp simp add: continuous_within Lim_within dist_norm) fix e::real assume "0 < e" moreover have "continuous (at z) (\x. Im (Ln x))" using z by (simp add: continuous_at_Ln complex_nonpos_Reals_iff) ultimately obtain d where d: "d>0" "\x. x \ z \ cmod (x - z) < d \ \Im (Ln x)\ < e" by (auto simp: continuous_within Lim_within dist_norm) { fix x assume "cmod (x - z) < Re z / 2" then have "\Re x - Re z\ < Re z / 2" by (metis le_less_trans abs_Re_le_cmod minus_complex.simps(1)) then have "0 < Re x" using z by linarith } then show "\d>0. \x. 0 \ Im x \ x \ z \ cmod (x - z) < d \ \Arg2pi x\ < e" apply (rule_tac x="min d (Re z / 2)" in exI) using z d by (auto simp: Arg2pi_eq_Im_Ln) qed qed lemma continuous_on_upperhalf_Arg2pi: "continuous_on ({z. 0 \ Im z} - {0}) Arg2pi" unfolding continuous_on_eq_continuous_within by (metis DiffE Diff_subset continuous_within_subset continuous_within_upperhalf_Arg2pi insertCI) lemma open_Arg2pi2pi_less_Int: assumes "0 \ s" "t \ 2*pi" shows "open ({y. s < Arg2pi y} \ {y. Arg2pi y < t})" proof - have 1: "continuous_on (UNIV - \\<^sub>\\<^sub>0) Arg2pi" using continuous_at_Arg2pi continuous_at_imp_continuous_within by (auto simp: continuous_on_eq_continuous_within) have 2: "open (UNIV - \\<^sub>\\<^sub>0 :: complex set)" by (simp add: open_Diff) have "open ({z. s < z} \ {z. z < t})" using open_lessThan [of t] open_greaterThan [of s] by (metis greaterThan_def lessThan_def open_Int) moreover have "{y. s < Arg2pi y} \ {y. Arg2pi y < t} \ - \\<^sub>\\<^sub>0" using assms by (auto simp: Arg2pi_real complex_nonneg_Reals_iff complex_is_Real_iff) ultimately show ?thesis using continuous_imp_open_vimage [OF 1 2, of "{z. Re z > s} \ {z. Re z < t}"] by auto qed lemma open_Arg2pi2pi_gt: "open {z. t < Arg2pi z}" proof (cases "t < 0") case True then have "{z. t < Arg2pi z} = UNIV" using Arg2pi_ge_0 less_le_trans by auto then show ?thesis by simp next case False then show ?thesis using open_Arg2pi2pi_less_Int [of t "2*pi"] Arg2pi_lt_2pi by auto qed lemma closed_Arg2pi2pi_le: "closed {z. Arg2pi z \ t}" using open_Arg2pi2pi_gt [of t] by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le) subsection\<^marker>\tag unimportant\\Complex Powers\ lemma powr_to_1 [simp]: "z powr 1 = (z::complex)" by (simp add: powr_def) lemma powr_nat: fixes n::nat and z::complex shows "z powr n = (if z = 0 then 0 else z^n)" by (simp add: exp_of_nat_mult powr_def) lemma norm_powr_real: "w \ \ \ 0 < Re w \ norm(w powr z) = exp(Re z * ln(Re w))" using Ln_Reals_eq norm_exp_eq_Re by (auto simp: Im_Ln_eq_0 powr_def norm_complex_def) lemma powr_complexpow [simp]: fixes x::complex shows "x \ 0 \ x powr (of_nat n) = x^n" by (induct n) (auto simp: ac_simps powr_add) lemma powr_complexnumeral [simp]: fixes x::complex shows "x \ 0 \ x powr (numeral n) = x ^ (numeral n)" by (metis of_nat_numeral powr_complexpow) lemma cnj_powr: assumes "Im a = 0 \ Re a \ 0" shows "cnj (a powr b) = cnj a powr cnj b" proof (cases "a = 0") case False with assms have "a \ \\<^sub>\\<^sub>0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff) with False show ?thesis by (simp add: powr_def exp_cnj cnj_Ln) qed simp lemma powr_real_real: assumes "w \ \" "z \ \" "0 < Re w" shows "w powr z = exp(Re z * ln(Re w))" proof - have "w \ 0" using assms by auto with assms show ?thesis by (simp add: powr_def Ln_Reals_eq of_real_exp) qed lemma powr_of_real: fixes x::real and y::real shows "0 \ x \ of_real x powr (of_real y::complex) = of_real (x powr y)" by (simp_all add: powr_def exp_eq_polar) lemma powr_of_int: fixes z::complex and n::int assumes "z\(0::complex)" shows "z powr of_int n = (if n\0 then z^nat n else inverse (z^nat (-n)))" by (metis assms not_le of_int_of_nat powr_complexpow powr_minus) lemma powr_Reals_eq: "\x \ \; y \ \; Re x \ 0\ \ x powr y = of_real (Re x powr Re y)" by (metis of_real_Re powr_of_real) lemma norm_powr_real_mono: "\w \ \; 1 < Re w\ \ cmod(w powr z1) \ cmod(w powr z2) \ Re z1 \ Re z2" by (auto simp: powr_def algebra_simps Reals_def Ln_of_real) lemma powr_times_real: "\x \ \; y \ \; 0 \ Re x; 0 \ Re y\ \ (x * y) powr z = x powr z * y powr z" by (auto simp: Reals_def powr_def Ln_times exp_add algebra_simps less_eq_real_def Ln_of_real) lemma Re_powr_le: "r \ \\<^sub>\\<^sub>0 \ Re (r powr z) \ Re r powr Re z" by (auto simp: powr_def nonneg_Reals_def order_trans [OF complex_Re_le_cmod]) lemma fixes w::complex shows Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \" and nonneg_Reals_powr [simp]: "\w \ \\<^sub>\\<^sub>0; z \ \\ \ w powr z \ \\<^sub>\\<^sub>0" 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)" proof (cases "x = 0") assume x: "x \ 0" hence "(-x) powr a = exp (a * 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 * \" by (simp add: Ln_minus Ln_of_real) also from x have "exp (a * ...) = cis pi powr (of_real (sgn x) * a) * of_real x powr a" by (simp add: powr_def exp_add algebra_simps Ln_of_real cis_conv_exp) also note cis_pi finally show ?thesis by simp qed simp_all lemma has_field_derivative_powr: fixes z :: complex assumes "z \ \\<^sub>\\<^sub>0" shows "((\z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)" proof (cases "z=0") case False then have \
: "exp (s * Ln z) * inverse z = exp ((s - 1) * Ln z)" by (simp add: divide_complex_def exp_diff left_diff_distrib') show ?thesis unfolding powr_def proof (rule has_field_derivative_transform_within) show "((\z. exp (s * Ln z)) has_field_derivative s * (if z = 0 then 0 else exp ((s - 1) * Ln z))) (at z)" by (intro derivative_eq_intros | simp add: assms False \
)+ qed (use False in auto) qed (use assms in auto) declare has_field_derivative_powr[THEN DERIV_chain2, derivative_intros] lemma has_field_derivative_powr_of_int: fixes z :: complex assumes gderiv:"(g has_field_derivative gd) (at z within S)" and "g z\0" shows "((\z. g z powr of_int n) has_field_derivative (n * g z powr (of_int n - 1) * gd)) (at z within S)" proof - define dd where "dd = of_int n * g z powr (of_int (n - 1)) * gd" obtain e where "e>0" and e_dist:"\y\S. dist z y < e \ g y \ 0" using DERIV_continuous[OF gderiv,THEN continuous_within_avoid] \g z\0\ by auto have ?thesis when "n\0" proof - define dd' where "dd' = of_int n * g z ^ (nat n - 1) * gd" have "dd=dd'" proof (cases "n=0") case False then have "n-1 \0" using \n\0\ by auto then have "g z powr (of_int (n - 1)) = g z ^ (nat n - 1)" using powr_of_int[OF \g z\0\,of "n-1"] by (simp add: nat_diff_distrib') then show ?thesis unfolding dd_def dd'_def by simp qed (simp add:dd_def dd'_def) then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within S)" by simp also have "... \ ((\z. g z ^ nat n) has_field_derivative dd') (at z within S)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within S. g x powr of_int n = g x ^ nat n" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." unfolding dd'_def using gderiv that by (auto intro!: derivative_eq_intros) finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S)" . then show ?thesis unfolding dd_def by simp qed moreover have ?thesis when "n<0" proof - define dd' where "dd' = of_int n / g z ^ (nat (1 - n)) * gd" have "dd=dd'" proof - have "g z powr of_int (n - 1) = inverse (g z ^ nat (1-n))" using powr_of_int[OF \g z\0\,of "n-1"] that by auto then show ?thesis unfolding dd_def dd'_def by (simp add: divide_inverse) qed then have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S) \ ((\z. g z powr of_int n) has_field_derivative dd') (at z within S)" by simp also have "... \ ((\z. inverse (g z ^ nat (-n))) has_field_derivative dd') (at z within S)" proof (rule has_field_derivative_cong_eventually) show "\\<^sub>F x in at z within S. g x powr of_int n = inverse (g x ^ nat (- n))" unfolding eventually_at apply (rule exI[where x=e]) using powr_of_int that \e>0\ e_dist by (simp add: dist_commute) qed (use powr_of_int \g z\0\ that in simp) also have "..." proof - have "nat (- n) + nat (1 - n) - Suc 0 = nat (- n) + nat (- n)" by auto then show ?thesis unfolding dd'_def using gderiv that \g z\0\ by (auto intro!: derivative_eq_intros simp add:field_split_simps power_add[symmetric]) qed finally have "((\z. g z powr of_int n) has_field_derivative dd) (at z within S)" . then show ?thesis unfolding dd_def by simp qed ultimately show ?thesis by force qed lemma field_differentiable_powr_of_int: fixes z :: complex assumes gderiv: "g field_differentiable (at z within S)" and "g z \ 0" shows "(\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 lemma holomorphic_on_powr_of_int [holomorphic_intros]: assumes holf: "f holomorphic_on S" and 0: "\z. z\S \ f z \ 0" shows "(\z. (f z) powr of_int n) holomorphic_on S" proof (cases "n\0") case True then have "?thesis \ (\z. (f z) ^ nat n) holomorphic_on S" by (metis (no_types, lifting) 0 holomorphic_cong powr_of_int) moreover have "(\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 \ (\z. inverse (f z) ^ nat (-n)) holomorphic_on S" by (metis (no_types, lifting) "0" holomorphic_cong power_inverse powr_of_int) moreover have "(\z. inverse (f z) ^ nat (-n)) holomorphic_on S" using assms by (auto intro!:holomorphic_intros) ultimately show ?thesis by auto qed lemma has_field_derivative_powr_right [derivative_intros]: "w \ 0 \ ((\z. w powr z) has_field_derivative Ln w * w powr z) (at z)" unfolding powr_def by (intro derivative_eq_intros | simp)+ lemma field_differentiable_powr_right [derivative_intros]: fixes w::complex shows "w \ 0 \ (\z. w powr z) field_differentiable (at z)" using field_differentiable_def has_field_derivative_powr_right by blast lemma holomorphic_on_powr_right [holomorphic_intros]: assumes "f holomorphic_on s" shows "(\z. w powr (f z)) holomorphic_on s" proof (cases "w = 0") case False with assms show ?thesis unfolding holomorphic_on_def field_differentiable_def by (metis (full_types) DERIV_chain' has_field_derivative_powr_right) qed simp lemma holomorphic_on_divide_gen [holomorphic_intros]: assumes f: "f holomorphic_on s" and g: "g holomorphic_on s" and 0: "\z z'. \z \ s; z' \ s\ \ g z = 0 \ g z' = 0" shows "(\z. f z / g z) holomorphic_on s" proof (cases "\z\s. g z = 0") case True with 0 have "g z = 0" if "z \ s" for z using that by blast then show ?thesis using g holomorphic_transform by auto next case False with 0 have "g z \ 0" if "z \ s" for z using that by blast with holomorphic_on_divide show ?thesis using f g by blast qed lemma norm_powr_real_powr: "w \ \ \ 0 \ Re w \ cmod (w powr z) = Re w powr Re z" by (metis dual_order.order_iff_strict norm_powr_real norm_zero of_real_0 of_real_Re powr_def) lemma tendsto_powr_complex: fixes f g :: "_ \ complex" assumes a: "a \ \\<^sub>\\<^sub>0" assumes f: "(f \ a) F" and g: "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" proof - from a have [simp]: "a \ 0" by auto from f g a have "((\z. exp (g z * ln (f z))) \ a powr b) F" (is ?P) by (auto intro!: tendsto_intros simp: powr_def) also { have "eventually (\z. z \ 0) (nhds a)" by (intro t1_space_nhds) simp_all with f have "eventually (\z. f z \ 0) F" using filterlim_iff by blast } hence "?P \ ((\z. f z powr g z) \ a powr b) F" by (intro tendsto_cong refl) (simp_all add: powr_def mult_ac) finally show ?thesis . qed lemma tendsto_powr_complex_0: fixes f g :: "'a \ complex" assumes f: "(f \ 0) F" and g: "(g \ b) F" and b: "Re b > 0" shows "((\z. f z powr g z) \ 0) F" proof (rule tendsto_norm_zero_cancel) define h where "h = (\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 \ 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))) \ pi" by simp from mult_left_mono[OF this, of "abs (Im (g z))"] have "abs (Im (g z) * Im (ln (f z))) \ c" by (simp add: abs_mult c_def) hence "-Im (g z) * Im (ln (f z)) \ c" by simp hence "norm (f z powr g z) \ h z" by (simp add: powr_def field_simps h_def c_def) } hence le: "norm (f z powr g z) \ h z" for z by (cases "f x = 0") (simp_all add: h_def) have g': "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all have "((\x. norm (f x)) \ 0) (inf F (principal {z. f z \ 0}))" by (subst tendsto_norm_zero_iff, rule tendsto_mono[OF _ f]) simp_all moreover { have "filterlim (\x. norm (f x)) (principal {0<..}) (principal {z. f z \ 0})" by (auto simp: filterlim_def) hence "filterlim (\x. norm (f x)) (principal {0<..}) (inf F (principal {z. f z \ 0}))" by (rule filterlim_mono) simp_all } ultimately have norm: "filterlim (\x. norm (f x)) (at_right 0) (inf F (principal {z. f z \ 0}))" by (simp add: filterlim_inf at_within_def) have A: "LIM x inf F (principal {z. f z \ 0}). Re (g x) * -ln (cmod (f x)) :> at_top" by (rule filterlim_tendsto_pos_mult_at_top tendsto_intros g' b filterlim_compose[OF filterlim_uminus_at_top_at_bot] filterlim_compose[OF ln_at_0] norm)+ have B: "LIM x inf F (principal {z. f z \ 0}). -\Im (g x)\ * pi + -(Re (g x) * ln (cmod (f x))) :> at_top" by (rule filterlim_tendsto_add_at_top tendsto_intros g')+ (insert A, simp_all) have C: "(h \ 0) F" unfolding h_def by (intro filterlim_If tendsto_const filterlim_compose[OF exp_at_bot]) (insert B, auto simp: filterlim_uminus_at_bot algebra_simps) show "((\x. norm (f x powr g x)) \ 0) F" by (rule Lim_null_comparison[OF always_eventually C]) (insert le, auto) qed lemma tendsto_powr_complex' [tendsto_intros]: fixes f g :: "_ \ complex" assumes "a \ \\<^sub>\\<^sub>0 \ (a = 0 \ Re b > 0)" and "(f \ a) F" "(g \ b) F" shows "((\z. f z powr g z) \ a powr b) F" using assms tendsto_powr_complex tendsto_powr_complex_0 by fastforce lemma tendsto_neg_powr_complex_of_real: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. complex_of_real (f x) powr s) \ 0) F" proof - have "((\x. norm (complex_of_real (f x) powr s)) \ 0) F" proof (rule Lim_transform_eventually) from assms(1) have "eventually (\x. f x \ 0) F" by (auto simp: filterlim_at_top) thus "eventually (\x. f x powr Re s = norm (of_real (f x) powr s)) F" by eventually_elim (simp add: norm_powr_real_powr) from assms show "((\x. f x powr Re s) \ 0) F" by (intro tendsto_neg_powr) qed thus ?thesis by (simp add: tendsto_norm_zero_iff) qed lemma tendsto_neg_powr_complex_of_nat: assumes "filterlim f at_top F" and "Re s < 0" shows "((\x. of_nat (f x) powr s) \ 0) F" proof - have "((\x. of_real (real (f x)) powr s) \ 0) F" using assms(2) by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real] filterlim_compose[OF _ assms(1)] filterlim_real_sequentially filterlim_ident) auto thus ?thesis by simp qed lemma continuous_powr_complex: assumes "f (netlimit F) \ \\<^sub>\\<^sub>0" "continuous F f" "continuous F g" shows "continuous F (\z. f z powr g z :: complex)" using assms unfolding continuous_def by (intro tendsto_powr_complex) simp_all lemma isCont_powr_complex [continuous_intros]: assumes "f z \ \\<^sub>\\<^sub>0" "isCont f z" "isCont g z" shows "isCont (\z. f z powr g z :: complex) z" using assms unfolding isCont_def by (intro tendsto_powr_complex) simp_all lemma continuous_on_powr_complex [continuous_intros]: assumes "A \ {z. Re (f z) \ 0 \ Im (f z) \ 0}" assumes "\z. z \ A \ f z = 0 \ Re (g z) > 0" assumes "continuous_on A f" "continuous_on A g" shows "continuous_on A (\z. f z powr g z)" unfolding continuous_on_def proof fix z assume z: "z \ A" show "((\z. f z powr g z) \ f z powr g z) (at z within A)" proof (cases "f z = 0") case False from assms(1,2) z have "Re (f z) \ 0 \ Im (f z) \ 0" "f z = 0 \ Re (g z) > 0" by auto with assms(3,4) z show ?thesis by (intro tendsto_powr_complex') (auto elim!: nonpos_Reals_cases simp: complex_eq_iff continuous_on_def) next case True with assms z show ?thesis by (auto intro!: tendsto_powr_complex_0 simp: continuous_on_def) qed qed subsection\<^marker>\tag unimportant\\Some Limits involving Logarithms\ lemma lim_Ln_over_power: fixes s::complex assumes "0 < Re s" shows "(\n. Ln (of_nat n) / of_nat n powr s) \ 0" proof (simp add: lim_sequentially dist_norm, clarify) fix e::real assume e: "0 < e" have "\xo>0. \x\xo. 0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" proof (rule_tac x="2/(e * (Re s)\<^sup>2)" in exI, safe) show "0 < 2 / (e * (Re s)\<^sup>2)" using e assms by (simp add: field_simps) next fix x::real assume x: "2 / (e * (Re s)\<^sup>2) \ x" have "2 / (e * (Re s)\<^sup>2) > 0" using e assms by simp with x have "x > 0" by linarith then have "x * 2 \ e * (x\<^sup>2 * (Re s)\<^sup>2)" using e assms x by (auto simp: power2_eq_square field_simps) also have "... < e * (2 + (x * (Re s * 2) + x\<^sup>2 * (Re s)\<^sup>2))" using e assms \x > 0\ by (auto simp: power2_eq_square field_simps add_pos_pos) finally show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2" by (auto simp: algebra_simps) qed then have "\xo>0. \x\xo. x / e < 1 + (Re s * x) + (1/2) * (Re s * x)^2" using e by (simp add: field_simps) then have "\xo>0. \x\xo. x / e < exp (Re s * x)" using assms by (force intro: less_le_trans [OF _ exp_lower_Taylor_quadratic]) then obtain xo where "xo > 0" and xo: "\x. x \ xo \ x < e * exp (Re s * x)" using e by (auto simp: field_simps) have "norm (Ln (of_nat n) / of_nat n powr s) < e" if "n \ nat \exp xo\" for n proof - have "ln (real n) \ xo" using that exp_gt_zero ln_ge_iff [of n] nat_ceiling_le_eq by fastforce then show ?thesis using e xo [of "ln n"] by (auto simp: norm_divide norm_powr_real field_split_simps) qed then show "\no. \n\no. norm (Ln (of_nat n) / of_nat n powr s) < e" by blast qed lemma lim_Ln_over_n: "((\n. Ln(of_nat n) / of_nat n) \ 0) sequentially" using lim_Ln_over_power [of 1] by simp lemma lim_ln_over_power: fixes s :: real assumes "0 < s" shows "((\n. ln n / (n powr s)) \ 0) sequentially" proof - have "(\n. ln (Suc n) / (Suc n) powr s) \ 0" using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide) then show ?thesis using filterlim_sequentially_Suc[of "\n::nat. ln n / n powr s"] by auto qed lemma lim_ln_over_n [tendsto_intros]: "((\n. ln(real_of_nat n) / of_nat n) \ 0) sequentially" using lim_ln_over_power [of 1] by auto lemma lim_log_over_n [tendsto_intros]: "(\n. log k n/n) \ 0" proof - have *: "log k n/n = (1/ln k) * (ln n / n)" for n unfolding log_def by auto have "(\n. (1/ln k) * (ln n / n)) \ (1/ln k) * 0" by (intro tendsto_intros) then show ?thesis unfolding * by auto qed lemma lim_1_over_complex_power: assumes "0 < Re s" shows "(\n. 1 / of_nat n powr s) \ 0" proof (rule Lim_null_comparison) have "\n>0. 3 \ n \ 1 \ ln (real_of_nat n)" using ln_272_gt_1 by (force intro: order_trans [of _ "ln (272/100)"]) then show "\\<^sub>F x in sequentially. cmod (1 / of_nat x powr s) \ cmod (Ln (of_nat x) / of_nat x powr s)" by (auto simp: norm_divide field_split_simps eventually_sequentially) show "(\n. cmod (Ln (of_nat n) / of_nat n powr s)) \ 0" using lim_Ln_over_power [OF assms] by (metis tendsto_norm_zero_iff) qed lemma lim_1_over_real_power: fixes s :: real assumes "0 < s" shows "((\n. 1 / (of_nat n powr s)) \ 0) sequentially" using lim_1_over_complex_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms 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: "((\n. 1 / Ln(of_nat n)) \ 0) sequentially" proof (clarsimp simp add: lim_sequentially dist_norm norm_divide field_split_simps) fix r::real assume "0 < r" have ir: "inverse (exp (inverse r)) > 0" by simp obtain n where n: "1 < of_nat n * inverse (exp (inverse r))" using ex_less_of_nat_mult [of _ 1, OF ir] by auto then have "exp (inverse r) < of_nat n" by (simp add: field_split_simps) then have "ln (exp (inverse r)) < ln (of_nat n)" by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff) with \0 < r\ have "1 < r * ln (real_of_nat n)" by (simp add: field_simps) moreover have "n > 0" using n using neq0_conv by fastforce ultimately show "\no. \k. Ln (of_nat k) \ 0 \ no \ k \ 1 < r * cmod (Ln (of_nat k))" using n \0 < r\ by (rule_tac x=n in exI) (force simp: field_split_simps intro: less_le_trans) qed lemma lim_1_over_ln: "((\n. 1 / ln(real_of_nat n)) \ 0) sequentially" 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) lemma lim_ln1_over_ln: "(\n. ln(Suc n) / ln n) \ 1" proof (rule Lim_transform_eventually) have "(\n. ln(1 + 1/n) / ln n) \ 0" proof (rule Lim_transform_bound) show "(inverse o real) \ 0" by (metis comp_def lim_inverse_n lim_explicit) show "\\<^sub>F n in sequentially. norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" proof fix n::nat assume n: "3 \ n" then have "ln 3 \ ln n" and ln0: "0 \ ln n" by auto with ln3_gt_1 have "1/ ln n \ 1" by (simp add: field_split_simps) moreover have "ln (1 + 1 / real n) \ 1/n" by (simp add: ln_add_one_self_le_self) ultimately have "ln (1 + 1 / real n) * (1 / ln n) \ (1/n) * 1" by (intro mult_mono) (use n in auto) then show "norm (ln (1 + 1 / n) / ln n) \ norm ((inverse \ real) n)" by (simp add: field_simps ln0) qed qed then show "(\n. 1 + ln(1 + 1/n) / ln n) \ 1" by (metis (full_types) add.right_neutral tendsto_add_const_iff) show "\\<^sub>F k in sequentially. 1 + ln (1 + 1 / k) / ln k = ln(Suc k) / ln k" by (simp add: field_split_simps ln_div eventually_sequentiallyI [of 2]) qed lemma lim_ln_over_ln1: "(\n. ln n / ln(Suc n)) \ 1" proof - have "(\n. inverse (ln(Suc n) / ln n)) \ inverse 1" by (rule tendsto_inverse [OF lim_ln1_over_ln]) auto then show ?thesis by simp qed subsection\<^marker>\tag unimportant\\Relation between Square Root and exp/ln, hence its derivative\ lemma csqrt_exp_Ln: assumes "z \ 0" 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) also have "... = z" using assms exp_Ln by blast finally have "csqrt z = csqrt ((exp (Ln z / 2))\<^sup>2)" by simp also have "... = exp (Ln z / 2)" apply (rule csqrt_square) using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms by (fastforce simp: Re_exp Im_exp ) finally show ?thesis using assms csqrt_square by simp qed lemma csqrt_inverse: assumes "z \ \\<^sub>\\<^sub>0" shows "csqrt (inverse z) = inverse (csqrt z)" proof (cases "z=0") case False then show ?thesis using assms csqrt_exp_Ln Ln_inverse exp_minus by (simp add: csqrt_exp_Ln Ln_inverse exp_minus) qed auto lemma cnj_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "cnj(csqrt z) = csqrt(cnj z)" proof (cases "z=0") case False then show ?thesis by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj) qed auto lemma has_field_derivative_csqrt: assumes "z \ \\<^sub>\\<^sub>0" shows "(csqrt has_field_derivative inverse(2 * csqrt z)) (at z)" proof - have z: "z \ 0" using assms by auto then have *: "inverse z = inverse (2*z) * 2" by (simp add: field_split_simps) have [simp]: "exp (Ln z / 2) * inverse z = inverse (csqrt z)" by (simp add: z field_simps csqrt_exp_Ln [symmetric]) (metis power2_csqrt power2_eq_square) have "Im z = 0 \ 0 < Re z" using assms complex_nonpos_Reals_iff not_less by blast with z have "((\z. exp (Ln z / 2)) has_field_derivative inverse (2 * csqrt z)) (at z)" by (force intro: derivative_eq_intros * simp add: assms) then show ?thesis proof (rule has_field_derivative_transform_within) show "\x. dist x z < cmod z \ exp (Ln x / 2) = csqrt x" by (metis csqrt_exp_Ln dist_0_norm less_irrefl) qed (use z in auto) qed lemma field_differentiable_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable at z" using field_differentiable_def has_field_derivative_csqrt by blast lemma field_differentiable_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ csqrt field_differentiable (at z within s)" using field_differentiable_at_csqrt field_differentiable_within_subset by blast lemma continuous_at_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z) csqrt" by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at) corollary\<^marker>\tag unimportant\ isCont_csqrt' [simp]: "\isCont f z; f z \ \\<^sub>\\<^sub>0\ \ isCont (\x. csqrt (f x)) z" by (blast intro: isCont_o2 [OF _ continuous_at_csqrt]) lemma continuous_within_csqrt: "z \ \\<^sub>\\<^sub>0 \ continuous (at z within s) csqrt" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_csqrt) lemma continuous_on_csqrt [continuous_intros]: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ continuous_on s csqrt" by (simp add: continuous_at_imp_continuous_on continuous_within_csqrt) lemma holomorphic_on_csqrt: "(\z. z \ s \ z \ \\<^sub>\\<^sub>0) \ csqrt holomorphic_on s" by (simp add: field_differentiable_within_csqrt holomorphic_on_def) lemma continuous_within_closed_nontrivial: "closed s \ a \ s ==> continuous (at a within s) f" using open_Compl by (force simp add: continuous_def eventually_at_topological filterlim_iff open_Collect_neg) lemma continuous_within_csqrt_posreal: "continuous (at z within (\ \ {w. 0 \ Re(w)})) csqrt" proof (cases "z \ \\<^sub>\\<^sub>0") case True have [simp]: "Im z = 0" and 0: "Re z < 0 \ z = 0" using True cnj.code complex_cnj_zero_iff by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce show ?thesis using 0 proof assume "Re z < 0" then show ?thesis by (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge]) next assume "z = 0" moreover have "\e. 0 < e \ \x'\\ \ {w. 0 \ Re w}. cmod x' < e^2 \ cmod (csqrt x') < e" by (auto simp: Reals_def real_less_lsqrt) ultimately show ?thesis using zero_less_power by (fastforce simp: continuous_within_eps_delta) qed qed (blast intro: continuous_within_csqrt) subsection\Complex arctangent\ text\The branch cut gives standard bounds in the real case.\ definition\<^marker>\tag important\ Arctan :: "complex \ complex" where "Arctan \ \z. (\/2) * Ln((1 - \*z) / (1 + \*z))" lemma Arctan_def_moebius: "Arctan z = \/2 * Ln (moebius (-\) 1 \ 1 z)" by (simp add: Arctan_def moebius_def add_ac) lemma Ln_conv_Arctan: assumes "z \ -1" shows "Ln z = -2*\ * Arctan (moebius 1 (- 1) (- \) (- \) z)" proof - have "Arctan (moebius 1 (- 1) (- \) (- \) z) = \/2 * Ln (moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z))" by (simp add: Arctan_def_moebius) also from assms have "\ * z \ \ * (-1)" by (subst mult_left_cancel) simp hence "\ * z - -\ \ 0" by (simp add: eq_neg_iff_add_eq_0) from moebius_inverse'[OF _ this, of 1 1] have "moebius (- \) 1 \ 1 (moebius 1 (- 1) (- \) (- \) z) = z" by simp finally show ?thesis by (simp add: field_simps) qed lemma Arctan_0 [simp]: "Arctan 0 = 0" by (simp add: Arctan_def) lemma Im_complex_div_lemma: "Im((1 - \*z) / (1 + \*z)) = 0 \ Re z = 0" by (auto simp: Im_complex_div_eq_0 algebra_simps) lemma Re_complex_div_lemma: "0 < Re((1 - \*z) / (1 + \*z)) \ norm z < 1" by (simp add: Re_complex_div_gt_0 algebra_simps cmod_def power2_eq_square) lemma tan_Arctan: assumes "z\<^sup>2 \ -1" shows [simp]:"tan(Arctan z) = z" proof - have "1 + \*z \ 0" by (metis assms complex_i_mult_minus i_squared minus_unique power2_eq_square power2_minus) moreover have "1 - \*z \ 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]) qed lemma Arctan_tan [simp]: assumes "\Re z\ < pi/2" shows "Arctan(tan z) = z" proof - have "Ln ((1 - \ * tan z) / (1 + \ * tan z)) = 2 * z / \" proof (rule Ln_unique) have ge_pi2: "\n::int. \of_int (2*n + 1) * pi/2\ \ pi/2" by (case_tac n rule: int_cases) (auto simp: abs_mult) have "exp (\*z)*exp (\*z) = -1 \ exp (2*\*z) = -1" by (metis distrib_right exp_add mult_2) also have "... \ exp (2*\*z) = exp (\*pi)" using cis_conv_exp cis_pi by auto also have "... \ exp (2*\*z - \*pi) = 1" by (metis (no_types) diff_add_cancel diff_minus_eq_add exp_add exp_minus_inverse mult.commute) also have "... \ Re(\*2*z - \*pi) = 0 \ (\n::int. Im(\*2*z - \*pi) = of_int (2 * n) * pi)" by (simp add: exp_eq_1) also have "... \ Im z = 0 \ (\n::int. 2 * Re z = of_int (2*n + 1) * pi)" by (simp add: algebra_simps) also have "... \ 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) finally have "exp (\*z)*exp (\*z) + 1 \ 0" by (auto simp: add.commute minus_unique) then show "exp (2 * z / \) = (1 - \ * tan z) / (1 + \ * tan z)" apply (simp add: tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps) by (simp add: algebra_simps flip: power2_eq_square exp_double) qed (use assms in auto) then show ?thesis by (auto simp: Arctan_def) qed lemma assumes "Re z = 0 \ \Im z\ < 1" shows Re_Arctan_bounds: "\Re(Arctan z)\ < pi/2" and has_field_derivative_Arctan: "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" proof - have nz0: "1 + \*z \ 0" using assms by (metis abs_one add_diff_cancel_left' complex_i_mult_minus diff_0 i_squared imaginary_unit.simps less_asym neg_equal_iff_equal) have "z \ -\" using assms by auto then have zz: "1 + z * z \ 0" by (metis abs_one assms i_squared imaginary_unit.simps less_irrefl minus_unique square_eq_iff) have nz1: "1 - \*z \ 0" using assms by (force simp add: i_times_eq_iff) have nz2: "inverse (1 + \*z) \ 0" using assms by (metis Im_complex_div_lemma Re_complex_div_lemma cmod_eq_Im divide_complex_def less_irrefl mult_zero_right zero_complex.simps(1) zero_complex.simps(2)) have nzi: "((1 - \*z) * inverse (1 + \*z)) \ 0" using nz1 nz2 by auto have "Im ((1 - \*z) / (1 + \*z)) = 0 \ 0 < Re ((1 - \*z) / (1 + \*z))" apply (simp add: divide_complex_def) apply (simp add: divide_simps split: if_split_asm) using assms apply (auto simp: algebra_simps abs_square_less_1 [unfolded power2_eq_square]) done then have *: "((1 - \*z) / (1 + \*z)) \ \\<^sub>\\<^sub>0" by (auto simp add: complex_nonpos_Reals_iff) show "\Re(Arctan z)\ < pi/2" unfolding Arctan_def divide_complex_def using mpi_less_Im_Ln [OF nzi] by (auto simp: abs_if intro!: Im_Ln_less_pi * [unfolded divide_complex_def]) show "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)" unfolding Arctan_def scaleR_conv_of_real apply (intro derivative_eq_intros | simp add: nz0 *)+ using nz1 zz apply (simp add: field_split_simps power2_eq_square) apply algebra done qed lemma field_differentiable_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable at z" using has_field_derivative_Arctan by (auto simp: field_differentiable_def) lemma field_differentiable_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ Arctan field_differentiable (at z within s)" using field_differentiable_at_Arctan field_differentiable_at_within by blast declare has_field_derivative_Arctan [derivative_intros] declare has_field_derivative_Arctan [THEN DERIV_chain2, derivative_intros] lemma continuous_at_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z) Arctan" by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Arctan) lemma continuous_within_Arctan: "(Re z = 0 \ \Im z\ < 1) \ continuous (at z within s) Arctan" using continuous_at_Arctan continuous_at_imp_continuous_within by blast lemma continuous_on_Arctan [continuous_intros]: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ continuous_on s Arctan" by (auto simp: continuous_at_imp_continuous_on continuous_within_Arctan) lemma holomorphic_on_Arctan: "(\z. z \ s \ Re z = 0 \ \Im z\ < 1) \ Arctan holomorphic_on s" by (simp add: field_differentiable_within_Arctan holomorphic_on_def) theorem Arctan_series: assumes z: "norm (z :: complex) < 1" defines "g \ \n. if odd n then -\*\^n / n else 0" defines "h \ \z n. (-1)^n / of_nat (2*n+1) * (z::complex)^(2*n+1)" shows "(\n. g n * z^n) sums Arctan z" and "h z sums Arctan z" proof - define G where [abs_def]: "G z = (\n. g n * z^n)" for z have summable: "summable (\n. g n * u^n)" if "norm u < 1" for u proof (cases "u = 0") assume u: "u \ 0" have "(\n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\n. ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))" proof fix n have "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal (((2*Suc n+1) / (Suc n)) / ((2*Suc n-1) / (Suc n)))" by (simp add: h_def norm_mult norm_power norm_divide field_split_simps power2_eq_square eval_nat_numeral del: of_nat_add of_nat_Suc) also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc n))" by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all? finally show "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n))))" . qed also have "\ \ ereal (inverse (norm u)^2) * ereal ((2 + 0) / (2 - 0))" by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all finally have "liminf (\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" by (simp add: field_split_simps) ultimately have A: "liminf (\n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp from u 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) thus "summable (\n. g n * u^n)" by (subst summable_mono_reindex[of "\n. 2*n+1", symmetric]) (auto simp: power_mult strict_mono_def g_def h_def elim!: oddE) qed (simp add: h_def) have "\c. \u\ball 0 1. Arctan u - G u = c" proof (rule has_field_derivative_zero_constant) fix u :: complex assume "u \ ball 0 1" hence u: "norm u < 1" by (simp) define K where "K = (norm u + 1) / 2" from u and abs_Im_le_cmod[of u] have Im_u: "\Im u\ < 1" by linarith from u have K: "0 \ K" "norm u < K" "K < 1" by (simp_all add: K_def) hence "(G has_field_derivative (\n. diffs g n * u ^ n)) (at u)" unfolding G_def by (intro termdiffs_strong[of _ "of_real K"] summable) simp_all also have "(\n. diffs g n * u^n) = (\n. if even n then (\*u)^n else 0)" by (intro ext) (simp_all del: of_nat_Suc add: g_def diffs_def power_mult_distrib) also have "suminf \ = (\n. (-(u^2))^n)" by (subst suminf_mono_reindex[of "\n. 2*n", symmetric]) (auto elim!: evenE simp: strict_mono_def power_mult power_mult_distrib) also from u have "norm u^2 < 1^2" by (intro power_strict_mono) simp_all hence "(\n. (-(u^2))^n) = inverse (1 + u^2)" by (subst suminf_geometric) (simp_all add: norm_power inverse_eq_divide) finally have "(G has_field_derivative inverse (1 + u\<^sup>2)) (at u)" . from DERIV_diff[OF has_field_derivative_Arctan this] Im_u u show "((\u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)" by (simp_all add: at_within_open[OF _ open_ball]) qed simp_all then obtain c where c: "\u. norm u < 1 \ Arctan u - G u = c" by auto from this[of 0] have "c = 0" by (simp add: G_def g_def) with c z have "Arctan z = G z" by simp with summable[OF z] show "(\n. g n * z^n) sums Arctan z" unfolding G_def by (simp add: sums_iff) thus "h z sums Arctan z" by (subst (asm) sums_mono_reindex[of "\n. 2*n+1", symmetric]) (auto elim!: oddE simp: strict_mono_def power_mult g_def h_def) qed text \A quickly-converging series for the logarithm, based on the arctangent.\ theorem ln_series_quadratic: assumes x: "x > (0::real)" shows "(\n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x" proof - define y :: complex where "y = of_real ((x-1)/(x+1))" from x have x': "complex_of_real x \ of_real (-1)" by (subst of_real_eq_iff) auto from x have "\x - 1\ < \x + 1\" by linarith hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1" by (simp add: norm_divide del: of_real_add of_real_diff) hence "norm (\ * y) < 1" unfolding y_def by (subst norm_mult) simp hence "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) sums ((-2*\) * Arctan (\*y))" by (intro Arctan_series sums_mult) simp_all also have "(\n. (-2*\) * ((-1)^n / of_nat (2*n+1) * (\*y)^(2*n+1))) = (\n. (-2*\) * ((-1)^n * (\*y*(-y\<^sup>2)^n)/of_nat (2*n+1)))" by (intro ext) (simp_all add: power_mult power_mult_distrib) also have "\ = (\n. 2*y* ((-1) * (-y\<^sup>2))^n/of_nat (2*n+1))" by (intro ext, subst power_mult_distrib) (simp add: algebra_simps power_mult) also have "\ = (\n. 2*y^(2*n+1) / of_nat (2*n+1))" by (subst power_add, subst power_mult) (simp add: mult_ac) also have "\ = (\n. of_real (2*((x-1)/(x+1))^(2*n+1) / of_nat (2*n+1)))" by (intro ext) (simp add: y_def) also have "\ * y = (of_real x - 1) / (-\ * (of_real x + 1))" by (subst divide_divide_eq_left [symmetric]) (simp add: y_def) also have "\ = moebius 1 (-1) (-\) (-\) (of_real x)" by (simp add: moebius_def algebra_simps) also from x' have "-2*\*Arctan \ = Ln (of_real x)" by (intro Ln_conv_Arctan [symmetric]) simp_all also from x have "\ = ln x" by (rule Ln_of_real) finally show ?thesis by (subst (asm) sums_of_real_iff) qed subsection\<^marker>\tag unimportant\ \Real arctangent\ lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0" proof - have ne: "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) have ne1: "1 + \ * complex_of_real x \ 0" using Complex_eq complex_eq_cancel_iff2 by fastforce have "Re (Ln ((1 - \ * x) * inverse (1 + \ * x))) = 0" apply (rule norm_exp_imaginary) using ne apply (simp add: ne1 cmod_def) apply (auto simp: field_split_simps) apply algebra done then show ?thesis unfolding Arctan_def divide_complex_def by (simp add: complex_eq_iff) qed lemma arctan_eq_Re_Arctan: "arctan x = Re (Arctan (of_real x))" proof (rule arctan_unique) have "(1 - \ * x) / (1 + \ * x) \ \\<^sub>\\<^sub>0" by (auto simp: Im_complex_div_lemma complex_nonpos_Reals_iff) then show "- (pi / 2) < Re (Arctan (complex_of_real x))" by (simp add: Arctan_def Im_Ln_less_pi) next have *: " (1 - \*x) / (1 + \*x) \ 0" by (simp add: field_split_simps) ( simp add: complex_eq_iff) show "Re (Arctan (complex_of_real x)) < pi / 2" using mpi_less_Im_Ln [OF *] 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) also have "... = x" proof - have "(complex_of_real x)\<^sup>2 \ - 1" by (metis diff_0_right minus_diff_eq mult_zero_left not_le of_real_1 of_real_eq_iff of_real_minus of_real_power power2_eq_square real_minus_mult_self_le zero_less_one) then show ?thesis by simp qed finally show "tan (Re (Arctan (complex_of_real x))) = x" . qed lemma Arctan_of_real: "Arctan (of_real x) = of_real (arctan x)" unfolding arctan_eq_Re_Arctan divide_complex_def by (simp add: complex_eq_iff) lemma Arctan_in_Reals [simp]: "z \ \ \ Arctan z \ \" by (metis Reals_cases Reals_of_real Arctan_of_real) declare arctan_one [simp] lemma arctan_less_pi4_pos: "x < 1 \ arctan x < pi/4" by (metis arctan_less_iff arctan_one) lemma arctan_less_pi4_neg: "-1 < x \ -(pi/4) < arctan x" by (metis arctan_less_iff arctan_minus arctan_one) lemma arctan_less_pi4: "\x\ < 1 \ \arctan x\ < pi/4" by (metis abs_less_iff arctan_less_pi4_pos arctan_minus) lemma arctan_le_pi4: "\x\ \ 1 \ \arctan x\ \ pi/4" by (metis abs_le_iff arctan_le_iff arctan_minus arctan_one) lemma abs_arctan: "\arctan x\ = arctan \x\" by (simp add: abs_if arctan_minus) lemma arctan_add_raw: assumes "\arctan x + arctan y\ < pi/2" shows "arctan x + arctan y = arctan((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) 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) qed lemma arctan_inverse: assumes "0 < x" shows "arctan(inverse x) = pi/2 - arctan x" proof - have "arctan(inverse x) = arctan(inverse(tan(arctan x)))" by (simp add: arctan) also have "... = arctan (tan (pi / 2 - arctan x))" by (simp add: tan_cot) also have "... = pi/2 - arctan x" proof - have "0 < pi - arctan x" using arctan_ubound [of x] pi_gt_zero by linarith with assms show ?thesis by (simp add: Transcendental.arctan_tan) qed finally show ?thesis . qed lemma arctan_add_small: assumes "\x * y\ < 1" shows "(arctan x + arctan y = arctan((x + y) / (1 - x * y)))" proof (cases "x = 0 \ y = 0") case False with assms have "\x\ < inverse \y\" by (simp add: field_split_simps abs_mult) with False have "\arctan x\ < pi / 2 - \arctan y\" using assms by (auto simp add: abs_arctan arctan_inverse [symmetric] arctan_less_iff) then show ?thesis by (intro arctan_add_raw) linarith qed auto lemma abs_arctan_le: fixes x::real shows "\arctan x\ \ \x\" proof - have 1: "\x. x \ \ \ cmod (inverse (1 + x\<^sup>2)) \ 1" by (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square) have "cmod (Arctan w - Arctan z) \ 1 * cmod (w-z)" if "w \ \" "z \ \" for w z apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1]) apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan]) using 1 that by (auto simp: Reals_def) then have "cmod (Arctan (of_real x) - Arctan 0) \ 1 * cmod (of_real x - 0)" using Reals_0 Reals_of_real by blast then show ?thesis by (simp add: Arctan_of_real) qed lemma arctan_le_self: "0 \ x \ arctan x \ x" by (metis abs_arctan_le abs_of_nonneg zero_le_arctan_iff) lemma abs_tan_ge: "\x\ < pi/2 \ \x\ \ \tan x\" by (metis abs_arctan_le abs_less_iff arctan_tan minus_less_iff) lemma arctan_bounds: assumes "0 \ x" "x < 1" shows arctan_lower_bound: "(\k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \ arctan x" (is "(\k<_. (- 1)^ k * ?a k) \ _") and arctan_upper_bound: "arctan x \ (\k<2 * n + 1. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" proof - have tendsto_zero: "?a \ 0" proof (rule tendsto_eq_rhs) show "(\k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \ 0 * 0" using assms by (intro tendsto_mult real_tendsto_divide_at_top) (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top) qed simp have nonneg: "0 \ ?a n" for n by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms) have le: "?a (Suc n) \ ?a n" for n by (rule mult_mono[OF _ power_decreasing]) (auto simp: field_split_simps assms less_imp_le) from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n] summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n] assms show "(\k<2*n. (- 1)^ k * ?a k) \ arctan x" "arctan x \ (\k<2 * n + 1. (- 1)^ k * ?a k)" by (auto simp: arctan_series) qed subsection\<^marker>\tag unimportant\ \Bounds on pi using real arctangent\ lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)" using machin by simp lemma pi_approx: "3.141592653588 \ pi" "pi \ 3.1415926535899" unfolding pi_machin using arctan_bounds[of "1/5" 4] arctan_bounds[of "1/239" 4] by (simp_all add: eval_nat_numeral) lemma pi_gt3: "pi > 3" using pi_approx by simp subsection\Inverse Sine\ definition\<^marker>\tag important\ Arcsin :: "complex \ complex" where "Arcsin \ \z. -\ * Ln(\ * z + csqrt(1 - z\<^sup>2))" lemma Arcsin_body_lemma: "\ * z + csqrt(1 - z\<^sup>2) \ 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) lemma Arcsin_range_lemma: "\Re z\ < 1 \ 0 < Re(\ * z + csqrt(1 - z\<^sup>2))" using Complex.cmod_power2 [of z, symmetric] by (simp add: real_less_rsqrt algebra_simps Re_power2 cmod_square_less_1_plus) lemma Re_Arcsin: "Re(Arcsin z) = Im (Ln (\ * z + csqrt(1 - z\<^sup>2)))" by (simp add: Arcsin_def) lemma Im_Arcsin: "Im(Arcsin z) = - ln (cmod (\ * z + csqrt (1 - z\<^sup>2)))" by (simp add: Arcsin_def Arcsin_body_lemma) lemma one_minus_z2_notin_nonpos_Reals: assumes "Im z = 0 \ \Re z\ < 1" shows "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" proof (cases "Im z = 0") case True with assms show ?thesis by (simp add: complex_nonpos_Reals_iff flip: abs_square_less_1) next case False have "\ (Im z)\<^sup>2 \ - 1" using False power2_less_eq_zero_iff by fastforce with False show ?thesis by (auto simp add: complex_nonpos_Reals_iff Re_power2 Im_power2) qed lemma isCont_Arcsin_lemma: assumes le0: "Re (\ * z + csqrt (1 - z\<^sup>2)) \ 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp: cmod_def abs_square_less_1 [symmetric]) next case False have leim: "(cmod (1 - z\<^sup>2) + (1 - Re (z\<^sup>2))) / 2 \ (Im z)\<^sup>2" using le0 sqrt_le_D by fastforce have neq: "(cmod z)\<^sup>2 \ 1 + cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 = 1 + sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2" using leim cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1] by (simp add: norm_power Re_power2 norm_minus_commute [of 1]) ultimately show False by (simp add: Re_power2 Im_power2 cmod_power2) qed lemma isCont_Arcsin: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arcsin z" proof - have 1: "\ * z + csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff) have 2: "1 - z\<^sup>2 \ \\<^sub>\\<^sub>0" by (simp add: one_minus_z2_notin_nonpos_Reals assms) show ?thesis using assms unfolding Arcsin_def by (intro isCont_Ln' isCont_csqrt' continuous_intros 1 2) qed lemma isCont_Arcsin' [simp]: shows "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arcsin (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arcsin]) lemma sin_Arcsin [simp]: "sin(Arcsin z) = z" proof - have "\*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \ (\*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ (\*z) + csqrt (1 - z\<^sup>2) = 0" by (metis Arcsin_body_lemma distrib_right no_zero_divisors zero_neq_numeral) ultimately show ?thesis apply (simp add: sin_exp_eq Arcsin_def Arcsin_body_lemma exp_minus divide_simps) apply (simp add: algebra_simps) apply (simp add: power2_eq_square [symmetric] algebra_simps) done qed lemma Re_eq_pihalf_lemma: "\Re z\ = pi/2 \ Im z = 0 \ Re ((exp (\*z) + inverse (exp (\*z))) / 2) = 0 \ 0 \ Im ((exp (\*z) + inverse (exp (\*z))) / 2)" apply (simp add: cos_i_times [symmetric] Re_cos Im_cos abs_if del: eq_divide_eq_numeral1) by (metis cos_minus cos_pi_half) lemma Re_less_pihalf_lemma: assumes "\Re z\ < pi / 2" shows "0 < Re ((exp (\*z) + inverse (exp (\*z))) / 2)" proof - have "0 < cos (Re z)" using assms using cos_gt_zero_pi by auto then show ?thesis by (simp add: cos_i_times [symmetric] Re_cos Im_cos add_pos_pos) qed lemma Arcsin_sin: assumes "\Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)" shows "Arcsin(sin z) = z" proof - have "Arcsin(sin z) = - (\ * Ln (csqrt (1 - (\ * (exp (\*z) - inverse (exp (\*z))))\<^sup>2 / 4) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: sin_exp_eq Arcsin_def exp_minus power_divide) also have "... = - (\ * Ln (csqrt (((exp (\*z) + inverse (exp (\*z)))/2)\<^sup>2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" by (simp add: field_simps power2_eq_square) also have "... = - (\ * Ln (((exp (\*z) + inverse (exp (\*z)))/2) - (inverse (exp (\*z)) - exp (\*z)) / 2))" apply (subst csqrt_square) using assms Re_eq_pihalf_lemma Re_less_pihalf_lemma by auto also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms by (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm) finally show ?thesis . qed lemma Arcsin_unique: "\sin z = w; \Re z\ < pi/2 \ (\Re z\ = pi/2 \ Im z = 0)\ \ Arcsin w = z" by (metis Arcsin_sin) lemma Arcsin_0 [simp]: "Arcsin 0 = 0" by (metis Arcsin_sin norm_zero pi_half_gt_zero real_norm_def sin_zero zero_complex.simps(1)) lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2" by (metis Arcsin_sin Im_complex_of_real Re_complex_of_real numeral_One of_real_numeral pi_half_ge_zero real_sqrt_abs real_sqrt_pow2 real_sqrt_power sin_of_real sin_pi_half) lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)" by (metis Arcsin_1 Arcsin_sin Im_complex_of_real Re_complex_of_real abs_of_nonneg of_real_minus pi_half_ge_zero power2_minus real_sqrt_abs sin_Arcsin sin_minus) lemma has_field_derivative_Arcsin: assumes "Im z = 0 \ \Re z\ < 1" shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)" proof - have "(sin (Arcsin z))\<^sup>2 \ 1" using assms one_minus_z2_notin_nonpos_Reals by force then have "cos (Arcsin z) \ 0" by (metis diff_0_right power_zero_numeral sin_squared_eq) then show ?thesis by (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) (auto intro: isCont_Arcsin assms) qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable at z" using field_differentiable_def has_field_derivative_Arcsin by blast lemma field_differentiable_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ Arcsin field_differentiable (at z within s)" using field_differentiable_at_Arcsin field_differentiable_within_subset by blast lemma continuous_within_Arcsin: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arcsin" using continuous_at_imp_continuous_within isCont_Arcsin by blast lemma continuous_on_Arcsin [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arcsin" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arcsin: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arcsin holomorphic_on s" by (simp add: field_differentiable_within_Arcsin holomorphic_on_def) subsection\Inverse Cosine\ definition\<^marker>\tag important\ Arccos :: "complex \ complex" where "Arccos \ \z. -\ * Ln(z + \ * csqrt(1 - z\<^sup>2))" lemma Arccos_range_lemma: "\Re z\ < 1 \ 0 < Im(z + \ * csqrt(1 - z\<^sup>2))" using Arcsin_range_lemma [of "-z"] by simp lemma Arccos_body_lemma: "z + \ * csqrt(1 - z\<^sup>2) \ 0" using Arcsin_body_lemma [of z] by (metis Arcsin_body_lemma complex_i_mult_minus diff_minus_eq_add power2_minus right_minus_eq) lemma Re_Arccos: "Re(Arccos z) = Im (Ln (z + \ * csqrt(1 - z\<^sup>2)))" by (simp add: Arccos_def) lemma Im_Arccos: "Im(Arccos z) = - ln (cmod (z + \ * csqrt (1 - z\<^sup>2)))" by (simp add: Arccos_def Arccos_body_lemma) text\A very tricky argument to find!\ lemma isCont_Arccos_lemma: assumes eq0: "Im (z + \ * csqrt (1 - z\<^sup>2)) = 0" and "(Im z = 0 \ \Re z\ < 1)" shows False proof (cases "Im z = 0") case True then show ?thesis using assms by (fastforce simp add: cmod_def abs_square_less_1 [symmetric]) next case False have Imz: "Im z = - sqrt ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)" using eq0 abs_Re_le_cmod [of "1-z\<^sup>2"] by (simp add: Re_power2 algebra_simps) have "(cmod z)\<^sup>2 - 1 \ cmod (1 - z\<^sup>2)" proof (clarsimp simp add: cmod_def) assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 - 1 = sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" then have "((Re z)\<^sup>2 + (Im z)\<^sup>2 - 1)\<^sup>2 = ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)" by simp then show False using False by (simp add: power2_eq_square algebra_simps) qed moreover have "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2" using abs_Re_le_cmod [of "1-z\<^sup>2"] by (subst Imz) (simp add: Re_power2) ultimately show False by (simp add: cmod_power2) qed lemma isCont_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "isCont Arccos z" proof - have "z + \ * csqrt (1 - z\<^sup>2) \ \\<^sub>\\<^sub>0" by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms) with assms show ?thesis unfolding Arccos_def by (simp_all add: one_minus_z2_notin_nonpos_Reals assms) qed lemma isCont_Arccos' [simp]: "isCont f z \ (Im (f z) = 0 \ \Re (f z)\ < 1) \ isCont (\x. Arccos (f x)) z" by (blast intro: isCont_o2 [OF _ isCont_Arccos]) lemma cos_Arccos [simp]: "cos(Arccos z) = z" proof - have "z*2 + \ * (2 * csqrt (1 - z\<^sup>2)) = 0 \ z*2 + \ * csqrt (1 - z\<^sup>2)*2 = 0" by (simp add: algebra_simps) \ \Cancelling a factor of 2\ moreover have "... \ z + \ * csqrt (1 - z\<^sup>2) = 0" by (metis distrib_right mult_eq_0_iff zero_neq_numeral) ultimately show ?thesis by (simp add: cos_exp_eq Arccos_def Arccos_body_lemma exp_minus field_simps flip: power2_eq_square) qed lemma Arccos_cos: assumes "0 < Re z \ Re z < pi \ Re z = 0 \ 0 \ Im z \ Re z = pi \ Im z \ 0" shows "Arccos(cos z) = z" proof - have *: "((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z))) = sin z" by (simp add: sin_exp_eq exp_minus field_simps power2_eq_square) have "1 - (exp (\ * z) + inverse (exp (\ * z)))\<^sup>2 / 4 = ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2" by (simp add: field_simps power2_eq_square) then have "Arccos(cos z) = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * csqrt (((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))\<^sup>2)))" by (simp add: cos_exp_eq Arccos_def exp_minus power_divide) also have "... = - (\ * Ln ((exp (\ * z) + inverse (exp (\ * z))) / 2 + \ * ((\ - (exp (\ * z))\<^sup>2 * \) / (2 * exp (\ * z)))))" apply (subst csqrt_square) using assms Re_sin_pos [of z] Im_sin_nonneg [of z] Im_sin_nonneg2 [of z] by (auto simp: * Re_sin Im_sin) also have "... = - (\ * Ln (exp (\*z)))" by (simp add: field_simps power2_eq_square) also have "... = z" using assms by (subst Complex_Transcendental.Ln_exp, auto) finally show ?thesis . qed lemma Arccos_unique: "\cos z = w; 0 < Re z \ Re z < pi \ Re z = 0 \ 0 \ Im z \ Re z = pi \ Im z \ 0\ \ Arccos w = z" using Arccos_cos by blast lemma Arccos_0 [simp]: "Arccos 0 = pi/2" by (rule Arccos_unique) auto lemma Arccos_1 [simp]: "Arccos 1 = 0" by (rule Arccos_unique) auto lemma Arccos_minus1: "Arccos(-1) = pi" by (rule Arccos_unique) auto lemma has_field_derivative_Arccos: assumes "(Im z = 0 \ \Re z\ < 1)" shows "(Arccos has_field_derivative - inverse(sin(Arccos z))) (at z)" proof - have "x\<^sup>2 \ -1" for x::real by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x__]^2)))))") with assms have "(cos (Arccos z))\<^sup>2 \ 1" by (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1) then have "- sin (Arccos z) \ 0" by (metis cos_squared_eq diff_0_right mult_zero_left neg_0_equal_iff_equal power2_eq_square) then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)" by (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]]) (auto intro: isCont_Arccos assms) then show ?thesis by simp qed declare has_field_derivative_Arcsin [derivative_intros] declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros] lemma field_differentiable_at_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable at z" using field_differentiable_def has_field_derivative_Arccos by blast lemma field_differentiable_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ Arccos field_differentiable (at z within s)" using field_differentiable_at_Arccos field_differentiable_within_subset by blast lemma continuous_within_Arccos: "(Im z = 0 \ \Re z\ < 1) \ continuous (at z within s) Arccos" using continuous_at_imp_continuous_within isCont_Arccos by blast lemma continuous_on_Arccos [continuous_intros]: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ continuous_on s Arccos" by (simp add: continuous_at_imp_continuous_on) lemma holomorphic_on_Arccos: "(\z. z \ s \ Im z = 0 \ \Re z\ < 1) \ Arccos holomorphic_on s" by (simp add: field_differentiable_within_Arccos holomorphic_on_def) subsection\<^marker>\tag unimportant\\Upper and Lower Bounds for Inverse Sine and Cosine\ lemma Arcsin_bounds: "\Re z\ < 1 \ \Re(Arcsin z)\ < pi/2" unfolding Re_Arcsin by (blast intro: Re_Ln_pos_lt_imp Arcsin_range_lemma) lemma Arccos_bounds: "\Re z\ < 1 \ 0 < Re(Arccos z) \ Re(Arccos z) < pi" unfolding Re_Arccos by (blast intro!: Im_Ln_pos_lt_imp Arccos_range_lemma) lemma Re_Arccos_bounds: "-pi < Re(Arccos z) \ Re(Arccos z) \ pi" unfolding Re_Arccos by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arccos_body_lemma) lemma Re_Arccos_bound: "\Re(Arccos z)\ \ pi" by (meson Re_Arccos_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma Im_Arccos_bound: "\Im (Arccos w)\ \ cmod w" proof - have "(Im (Arccos w))\<^sup>2 \ (cmod (cos (Arccos w)))\<^sup>2 - (cos (Re (Arccos w)))\<^sup>2" using norm_cos_squared [of "Arccos w"] real_le_abs_sinh [of "Im (Arccos w)"] by (simp only: abs_le_square_iff) (simp add: field_split_simps) also have "... \ (cmod w)\<^sup>2" by (auto simp: cmod_power2) finally show ?thesis using abs_le_square_iff by force qed lemma Re_Arcsin_bounds: "-pi < Re(Arcsin z) & Re(Arcsin z) \ pi" unfolding Re_Arcsin by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arcsin_body_lemma) lemma Re_Arcsin_bound: "\Re(Arcsin z)\ \ pi" by (meson Re_Arcsin_bounds abs_le_iff less_eq_real_def minus_less_iff) lemma norm_Arccos_bounded: fixes w :: complex shows "norm (Arccos w) \ pi + norm w" proof - have Re: "(Re (Arccos w))\<^sup>2 \ pi\<^sup>2" "(Im (Arccos w))\<^sup>2 \ (cmod w)\<^sup>2" using Re_Arccos_bound [of w] Im_Arccos_bound [of w] abs_le_square_iff by force+ have "Arccos w \ Arccos w \ pi\<^sup>2 + (cmod w)\<^sup>2" using Re by (simp add: dot_square_norm cmod_power2 [of "Arccos w"]) then have "cmod (Arccos w) \ pi + cmod (cos (Arccos w))" apply (simp add: norm_le_square) by (metis dot_square_norm norm_ge_zero norm_le_square pi_ge_zero triangle_lemma) then show "cmod (Arccos w) \ pi + cmod w" by auto qed subsection\<^marker>\tag unimportant\\Interrelations between Arcsin and Arccos\ lemma cos_Arcsin_nonzero: assumes "z\<^sup>2 \ 1" shows "cos(Arcsin z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)" by (simp add: algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ z\<^sup>2 - 1" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (z\<^sup>2 - 1)\<^sup>2" by simp then have "z\<^sup>2 * (z\<^sup>2 - 1) = (z\<^sup>2 - 1)*(z\<^sup>2 - 1)" using eq power2_eq_square by auto then show False using assms by simp qed then have "1 + \ * z * (csqrt (1 - z * z)) \ z\<^sup>2" by (metis add_minus_cancel power2_eq_square uminus_add_conv_diff) then have "2*(1 + \ * z * (csqrt (1 - z * z))) \ 2*z\<^sup>2" (*FIXME cancel_numeral_factor*) by (metis mult_cancel_left zero_neq_numeral) then have "(\ * z + csqrt (1 - z\<^sup>2))\<^sup>2 \ -1" using assms apply (simp add: power2_sum) apply (simp add: power2_eq_square algebra_simps) done then show ?thesis apply (simp add: cos_exp_eq Arcsin_def exp_minus) apply (simp add: divide_simps Arcsin_body_lemma) apply (metis add.commute minus_unique power2_eq_square) done qed lemma sin_Arccos_nonzero: assumes "z\<^sup>2 \ 1" shows "sin(Arccos z) \ 0" proof - have eq: "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)" by (simp add: algebra_simps) have "\ * z * (csqrt (1 - z\<^sup>2)) \ 1 - z\<^sup>2" proof assume "\ * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2" then have "(\ * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (1 - z\<^sup>2)\<^sup>2" by simp then have "-(z\<^sup>2) * (1 - z\<^sup>2) = (1 - z\<^sup>2)*(1 - z\<^sup>2)" using eq power2_eq_square by auto then have "-(z\<^sup>2) = (1 - z\<^sup>2)" using assms by (metis add.commute add.right_neutral diff_add_cancel mult_right_cancel) then show False using assms by simp qed then have "z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2)) \ 1" by (simp add: algebra_simps) then have "2*(z\<^sup>2 + \ * z * (csqrt (1 - z\<^sup>2))) \ 2*1" by (metis mult_cancel_left2 zero_neq_numeral) (*FIXME cancel_numeral_factor*) then have "(z + \ * csqrt (1 - z\<^sup>2))\<^sup>2 \ 1" using assms by (metis Arccos_def add.commute add.left_neutral cancel_comm_monoid_add_class.diff_cancel cos_Arccos csqrt_0 mult_zero_right) then show ?thesis apply (simp add: sin_exp_eq Arccos_def exp_minus) apply (simp add: divide_simps Arccos_body_lemma) apply (simp add: power2_eq_square) done qed lemma cos_sin_csqrt: assumes "0 < cos(Re z) \ cos(Re z) = 0 \ Im z * sin(Re z) \ 0" shows "cos z = csqrt(1 - (sin z)\<^sup>2)" proof (rule csqrt_unique [THEN sym]) show "(cos z)\<^sup>2 = 1 - (sin z)\<^sup>2" by (simp add: cos_squared_eq) qed (use assms in \auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff\) lemma sin_cos_csqrt: assumes "0 < sin(Re z) \ sin(Re z) = 0 \ 0 \ Im z * cos(Re z)" shows "sin z = csqrt(1 - (cos z)\<^sup>2)" proof (rule csqrt_unique [THEN sym]) show "(sin z)\<^sup>2 = 1 - (cos z)\<^sup>2" by (simp add: sin_squared_eq) qed (use assms in \auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff\) lemma Arcsin_Arccos_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arcsin z = Arccos(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma Arccos_Arcsin_csqrt_pos: "(0 < Re z | Re z = 0 & 0 \ Im z) \ Arccos z = Arcsin(csqrt(1 - z\<^sup>2))" by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute) lemma sin_Arccos: "0 < Re z | Re z = 0 & 0 \ Im z \ sin(Arccos z) = csqrt(1 - z\<^sup>2)" by (simp add: Arccos_Arcsin_csqrt_pos) lemma cos_Arcsin: "0 < Re z | Re z = 0 & 0 \ Im z \ cos(Arcsin z) = csqrt(1 - z\<^sup>2)" by (simp add: Arcsin_Arccos_csqrt_pos) subsection\<^marker>\tag unimportant\\Relationship with Arcsin on the Real Numbers\ lemma Im_Arcsin_of_real: assumes "\x\ \ 1" shows "Im (Arcsin (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (\ * of_real x + csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arcsin exp_minus) qed corollary\<^marker>\tag unimportant\ Arcsin_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arcsin z \ \" by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arcsin_eq_Re_Arcsin: assumes "\x\ \ 1" shows "arcsin x = Re (Arcsin (of_real x))" unfolding arcsin_def proof (rule the_equality, safe) show "- (pi / 2) \ Re (Arcsin (complex_of_real x))" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "Re (Arcsin (complex_of_real x)) \ pi / 2" using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arcsin) next show "sin (Re (Arcsin (complex_of_real x))) = x" using Re_sin [of "Arcsin (of_real x)"] Arcsin_body_lemma [of "of_real x"] by (simp add: Im_Arcsin_of_real assms) next fix x' assume "- (pi / 2) \ x'" "x' \ pi / 2" "x = sin x'" then show "x' = Re (Arcsin (complex_of_real (sin x')))" unfolding sin_of_real [symmetric] by (subst Arcsin_sin) auto qed lemma of_real_arcsin: "\x\ \ 1 \ of_real(arcsin x) = Arcsin(of_real x)" by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Relationship with Arccos on the Real Numbers\ lemma Im_Arccos_of_real: assumes "\x\ \ 1" shows "Im (Arccos (of_real x)) = 0" proof - have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by (simp add: of_real_sqrt del: csqrt_of_real_nonneg) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2))^2 = 1" using assms abs_square_le_1 by (force simp add: Complex.cmod_power2) then have "cmod (of_real x + \ * csqrt (1 - (of_real x)\<^sup>2)) = 1" by (simp add: norm_complex_def) then show ?thesis by (simp add: Im_Arccos exp_minus) qed corollary\<^marker>\tag unimportant\ Arccos_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arccos z \ \" by (metis Im_Arccos_of_real Re_complex_of_real Reals_cases complex_is_Real_iff) lemma arccos_eq_Re_Arccos: assumes "\x\ \ 1" shows "arccos x = Re (Arccos (of_real x))" unfolding arccos_def proof (rule the_equality, safe) show "0 \ Re (Arccos (complex_of_real x))" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "Re (Arccos (complex_of_real x)) \ pi" using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"] by (auto simp: Complex.in_Reals_norm Re_Arccos) next show "cos (Re (Arccos (complex_of_real x))) = x" using Re_cos [of "Arccos (of_real x)"] Arccos_body_lemma [of "of_real x"] by (simp add: Im_Arccos_of_real assms) next fix x' assume "0 \ x'" "x' \ pi" "x = cos x'" then show "x' = Re (Arccos (complex_of_real (cos x')))" unfolding cos_of_real [symmetric] by (subst Arccos_cos) auto qed lemma of_real_arccos: "\x\ \ 1 \ of_real(arccos x) = Arccos(of_real x)" by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0) subsection\<^marker>\tag unimportant\\Some interrelationships among the real inverse trig functions\ lemma arccos_arctan: assumes "-1 < x" "x < 1" shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))" proof - have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0" proof (rule sin_eq_0_pi) show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)" using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi" using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"] arccos_bounded [of x] assms by (simp add: algebra_simps) next show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0" using assms by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan power2_eq_square square_eq_1_iff) qed then show ?thesis by simp qed lemma arcsin_plus_arccos: assumes "-1 \ x" "x \ 1" shows "arcsin x + arccos x = pi/2" proof - have "arcsin x = pi/2 - arccos x" apply (rule sin_inj_pi) using assms arcsin [OF assms] arccos [OF assms] by (auto simp: algebra_simps sin_diff) then show ?thesis by (simp add: algebra_simps) qed lemma arcsin_arccos_eq: "-1 \ x \ x \ 1 \ arcsin x = pi/2 - arccos x" using arcsin_plus_arccos by force lemma arccos_arcsin_eq: "-1 \ x \ x \ 1 \ arccos x = pi/2 - arcsin x" using arcsin_plus_arccos by force lemma arcsin_arctan: "-1 < x \ x < 1 \ arcsin x = arctan(x / sqrt(1 - x\<^sup>2))" by (simp add: arccos_arctan arcsin_arccos_eq) lemma csqrt_1_diff_eq: "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \ 1 then sqrt (1 - x^2) else \ * sqrt (x^2 - 1))" by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg) lemma arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arcsin_Arccos_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arcsin_arccos_sqrt_neg: "-1 \ x \ x \ 0 \ arcsin x = -arccos(sqrt(1 - x\<^sup>2))" using arcsin_arccos_sqrt_pos [of "-x"] by (simp add: arcsin_minus) lemma arccos_arcsin_sqrt_pos: "0 \ x \ x \ 1 \ arccos x = arcsin(sqrt(1 - x\<^sup>2))" apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_Arccos) apply (subst Arccos_Arcsin_csqrt_pos) apply (auto simp: power_le_one csqrt_1_diff_eq) done lemma arccos_arcsin_sqrt_neg: "-1 \ x \ x \ 0 \ arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))" using arccos_arcsin_sqrt_pos [of "-x"] by (simp add: arccos_minus) subsection\<^marker>\tag unimportant\\Continuity results for arcsin and arccos\ lemma continuous_on_Arcsin_real [continuous_intros]: "continuous_on {w \ \. \Re w\ \ 1} Arcsin" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arcsin (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arcsin (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arcsin_eq_Re_Arcsin) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arcsin [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arcsin_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arcsin" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arcsin_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma continuous_on_Arccos_real: "continuous_on {w \ \. \Re w\ \ 1} Arccos" proof - have "continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (arccos (Re x))) = continuous_on {w \ \. \Re w\ \ 1} (\x. complex_of_real (Re (Arccos (of_real (Re x)))))" by (rule continuous_on_cong [OF refl]) (simp add: arccos_eq_Re_Arccos) also have "... = ?thesis" by (rule continuous_on_cong [OF refl]) simp finally show ?thesis using continuous_on_arccos [OF continuous_on_Re [OF continuous_on_id], of "{w \ \. \Re w\ \ 1}"] continuous_on_of_real by fastforce qed lemma continuous_within_Arccos_real: "continuous (at z within {w \ \. \Re w\ \ 1}) Arccos" proof (cases "z \ {w \ \. \Re w\ \ 1}") case True then show ?thesis using continuous_on_Arccos_real continuous_on_eq_continuous_within by blast next case False with closed_real_abs_le [of 1] show ?thesis by (rule continuous_within_closed_nontrivial) qed lemma sinh_ln_complex: "x \ 0 \ sinh (ln x :: complex) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus scaleR_conv_of_real exp_of_real) lemma cosh_ln_complex: "x \ 0 \ cosh (ln x :: complex) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus scaleR_conv_of_real) lemma tanh_ln_complex: "x \ 0 \ tanh (ln x :: complex) = (x ^ 2 - 1) / (x ^ 2 + 1)" by (simp add: tanh_def sinh_ln_complex cosh_ln_complex divide_simps power2_eq_square) subsection\Roots of unity\ theorem complex_root_unity: fixes j::nat assumes "n \ 0" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n)^n = 1" proof - have *: "of_nat j * (complex_of_real pi * 2) = complex_of_real (2 * real j * pi)" by (simp) then show ?thesis apply (simp add: exp_of_nat_mult [symmetric] mult_ac exp_Euler) apply (simp only: * cos_of_real sin_of_real) apply simp done qed lemma complex_root_unity_eq: fixes j::nat and k::nat assumes "1 \ n" shows "(exp(2 * of_real pi * \ * of_nat j / of_nat n) = exp(2 * of_real pi * \ * of_nat k / of_nat n) \ j mod n = k mod n)" proof - have "(\z::int. \ * (of_nat j * (of_real pi * 2)) = \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ (\z::int. of_nat j * (\ * (of_real pi * 2)) = (of_nat k + of_nat n * of_int z) * (\ * (of_real pi * 2)))" by (simp add: algebra_simps) also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * (of_int z :: complex))" by simp also have "... \ (\z::int. of_nat j = of_nat k + of_nat n * z)" by (metis (mono_tags, hide_lams) of_int_add of_int_eq_iff of_int_mult of_int_of_nat_eq) also have "... \ int j mod int n = int k mod int n" by (auto simp: mod_eq_dvd_iff dvd_def algebra_simps) also have "... \ j mod n = k mod n" by (metis of_nat_eq_iff zmod_int) finally have "(\z. \ * (of_nat j * (of_real pi * 2)) = \ * (of_nat k * (of_real pi * 2)) + \ * (of_int z * (of_nat n * (of_real pi * 2)))) \ j mod n = k mod n" . note * = this show ?thesis using assms by (simp add: exp_eq field_split_simps *) qed corollary bij_betw_roots_unity: "bij_betw (\j. exp(2 * of_real pi * \ * of_nat j / of_nat n)) {.. * of_nat j / of_nat n) | j. j < n}" by (auto simp: bij_betw_def inj_on_def complex_root_unity_eq) lemma complex_root_unity_eq_1: fixes j::nat and k::nat assumes "1 \ n" shows "exp(2 * of_real pi * \ * of_nat j / of_nat n) = 1 \ n dvd j" proof - have "1 = exp(2 * of_real pi * \ * (of_nat n / of_nat n))" using assms by simp then have "exp(2 * of_real pi * \ * (of_nat j / of_nat n)) = 1 \ j mod n = n mod n" using complex_root_unity_eq [of n j n] assms by simp then show ?thesis by auto qed lemma finite_complex_roots_unity_explicit: "finite {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n}" by simp lemma card_complex_roots_unity_explicit: "card {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j::nat. j < n} = n" by (simp add: Finite_Set.bij_betw_same_card [OF bij_betw_roots_unity, symmetric]) lemma complex_roots_unity: assumes "1 \ n" shows "{z::complex. z^n = 1} = {exp(2 * of_real pi * \ * of_nat j / of_nat n) | j. j < n}" apply (rule Finite_Set.card_seteq [symmetric]) using assms apply (auto simp: card_complex_roots_unity_explicit finite_roots_unity complex_root_unity card_roots_unity) done lemma card_complex_roots_unity: "1 \ n \ card {z::complex. z^n = 1} = n" by (simp add: card_complex_roots_unity_explicit complex_roots_unity) lemma complex_not_root_unity: "1 \ n \ \u::complex. norm u = 1 \ u^n \ 1" apply (rule_tac x="exp (of_real pi * \ * of_real (1 / n))" in exI) apply (auto simp: Re_complex_div_eq_0 exp_of_nat_mult [symmetric] mult_ac exp_Euler) done end diff --git a/src/HOL/Complex.thy b/src/HOL/Complex.thy --- a/src/HOL/Complex.thy +++ b/src/HOL/Complex.thy @@ -1,1324 +1,1327 @@ (* Title: HOL/Complex.thy Author: Jacques D. Fleuriot, 2001 University of Edinburgh Author: Lawrence C Paulson, 2003/4 *) section \Complex Numbers: Rectangular and Polar Representations\ theory Complex imports Transcendental begin text \ We use the \<^theory_text>\codatatype\ command to define the type of complex numbers. This allows us to use \<^theory_text>\primcorec\ to define complex functions by defining their real and imaginary result separately. \ codatatype complex = Complex (Re: real) (Im: real) lemma complex_surj: "Complex (Re z) (Im z) = z" by (rule complex.collapse) lemma complex_eqI [intro?]: "Re x = Re y \ Im x = Im y \ x = y" by (rule complex.expand) simp lemma complex_eq_iff: "x = y \ Re x = Re y \ Im x = Im y" by (auto intro: complex.expand) subsection \Addition and Subtraction\ instantiation complex :: ab_group_add begin primcorec zero_complex where "Re 0 = 0" | "Im 0 = 0" primcorec plus_complex where "Re (x + y) = Re x + Re y" | "Im (x + y) = Im x + Im y" primcorec uminus_complex where "Re (- x) = - Re x" | "Im (- x) = - Im x" primcorec minus_complex where "Re (x - y) = Re x - Re y" | "Im (x - y) = Im x - Im y" instance by standard (simp_all add: complex_eq_iff) end subsection \Multiplication and Division\ instantiation complex :: field begin primcorec one_complex where "Re 1 = 1" | "Im 1 = 0" primcorec times_complex where "Re (x * y) = Re x * Re y - Im x * Im y" | "Im (x * y) = Re x * Im y + Im x * Re y" primcorec inverse_complex where "Re (inverse x) = Re x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)" | "Im (inverse x) = - Im x / ((Re x)\<^sup>2 + (Im x)\<^sup>2)" definition "x div y = x * inverse y" for x y :: complex instance by standard (simp_all add: complex_eq_iff divide_complex_def distrib_left distrib_right right_diff_distrib left_diff_distrib power2_eq_square add_divide_distrib [symmetric]) end lemma Re_divide: "Re (x / y) = (Re x * Re y + Im x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)" by (simp add: divide_complex_def add_divide_distrib) lemma Im_divide: "Im (x / y) = (Im x * Re y - Re x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)" by (simp add: divide_complex_def diff_divide_distrib) lemma Complex_divide: "(x / y) = Complex ((Re x * Re y + Im x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2)) ((Im x * Re y - Re x * Im y) / ((Re y)\<^sup>2 + (Im y)\<^sup>2))" by (metis Im_divide Re_divide complex_surj) lemma Re_power2: "Re (x ^ 2) = (Re x)^2 - (Im x)^2" by (simp add: power2_eq_square) lemma Im_power2: "Im (x ^ 2) = 2 * Re x * Im x" by (simp add: power2_eq_square) lemma Re_power_real [simp]: "Im x = 0 \ Re (x ^ n) = Re x ^ n " by (induct n) simp_all lemma Im_power_real [simp]: "Im x = 0 \ Im (x ^ n) = 0" by (induct n) simp_all subsection \Scalar Multiplication\ instantiation complex :: real_field begin primcorec scaleR_complex where "Re (scaleR r x) = r * Re x" | "Im (scaleR r x) = r * Im x" instance proof fix a b :: real and x y :: complex show "scaleR a (x + y) = scaleR a x + scaleR a y" by (simp add: complex_eq_iff distrib_left) show "scaleR (a + b) x = scaleR a x + scaleR b x" by (simp add: complex_eq_iff distrib_right) show "scaleR a (scaleR b x) = scaleR (a * b) x" by (simp add: complex_eq_iff mult.assoc) show "scaleR 1 x = x" by (simp add: complex_eq_iff) show "scaleR a x * y = scaleR a (x * y)" by (simp add: complex_eq_iff algebra_simps) show "x * scaleR a y = scaleR a (x * y)" by (simp add: complex_eq_iff algebra_simps) qed end subsection \Numerals, Arithmetic, and Embedding from R\ abbreviation complex_of_real :: "real \ complex" where "complex_of_real \ of_real" declare [[coercion "of_real :: real \ complex"]] declare [[coercion "of_rat :: rat \ complex"]] declare [[coercion "of_int :: int \ complex"]] declare [[coercion "of_nat :: nat \ complex"]] lemma complex_Re_of_nat [simp]: "Re (of_nat n) = of_nat n" by (induct n) simp_all lemma complex_Im_of_nat [simp]: "Im (of_nat n) = 0" by (induct n) simp_all lemma complex_Re_of_int [simp]: "Re (of_int z) = of_int z" by (cases z rule: int_diff_cases) simp lemma complex_Im_of_int [simp]: "Im (of_int z) = 0" by (cases z rule: int_diff_cases) simp lemma complex_Re_numeral [simp]: "Re (numeral v) = numeral v" using complex_Re_of_int [of "numeral v"] by simp lemma complex_Im_numeral [simp]: "Im (numeral v) = 0" using complex_Im_of_int [of "numeral v"] by simp lemma Re_complex_of_real [simp]: "Re (complex_of_real z) = z" by (simp add: of_real_def) lemma Im_complex_of_real [simp]: "Im (complex_of_real z) = 0" by (simp add: of_real_def) lemma Re_divide_numeral [simp]: "Re (z / numeral w) = Re z / numeral w" by (simp add: Re_divide sqr_conv_mult) lemma Im_divide_numeral [simp]: "Im (z / numeral w) = Im z / numeral w" by (simp add: Im_divide sqr_conv_mult) lemma Re_divide_of_nat [simp]: "Re (z / of_nat n) = Re z / of_nat n" by (cases n) (simp_all add: Re_divide field_split_simps power2_eq_square del: of_nat_Suc) lemma Im_divide_of_nat [simp]: "Im (z / of_nat n) = Im z / of_nat n" by (cases n) (simp_all add: Im_divide field_split_simps power2_eq_square del: of_nat_Suc) lemma of_real_Re [simp]: "z \ \ \ of_real (Re z) = z" by (auto simp: Reals_def) lemma complex_Re_fact [simp]: "Re (fact n) = fact n" proof - have "(fact n :: complex) = of_real (fact n)" by simp also have "Re \ = fact n" by (subst Re_complex_of_real) simp_all finally show ?thesis . qed lemma complex_Im_fact [simp]: "Im (fact n) = 0" by (subst of_nat_fact [symmetric]) (simp only: complex_Im_of_nat) lemma Re_prod_Reals: "(\x. x \ A \ f x \ \) \ Re (prod f A) = prod (\x. Re (f x)) A" proof (induction A rule: infinite_finite_induct) case (insert x A) hence "Re (prod f (insert x A)) = Re (f x) * Re (prod f A) - Im (f x) * Im (prod f A)" by simp also from insert.prems have "f x \ \" by simp hence "Im (f x) = 0" by (auto elim!: Reals_cases) also have "Re (prod f A) = (\x\A. Re (f x))" by (intro insert.IH insert.prems) auto finally show ?case using insert.hyps by simp qed auto subsection \The Complex Number $i$\ primcorec imaginary_unit :: complex ("\") where "Re \ = 0" | "Im \ = 1" lemma Complex_eq: "Complex a b = a + \ * b" by (simp add: complex_eq_iff) lemma complex_eq: "a = Re a + \ * Im a" by (simp add: complex_eq_iff) lemma fun_complex_eq: "f = (\x. Re (f x) + \ * Im (f x))" by (simp add: fun_eq_iff complex_eq) lemma i_squared [simp]: "\ * \ = -1" by (simp add: complex_eq_iff) lemma power2_i [simp]: "\\<^sup>2 = -1" by (simp add: power2_eq_square) lemma inverse_i [simp]: "inverse \ = - \" by (rule inverse_unique) simp lemma divide_i [simp]: "x / \ = - \ * x" by (simp add: divide_complex_def) lemma complex_i_mult_minus [simp]: "\ * (\ * x) = - x" by (simp add: mult.assoc [symmetric]) lemma complex_i_not_zero [simp]: "\ \ 0" by (simp add: complex_eq_iff) lemma complex_i_not_one [simp]: "\ \ 1" by (simp add: complex_eq_iff) lemma complex_i_not_numeral [simp]: "\ \ numeral w" by (simp add: complex_eq_iff) lemma complex_i_not_neg_numeral [simp]: "\ \ - numeral w" by (simp add: complex_eq_iff) lemma complex_split_polar: "\r a. z = complex_of_real r * (cos a + \ * sin a)" by (simp add: complex_eq_iff polar_Ex) lemma i_even_power [simp]: "\ ^ (n * 2) = (-1) ^ n" by (metis mult.commute power2_i power_mult) lemma Re_i_times [simp]: "Re (\ * z) = - Im z" by simp lemma Im_i_times [simp]: "Im (\ * z) = Re z" by simp lemma i_times_eq_iff: "\ * w = z \ w = - (\ * z)" by auto lemma divide_numeral_i [simp]: "z / (numeral n * \) = - (\ * z) / numeral n" by (metis divide_divide_eq_left divide_i mult.commute mult_minus_right) lemma imaginary_eq_real_iff [simp]: assumes "y \ Reals" "x \ Reals" shows "\ * y = x \ x=0 \ y=0" using assms unfolding Reals_def apply clarify apply (rule iffI) apply (metis Im_complex_of_real Im_i_times Re_complex_of_real mult_eq_0_iff of_real_0) by simp lemma real_eq_imaginary_iff [simp]: assumes "y \ Reals" "x \ Reals" shows "x = \ * y \ x=0 \ y=0" using assms imaginary_eq_real_iff by fastforce subsection \Vector Norm\ instantiation complex :: real_normed_field begin definition "norm z = sqrt ((Re z)\<^sup>2 + (Im z)\<^sup>2)" abbreviation cmod :: "complex \ real" where "cmod \ norm" definition complex_sgn_def: "sgn x = x /\<^sub>R cmod x" definition dist_complex_def: "dist x y = cmod (x - y)" definition uniformity_complex_def [code del]: "(uniformity :: (complex \ complex) filter) = (INF e\{0 <..}. principal {(x, y). dist x y < e})" definition open_complex_def [code del]: "open (U :: complex set) \ (\x\U. eventually (\(x', y). x' = x \ y \ U) uniformity)" instance proof fix r :: real and x y :: complex and S :: "complex set" show "(norm x = 0) = (x = 0)" by (simp add: norm_complex_def complex_eq_iff) show "norm (x + y) \ norm x + norm y" by (simp add: norm_complex_def complex_eq_iff real_sqrt_sum_squares_triangle_ineq) show "norm (scaleR r x) = \r\ * norm x" by (simp add: norm_complex_def complex_eq_iff power_mult_distrib distrib_left [symmetric] real_sqrt_mult) show "norm (x * y) = norm x * norm y" by (simp add: norm_complex_def complex_eq_iff real_sqrt_mult [symmetric] power2_eq_square algebra_simps) qed (rule complex_sgn_def dist_complex_def open_complex_def uniformity_complex_def)+ end declare uniformity_Abort[where 'a = complex, code] lemma norm_ii [simp]: "norm \ = 1" by (simp add: norm_complex_def) lemma cmod_unit_one: "cmod (cos a + \ * sin a) = 1" by (simp add: norm_complex_def) lemma cmod_complex_polar: "cmod (r * (cos a + \ * sin a)) = \r\" by (simp add: norm_mult cmod_unit_one) lemma complex_Re_le_cmod: "Re x \ cmod x" unfolding norm_complex_def by (rule real_sqrt_sum_squares_ge1) lemma complex_mod_minus_le_complex_mod: "- cmod x \ cmod x" by (rule order_trans [OF _ norm_ge_zero]) simp lemma complex_mod_triangle_ineq2: "cmod (b + a) - cmod b \ cmod a" by (rule ord_le_eq_trans [OF norm_triangle_ineq2]) simp lemma abs_Re_le_cmod: "\Re x\ \ cmod x" by (simp add: norm_complex_def) lemma abs_Im_le_cmod: "\Im x\ \ cmod x" by (simp add: norm_complex_def) lemma cmod_le: "cmod z \ \Re z\ + \Im z\" apply (subst complex_eq) apply (rule order_trans) apply (rule norm_triangle_ineq) apply (simp add: norm_mult) done lemma cmod_eq_Re: "Im z = 0 \ cmod z = \Re z\" by (simp add: norm_complex_def) lemma cmod_eq_Im: "Re z = 0 \ cmod z = \Im z\" by (simp add: norm_complex_def) lemma cmod_power2: "(cmod z)\<^sup>2 = (Re z)\<^sup>2 + (Im z)\<^sup>2" by (simp add: norm_complex_def) lemma cmod_plus_Re_le_0_iff: "cmod z + Re z \ 0 \ Re z = - cmod z" using abs_Re_le_cmod[of z] by auto lemma cmod_Re_le_iff: "Im x = Im y \ cmod x \ cmod y \ \Re x\ \ \Re y\" by (metis add.commute add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) lemma cmod_Im_le_iff: "Re x = Re y \ cmod x \ cmod y \ \Im x\ \ \Im y\" by (metis add_le_cancel_left norm_complex_def real_sqrt_abs real_sqrt_le_iff) lemma Im_eq_0: "\Re z\ = cmod z \ Im z = 0" by (subst (asm) power_eq_iff_eq_base[symmetric, where n=2]) (auto simp add: norm_complex_def) lemma abs_sqrt_wlog: "(\x. x \ 0 \ P x (x\<^sup>2)) \ P \x\ (x\<^sup>2)" for x::"'a::linordered_idom" by (metis abs_ge_zero power2_abs) lemma complex_abs_le_norm: "\Re z\ + \Im z\ \ sqrt 2 * norm z" unfolding norm_complex_def apply (rule abs_sqrt_wlog [where x="Re z"]) apply (rule abs_sqrt_wlog [where x="Im z"]) apply (rule power2_le_imp_le) apply (simp_all add: power2_sum add.commute sum_squares_bound real_sqrt_mult [symmetric]) done lemma complex_unit_circle: "z \ 0 \ (Re z / cmod z)\<^sup>2 + (Im z / cmod z)\<^sup>2 = 1" by (simp add: norm_complex_def complex_eq_iff power2_eq_square add_divide_distrib [symmetric]) text \Properties of complex signum.\ lemma sgn_eq: "sgn z = z / complex_of_real (cmod z)" by (simp add: sgn_div_norm divide_inverse scaleR_conv_of_real mult.commute) lemma Re_sgn [simp]: "Re(sgn z) = Re(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) lemma Im_sgn [simp]: "Im(sgn z) = Im(z)/cmod z" by (simp add: complex_sgn_def divide_inverse) subsection \Absolute value\ instantiation complex :: field_abs_sgn begin definition abs_complex :: "complex \ complex" where "abs_complex = of_real \ norm" instance apply standard apply (auto simp add: abs_complex_def complex_sgn_def norm_mult) apply (auto simp add: scaleR_conv_of_real field_simps) done end subsection \Completeness of the Complexes\ lemma bounded_linear_Re: "bounded_linear Re" by (rule bounded_linear_intro [where K=1]) (simp_all add: norm_complex_def) lemma bounded_linear_Im: "bounded_linear Im" by (rule bounded_linear_intro [where K=1]) (simp_all add: norm_complex_def) lemmas Cauchy_Re = bounded_linear.Cauchy [OF bounded_linear_Re] lemmas Cauchy_Im = bounded_linear.Cauchy [OF bounded_linear_Im] lemmas tendsto_Re [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Re] lemmas tendsto_Im [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_Im] lemmas isCont_Re [simp] = bounded_linear.isCont [OF bounded_linear_Re] lemmas isCont_Im [simp] = bounded_linear.isCont [OF bounded_linear_Im] lemmas continuous_Re [simp] = bounded_linear.continuous [OF bounded_linear_Re] lemmas continuous_Im [simp] = bounded_linear.continuous [OF bounded_linear_Im] lemmas continuous_on_Re [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Re] lemmas continuous_on_Im [continuous_intros] = bounded_linear.continuous_on[OF bounded_linear_Im] lemmas has_derivative_Re [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Re] lemmas has_derivative_Im [derivative_intros] = bounded_linear.has_derivative[OF bounded_linear_Im] lemmas sums_Re = bounded_linear.sums [OF bounded_linear_Re] lemmas sums_Im = bounded_linear.sums [OF bounded_linear_Im] lemma tendsto_Complex [tendsto_intros]: "(f \ a) F \ (g \ b) F \ ((\x. Complex (f x) (g x)) \ Complex a b) F" unfolding Complex_eq by (auto intro!: tendsto_intros) lemma tendsto_complex_iff: "(f \ x) F \ (((\x. Re (f x)) \ Re x) F \ ((\x. Im (f x)) \ Im x) F)" proof safe assume "((\x. Re (f x)) \ Re x) F" "((\x. Im (f x)) \ Im x) F" from tendsto_Complex[OF this] show "(f \ x) F" unfolding complex.collapse . qed (auto intro: tendsto_intros) lemma continuous_complex_iff: "continuous F f \ continuous F (\x. Re (f x)) \ continuous F (\x. Im (f x))" by (simp only: continuous_def tendsto_complex_iff) lemma continuous_on_of_real_o_iff [simp]: "continuous_on S (\x. complex_of_real (g x)) = continuous_on S g" using continuous_on_Re continuous_on_of_real by fastforce lemma continuous_on_of_real_id [simp]: "continuous_on S (of_real :: real \ 'a::real_normed_algebra_1)" by (rule continuous_on_of_real [OF continuous_on_id]) lemma has_vector_derivative_complex_iff: "(f has_vector_derivative x) F \ ((\x. Re (f x)) has_field_derivative (Re x)) F \ ((\x. Im (f x)) has_field_derivative (Im x)) F" by (simp add: has_vector_derivative_def has_field_derivative_def has_derivative_def tendsto_complex_iff algebra_simps bounded_linear_scaleR_left bounded_linear_mult_right) lemma has_field_derivative_Re[derivative_intros]: "(f has_vector_derivative D) F \ ((\x. Re (f x)) has_field_derivative (Re D)) F" unfolding has_vector_derivative_complex_iff by safe lemma has_field_derivative_Im[derivative_intros]: "(f has_vector_derivative D) F \ ((\x. Im (f x)) has_field_derivative (Im D)) F" unfolding has_vector_derivative_complex_iff by safe instance complex :: banach proof fix X :: "nat \ complex" assume X: "Cauchy X" then have "(\n. Complex (Re (X n)) (Im (X n))) \ Complex (lim (\n. Re (X n))) (lim (\n. Im (X n)))" by (intro tendsto_Complex convergent_LIMSEQ_iff[THEN iffD1] Cauchy_convergent_iff[THEN iffD1] Cauchy_Re Cauchy_Im) then show "convergent X" unfolding complex.collapse by (rule convergentI) qed declare DERIV_power[where 'a=complex, unfolded of_nat_def[symmetric], derivative_intros] subsection \Complex Conjugation\ primcorec cnj :: "complex \ complex" where "Re (cnj z) = Re z" | "Im (cnj z) = - Im z" lemma complex_cnj_cancel_iff [simp]: "cnj x = cnj y \ x = y" by (simp add: complex_eq_iff) lemma complex_cnj_cnj [simp]: "cnj (cnj z) = z" by (simp add: complex_eq_iff) lemma complex_cnj_zero [simp]: "cnj 0 = 0" by (simp add: complex_eq_iff) lemma complex_cnj_zero_iff [iff]: "cnj z = 0 \ z = 0" by (simp add: complex_eq_iff) lemma complex_cnj_one_iff [simp]: "cnj z = 1 \ z = 1" by (simp add: complex_eq_iff) lemma complex_cnj_add [simp]: "cnj (x + y) = cnj x + cnj y" by (simp add: complex_eq_iff) lemma cnj_sum [simp]: "cnj (sum f s) = (\x\s. cnj (f x))" by (induct s rule: infinite_finite_induct) auto lemma complex_cnj_diff [simp]: "cnj (x - y) = cnj x - cnj y" by (simp add: complex_eq_iff) lemma complex_cnj_minus [simp]: "cnj (- x) = - cnj x" by (simp add: complex_eq_iff) lemma complex_cnj_one [simp]: "cnj 1 = 1" by (simp add: complex_eq_iff) lemma complex_cnj_mult [simp]: "cnj (x * y) = cnj x * cnj y" by (simp add: complex_eq_iff) lemma cnj_prod [simp]: "cnj (prod f s) = (\x\s. cnj (f x))" by (induct s rule: infinite_finite_induct) auto lemma complex_cnj_inverse [simp]: "cnj (inverse x) = inverse (cnj x)" by (simp add: complex_eq_iff) lemma complex_cnj_divide [simp]: "cnj (x / y) = cnj x / cnj y" by (simp add: divide_complex_def) lemma complex_cnj_power [simp]: "cnj (x ^ n) = cnj x ^ n" by (induct n) simp_all lemma complex_cnj_of_nat [simp]: "cnj (of_nat n) = of_nat n" by (simp add: complex_eq_iff) lemma complex_cnj_of_int [simp]: "cnj (of_int z) = of_int z" by (simp add: complex_eq_iff) lemma complex_cnj_numeral [simp]: "cnj (numeral w) = numeral w" by (simp add: complex_eq_iff) lemma complex_cnj_neg_numeral [simp]: "cnj (- numeral w) = - numeral w" by (simp add: complex_eq_iff) lemma complex_cnj_scaleR [simp]: "cnj (scaleR r x) = scaleR r (cnj x)" by (simp add: complex_eq_iff) lemma complex_mod_cnj [simp]: "cmod (cnj z) = cmod z" by (simp add: norm_complex_def) lemma complex_cnj_complex_of_real [simp]: "cnj (of_real x) = of_real x" by (simp add: complex_eq_iff) lemma complex_cnj_i [simp]: "cnj \ = - \" by (simp add: complex_eq_iff) lemma complex_add_cnj: "z + cnj z = complex_of_real (2 * Re z)" by (simp add: complex_eq_iff) lemma complex_diff_cnj: "z - cnj z = complex_of_real (2 * Im z) * \" by (simp add: complex_eq_iff) lemma Ints_cnj [intro]: "x \ \ \ cnj x \ \" by (auto elim!: Ints_cases) lemma cnj_in_Ints_iff [simp]: "cnj x \ \ \ x \ \" using Ints_cnj[of x] Ints_cnj[of "cnj x"] by auto lemma complex_mult_cnj: "z * cnj z = complex_of_real ((Re z)\<^sup>2 + (Im z)\<^sup>2)" by (simp add: complex_eq_iff power2_eq_square) lemma cnj_add_mult_eq_Re: "z * cnj w + cnj z * w = 2 * Re (z * cnj w)" by (rule complex_eqI) auto lemma complex_mod_mult_cnj: "cmod (z * cnj z) = (cmod z)\<^sup>2" by (simp add: norm_mult power2_eq_square) lemma complex_mod_sqrt_Re_mult_cnj: "cmod z = sqrt (Re (z * cnj z))" by (simp add: norm_complex_def power2_eq_square) lemma complex_In_mult_cnj_zero [simp]: "Im (z * cnj z) = 0" by simp lemma complex_cnj_fact [simp]: "cnj (fact n) = fact n" by (subst of_nat_fact [symmetric], subst complex_cnj_of_nat) simp lemma complex_cnj_pochhammer [simp]: "cnj (pochhammer z n) = pochhammer (cnj z) n" by (induct n arbitrary: z) (simp_all add: pochhammer_rec) lemma bounded_linear_cnj: "bounded_linear cnj" using complex_cnj_add complex_cnj_scaleR by (rule bounded_linear_intro [where K=1]) simp lemma linear_cnj: "linear cnj" using bounded_linear.linear[OF bounded_linear_cnj] . lemmas tendsto_cnj [tendsto_intros] = bounded_linear.tendsto [OF bounded_linear_cnj] and isCont_cnj [simp] = bounded_linear.isCont [OF bounded_linear_cnj] and continuous_cnj [simp, continuous_intros] = bounded_linear.continuous [OF bounded_linear_cnj] and continuous_on_cnj [simp, continuous_intros] = bounded_linear.continuous_on [OF bounded_linear_cnj] and has_derivative_cnj [simp, derivative_intros] = bounded_linear.has_derivative [OF bounded_linear_cnj] lemma lim_cnj: "((\x. cnj(f x)) \ cnj l) F \ (f \ l) F" by (simp add: tendsto_iff dist_complex_def complex_cnj_diff [symmetric] del: complex_cnj_diff) lemma sums_cnj: "((\x. cnj(f x)) sums cnj l) \ (f sums l)" by (simp add: sums_def lim_cnj cnj_sum [symmetric] del: cnj_sum) lemma differentiable_cnj_iff: "(\z. cnj (f z)) differentiable at x within A \ f differentiable at x within A" proof assume "(\z. cnj (f z)) differentiable at x within A" then obtain D where "((\z. cnj (f z)) has_derivative D) (at x within A)" by (auto simp: differentiable_def) from has_derivative_cnj[OF this] show "f differentiable at x within A" by (auto simp: differentiable_def) next assume "f differentiable at x within A" then obtain D where "(f has_derivative D) (at x within A)" by (auto simp: differentiable_def) from has_derivative_cnj[OF this] show "(\z. cnj (f z)) differentiable at x within A" by (auto simp: differentiable_def) qed lemma has_vector_derivative_cnj [derivative_intros]: assumes "(f has_vector_derivative f') (at z within A)" shows "((\z. cnj (f z)) has_vector_derivative cnj f') (at z within A)" using assms by (auto simp: has_vector_derivative_complex_iff intro: derivative_intros) subsection \Basic Lemmas\ lemma complex_eq_0: "z=0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 = 0" by (metis zero_complex.sel complex_eqI sum_power2_eq_zero_iff) lemma complex_neq_0: "z\0 \ (Re z)\<^sup>2 + (Im z)\<^sup>2 > 0" by (metis complex_eq_0 less_numeral_extra(3) sum_power2_gt_zero_iff) lemma complex_norm_square: "of_real ((norm z)\<^sup>2) = z * cnj z" by (cases z) (auto simp: complex_eq_iff norm_complex_def power2_eq_square[symmetric] of_real_power[symmetric] simp del: of_real_power) lemma complex_div_cnj: "a / b = (a * cnj b) / (norm b)\<^sup>2" using complex_norm_square by auto lemma Re_complex_div_eq_0: "Re (a / b) = 0 \ Re (a * cnj b) = 0" by (auto simp add: Re_divide) lemma Im_complex_div_eq_0: "Im (a / b) = 0 \ Im (a * cnj b) = 0" by (auto simp add: Im_divide) lemma complex_div_gt_0: "(Re (a / b) > 0 \ Re (a * cnj b) > 0) \ (Im (a / b) > 0 \ Im (a * cnj b) > 0)" proof (cases "b = 0") case True then show ?thesis by auto next case False then have "0 < (Re b)\<^sup>2 + (Im b)\<^sup>2" by (simp add: complex_eq_iff sum_power2_gt_zero_iff) then show ?thesis by (simp add: Re_divide Im_divide zero_less_divide_iff) qed lemma Re_complex_div_gt_0: "Re (a / b) > 0 \ Re (a * cnj b) > 0" and Im_complex_div_gt_0: "Im (a / b) > 0 \ Im (a * cnj b) > 0" using complex_div_gt_0 by auto lemma Re_complex_div_ge_0: "Re (a / b) \ 0 \ Re (a * cnj b) \ 0" by (metis le_less Re_complex_div_eq_0 Re_complex_div_gt_0) lemma Im_complex_div_ge_0: "Im (a / b) \ 0 \ Im (a * cnj b) \ 0" by (metis Im_complex_div_eq_0 Im_complex_div_gt_0 le_less) lemma Re_complex_div_lt_0: "Re (a / b) < 0 \ Re (a * cnj b) < 0" by (metis less_asym neq_iff Re_complex_div_eq_0 Re_complex_div_gt_0) lemma Im_complex_div_lt_0: "Im (a / b) < 0 \ Im (a * cnj b) < 0" by (metis Im_complex_div_eq_0 Im_complex_div_gt_0 less_asym neq_iff) lemma Re_complex_div_le_0: "Re (a / b) \ 0 \ Re (a * cnj b) \ 0" by (metis not_le Re_complex_div_gt_0) lemma Im_complex_div_le_0: "Im (a / b) \ 0 \ Im (a * cnj b) \ 0" by (metis Im_complex_div_gt_0 not_le) lemma Re_divide_of_real [simp]: "Re (z / of_real r) = Re z / r" by (simp add: Re_divide power2_eq_square) lemma Im_divide_of_real [simp]: "Im (z / of_real r) = Im z / r" by (simp add: Im_divide power2_eq_square) lemma Re_divide_Reals [simp]: "r \ \ \ Re (z / r) = Re z / Re r" by (metis Re_divide_of_real of_real_Re) lemma Im_divide_Reals [simp]: "r \ \ \ Im (z / r) = Im z / Re r" by (metis Im_divide_of_real of_real_Re) lemma Re_sum[simp]: "Re (sum f s) = (\x\s. Re (f x))" by (induct s rule: infinite_finite_induct) auto lemma Im_sum[simp]: "Im (sum f s) = (\x\s. Im(f x))" by (induct s rule: infinite_finite_induct) auto lemma sums_complex_iff: "f sums x \ ((\x. Re (f x)) sums Re x) \ ((\x. Im (f x)) sums Im x)" unfolding sums_def tendsto_complex_iff Im_sum Re_sum .. lemma summable_complex_iff: "summable f \ summable (\x. Re (f x)) \ summable (\x. Im (f x))" unfolding summable_def sums_complex_iff[abs_def] by (metis complex.sel) lemma summable_complex_of_real [simp]: "summable (\n. complex_of_real (f n)) \ summable f" unfolding summable_complex_iff by simp lemma summable_Re: "summable f \ summable (\x. Re (f x))" unfolding summable_complex_iff by blast lemma summable_Im: "summable f \ summable (\x. Im (f x))" unfolding summable_complex_iff by blast lemma complex_is_Nat_iff: "z \ \ \ Im z = 0 \ (\i. Re z = of_nat i)" by (auto simp: Nats_def complex_eq_iff) lemma complex_is_Int_iff: "z \ \ \ Im z = 0 \ (\i. Re z = of_int i)" by (auto simp: Ints_def complex_eq_iff) lemma complex_is_Real_iff: "z \ \ \ Im z = 0" by (auto simp: Reals_def complex_eq_iff) lemma Reals_cnj_iff: "z \ \ \ cnj z = z" by (auto simp: complex_is_Real_iff complex_eq_iff) lemma in_Reals_norm: "z \ \ \ norm z = \Re z\" by (simp add: complex_is_Real_iff norm_complex_def) lemma Re_Reals_divide: "r \ \ \ Re (r / z) = Re r * Re z / (norm z)\<^sup>2" by (simp add: Re_divide complex_is_Real_iff cmod_power2) lemma Im_Reals_divide: "r \ \ \ Im (r / z) = -Re r * Im z / (norm z)\<^sup>2" by (simp add: Im_divide complex_is_Real_iff cmod_power2) lemma series_comparison_complex: fixes f:: "nat \ 'a::banach" assumes sg: "summable g" and "\n. g n \ \" "\n. Re (g n) \ 0" and fg: "\n. n \ N \ norm(f n) \ norm(g n)" shows "summable f" proof - have g: "\n. cmod (g n) = Re (g n)" using assms by (metis abs_of_nonneg in_Reals_norm) show ?thesis apply (rule summable_comparison_test' [where g = "\n. norm (g n)" and N=N]) using sg apply (auto simp: summable_def) apply (rule_tac x = "Re s" in exI) apply (auto simp: g sums_Re) apply (metis fg g) done qed subsection \Polar Form for Complex Numbers\ lemma complex_unimodular_polar: assumes "norm z = 1" obtains t where "0 \ t" "t < 2 * pi" "z = Complex (cos t) (sin t)" by (metis cmod_power2 one_power2 complex_surj sincos_total_2pi [of "Re z" "Im z"] assms) subsubsection \$\cos \theta + i \sin \theta$\ primcorec cis :: "real \ complex" where "Re (cis a) = cos a" | "Im (cis a) = sin a" lemma cis_zero [simp]: "cis 0 = 1" by (simp add: complex_eq_iff) lemma norm_cis [simp]: "norm (cis a) = 1" by (simp add: norm_complex_def) lemma sgn_cis [simp]: "sgn (cis a) = cis a" by (simp add: sgn_div_norm) lemma cis_2pi [simp]: "cis (2 * pi) = 1" by (simp add: cis.ctr complex_eq_iff) lemma cis_neq_zero [simp]: "cis a \ 0" by (metis norm_cis norm_zero zero_neq_one) lemma cis_cnj: "cnj (cis t) = cis (-t)" by (simp add: complex_eq_iff) lemma cis_mult: "cis a * cis b = cis (a + b)" by (simp add: complex_eq_iff cos_add sin_add) lemma DeMoivre: "(cis a) ^ n = cis (real n * a)" by (induct n) (simp_all add: algebra_simps cis_mult) lemma cis_inverse [simp]: "inverse (cis a) = cis (- a)" by (simp add: complex_eq_iff) lemma cis_divide: "cis a / cis b = cis (a - b)" by (simp add: divide_complex_def cis_mult) lemma cos_n_Re_cis_pow_n: "cos (real n * a) = Re (cis a ^ n)" by (auto simp add: DeMoivre) lemma sin_n_Im_cis_pow_n: "sin (real n * a) = Im (cis a ^ n)" by (auto simp add: DeMoivre) lemma cis_pi [simp]: "cis pi = -1" by (simp add: complex_eq_iff) lemma cis_pi_half[simp]: "cis (pi / 2) = \" by (simp add: cis.ctr complex_eq_iff) lemma cis_minus_pi_half[simp]: "cis (-(pi / 2)) = -\" by (simp add: cis.ctr complex_eq_iff) lemma cis_multiple_2pi[simp]: "n \ \ \ cis (2 * pi * n) = 1" by (auto elim!: Ints_cases simp: cis.ctr one_complex.ctr) subsubsection \$r(\cos \theta + i \sin \theta)$\ definition rcis :: "real \ real \ complex" where "rcis r a = complex_of_real r * cis a" lemma Re_rcis [simp]: "Re(rcis r a) = r * cos a" by (simp add: rcis_def) lemma Im_rcis [simp]: "Im(rcis r a) = r * sin a" by (simp add: rcis_def) lemma rcis_Ex: "\r a. z = rcis r a" by (simp add: complex_eq_iff polar_Ex) lemma complex_mod_rcis [simp]: "cmod (rcis r a) = \r\" by (simp add: rcis_def norm_mult) lemma cis_rcis_eq: "cis a = rcis 1 a" by (simp add: rcis_def) lemma rcis_mult: "rcis r1 a * rcis r2 b = rcis (r1 * r2) (a + b)" by (simp add: rcis_def cis_mult) lemma rcis_zero_mod [simp]: "rcis 0 a = 0" by (simp add: rcis_def) lemma rcis_zero_arg [simp]: "rcis r 0 = complex_of_real r" by (simp add: rcis_def) lemma rcis_eq_zero_iff [simp]: "rcis r a = 0 \ r = 0" by (simp add: rcis_def) lemma DeMoivre2: "(rcis r a) ^ n = rcis (r ^ n) (real n * a)" by (simp add: rcis_def power_mult_distrib DeMoivre) lemma rcis_inverse: "inverse(rcis r a) = rcis (1 / r) (- a)" by (simp add: divide_inverse rcis_def) lemma rcis_divide: "rcis r1 a / rcis r2 b = rcis (r1 / r2) (a - b)" by (simp add: rcis_def cis_divide [symmetric]) - subsubsection \Complex exponential\ lemma exp_Reals_eq: assumes "z \ \" shows "exp z = of_real (exp (Re z))" using assms by (auto elim!: Reals_cases simp: exp_of_real) lemma cis_conv_exp: "cis b = exp (\ * b)" proof - have "(\ * complex_of_real b) ^ n /\<^sub>R fact n = of_real (cos_coeff n * b^n) + \ * of_real (sin_coeff n * b^n)" for n :: nat proof - have "\ ^ n = fact n *\<^sub>R (cos_coeff n + \ * sin_coeff n)" by (induct n) (simp_all add: sin_coeff_Suc cos_coeff_Suc complex_eq_iff Re_divide Im_divide field_simps power2_eq_square add_nonneg_eq_0_iff) then show ?thesis by (simp add: field_simps) qed then show ?thesis using sin_converges [of b] cos_converges [of b] by (auto simp add: Complex_eq cis.ctr exp_def simp del: of_real_mult intro!: sums_unique sums_add sums_mult sums_of_real) qed lemma exp_eq_polar: "exp z = exp (Re z) * cis (Im z)" unfolding cis_conv_exp exp_of_real [symmetric] mult_exp_exp by (cases z) (simp add: Complex_eq) lemma Re_exp: "Re (exp z) = exp (Re z) * cos (Im z)" unfolding exp_eq_polar by simp lemma Im_exp: "Im (exp z) = exp (Re z) * sin (Im z)" unfolding exp_eq_polar by simp lemma norm_cos_sin [simp]: "norm (Complex (cos t) (sin t)) = 1" by (simp add: norm_complex_def) lemma norm_exp_eq_Re [simp]: "norm (exp z) = exp (Re z)" by (simp add: cis.code cmod_complex_polar exp_eq_polar Complex_eq) lemma complex_exp_exists: "\a r. z = complex_of_real r * exp a" apply (insert rcis_Ex [of z]) apply (auto simp add: exp_eq_polar rcis_def mult.assoc [symmetric]) apply (rule_tac x = "\ * complex_of_real a" in exI) apply auto done lemma exp_pi_i [simp]: "exp (of_real pi * \) = -1" by (metis cis_conv_exp cis_pi mult.commute) lemma exp_pi_i' [simp]: "exp (\ * of_real pi) = -1" using cis_conv_exp cis_pi by auto lemma exp_two_pi_i [simp]: "exp (2 * of_real pi * \) = 1" by (simp add: exp_eq_polar complex_eq_iff) lemma exp_two_pi_i' [simp]: "exp (\ * (of_real pi * 2)) = 1" by (metis exp_two_pi_i mult.commute) lemma continuous_on_cis [continuous_intros]: "continuous_on A f \ continuous_on A (\x. cis (f x))" by (auto simp: cis_conv_exp intro!: continuous_intros) subsubsection \Complex argument\ -definition arg :: "complex \ real" - where "arg z = (if z = 0 then 0 else (SOME a. sgn z = cis a \ - pi < a \ a \ pi))" +definition Arg :: "complex \ real" + where "Arg z = (if z = 0 then 0 else (SOME a. sgn z = cis a \ - pi < a \ a \ pi))" -lemma arg_zero: "arg 0 = 0" - by (simp add: arg_def) +lemma Arg_zero: "Arg 0 = 0" + by (simp add: Arg_def) lemma arg_unique: assumes "sgn z = cis x" and "-pi < x" and "x \ pi" - shows "arg z = x" + shows "Arg z = x" proof - from assms have "z \ 0" by auto have "(SOME a. sgn z = cis a \ -pi < a \ a \ pi) = x" proof fix a define d where "d = a - x" assume a: "sgn z = cis a \ - pi < a \ a \ pi" from a assms have "- (2*pi) < d \ d < 2*pi" unfolding d_def by simp moreover from a assms have "cos a = cos x" and "sin a = sin x" by (simp_all add: complex_eq_iff) then have cos: "cos d = 1" by (simp add: d_def cos_diff) moreover from cos have "sin d = 0" by (rule cos_one_sin_zero) ultimately have "d = 0" by (auto simp: sin_zero_iff elim!: evenE dest!: less_2_cases) then show "a = x" by (simp add: d_def) qed (simp add: assms del: Re_sgn Im_sgn) - with \z \ 0\ show "arg z = x" - by (simp add: arg_def) + with \z \ 0\ show "Arg z = x" + by (simp add: Arg_def) qed lemma arg_correct: assumes "z \ 0" - shows "sgn z = cis (arg z) \ -pi < arg z \ arg z \ pi" -proof (simp add: arg_def assms, rule someI_ex) + shows "sgn z = cis (Arg z) \ -pi < Arg z \ Arg z \ pi" +proof (simp add: Arg_def assms, rule someI_ex) obtain r a where z: "z = rcis r a" using rcis_Ex by fast with assms have "r \ 0" by auto define b where "b = (if 0 < r then a else a + pi)" have b: "sgn z = cis b" using \r \ 0\ by (simp add: z b_def rcis_def of_real_def sgn_scaleR sgn_if complex_eq_iff) have cis_2pi_nat: "cis (2 * pi * real_of_nat n) = 1" for n by (induct n) (simp_all add: distrib_left cis_mult [symmetric] complex_eq_iff) have cis_2pi_int: "cis (2 * pi * real_of_int x) = 1" for x by (cases x rule: int_diff_cases) (simp add: right_diff_distrib cis_divide [symmetric] cis_2pi_nat) define c where "c = b - 2 * pi * of_int \(b - pi) / (2 * pi)\" have "sgn z = cis c" by (simp add: b c_def cis_divide [symmetric] cis_2pi_int) moreover have "- pi < c \ c \ pi" using ceiling_correct [of "(b - pi) / (2*pi)"] by (simp add: c_def less_divide_eq divide_le_eq algebra_simps del: le_of_int_ceiling) ultimately show "\a. sgn z = cis a \ -pi < a \ a \ pi" by fast qed -lemma arg_bounded: "- pi < arg z \ arg z \ pi" - by (cases "z = 0") (simp_all add: arg_zero arg_correct) +lemma Arg_bounded: "- pi < Arg z \ Arg z \ pi" + by (cases "z = 0") (simp_all add: Arg_zero arg_correct) -lemma cis_arg: "z \ 0 \ cis (arg z) = sgn z" +lemma cis_Arg: "z \ 0 \ cis (Arg z) = sgn z" by (simp add: arg_correct) -lemma rcis_cmod_arg: "rcis (cmod z) (arg z) = z" - by (cases "z = 0") (simp_all add: rcis_def cis_arg sgn_div_norm of_real_def) +lemma rcis_cmod_Arg: "rcis (cmod z) (Arg z) = z" + by (cases "z = 0") (simp_all add: rcis_def cis_Arg sgn_div_norm of_real_def) -lemma cos_arg_i_mult_zero [simp]: "y \ 0 \ Re y = 0 \ cos (arg y) = 0" - using cis_arg [of y] by (simp add: complex_eq_iff) +lemma rcis_cnj: + shows "cnj a = rcis (cmod a) (- Arg a)" + by (metis cis_cnj complex_cnj_complex_of_real complex_cnj_mult rcis_cmod_Arg rcis_def) -lemma arg_ii [simp]: "arg \ = pi/2" +lemma cos_Arg_i_mult_zero [simp]: "y \ 0 \ Re y = 0 \ cos (Arg y) = 0" + using cis_Arg [of y] by (simp add: complex_eq_iff) + +lemma Arg_ii [simp]: "Arg \ = pi/2" by (rule arg_unique; simp add: sgn_eq) -lemma arg_minus_ii [simp]: "arg (-\) = -pi/2" +lemma Arg_minus_ii [simp]: "Arg (-\) = -pi/2" proof (rule arg_unique) show "sgn (- \) = cis (- pi / 2)" by (simp add: sgn_eq) show "- pi / 2 \ pi" using pi_not_less_zero by linarith qed auto subsection \Complex n-th roots\ lemma bij_betw_roots_unity: assumes "n > 0" shows "bij_betw (\k. cis (2 * pi * real k / real n)) {.. = cis (2*pi*(real k - real l)/n)" using assms by (simp add: field_simps cis_divide) finally have "cos (2*pi*(real k - real l) / n) = 1" by (simp add: complex_eq_iff) then obtain m :: int where "2 * pi * (real k - real l) / real n = real_of_int m * 2 * pi" by (subst (asm) cos_one_2pi_int) blast hence "real_of_int (int k - int l) = real_of_int (m * int n)" unfolding of_int_diff of_int_mult using assms by (simp add: nonzero_divide_eq_eq) also note of_int_eq_iff finally have *: "abs m * n = abs (int k - int l)" by (simp add: abs_mult) also have "\ < int n" using kl by linarith finally have "m = 0" using assms by simp with * show "k = l" by simp qed have subset: "?f ` {.. {z. z ^ n = 1}" proof safe fix k :: nat have "cis (2 * pi * real k / real n) ^ n = cis (2 * pi) ^ k" using assms by (simp add: DeMoivre mult_ac) also have "cis (2 * pi) = 1" by (simp add: complex_eq_iff) finally show "?f k ^ n = 1" by simp qed have "n = card {.. \ card {z::complex. z ^ n = 1}" by (intro card_inj_on_le[OF inj]) (auto simp: finite_roots_unity) finally have card: "card {z::complex. z ^ n = 1} = n" using assms by (intro antisym card_roots_unity) auto have "card (?f ` {.. 0" shows "card {z::complex. z ^ n = 1} = n" using bij_betw_same_card [OF bij_betw_roots_unity [OF assms]] by simp lemma bij_betw_nth_root_unity: fixes c :: complex and n :: nat assumes c: "c \ 0" and n: "n > 0" - defines "c' \ root n (norm c) * cis (arg c / n)" + defines "c' \ root n (norm c) * cis (Arg c / n)" shows "bij_betw (\z. c' * z) {z. z ^ n = 1} {z. z ^ n = c}" proof - - have "c' ^ n = of_real (root n (norm c) ^ n) * cis (arg c)" + have "c' ^ n = of_real (root n (norm c) ^ n) * cis (Arg c)" unfolding of_real_power using n by (simp add: c'_def power_mult_distrib DeMoivre) also from n have "root n (norm c) ^ n = norm c" by simp - also from c have "of_real \ * cis (arg c) = c" by (simp add: cis_arg Complex.sgn_eq) + also from c have "of_real \ * cis (Arg c) = c" by (simp add: cis_Arg Complex.sgn_eq) finally have [simp]: "c' ^ n = c" . show ?thesis unfolding bij_betw_def inj_on_def proof safe fix z :: complex assume "z ^ n = 1" hence "(c' * z) ^ n = c' ^ n" by (simp add: power_mult_distrib) - also have "c' ^ n = of_real (root n (norm c) ^ n) * cis (arg c)" + also have "c' ^ n = of_real (root n (norm c) ^ n) * cis (Arg c)" unfolding of_real_power using n by (simp add: c'_def power_mult_distrib DeMoivre) also from n have "root n (norm c) ^ n = norm c" by simp - also from c have "\ * cis (arg c) = c" by (simp add: cis_arg Complex.sgn_eq) + also from c have "\ * cis (Arg c) = c" by (simp add: cis_Arg Complex.sgn_eq) finally show "(c' * z) ^ n = c" . next fix z assume z: "c = z ^ n" define z' where "z' = z / c'" from c and n have "c' \ 0" by (auto simp: c'_def) with n c have "z = c' * z'" and "z' ^ n = 1" by (auto simp: z'_def power_divide z) thus "z \ (\z. c' * z) ` {z. z ^ n = 1}" by blast qed (insert c n, auto simp: c'_def) qed lemma finite_nth_roots [intro]: assumes "n > 0" shows "finite {z::complex. z ^ n = c}" proof (cases "c = 0") case True with assms have "{z::complex. z ^ n = c} = {0}" by auto thus ?thesis by simp next case False from assms have "finite {z::complex. z ^ n = 1}" by (intro finite_roots_unity) simp_all also have "?this \ ?thesis" by (rule bij_betw_finite, rule bij_betw_nth_root_unity) fact+ finally show ?thesis . qed lemma card_nth_roots: assumes "c \ 0" "n > 0" shows "card {z::complex. z ^ n = c} = n" proof - have "card {z. z ^ n = c} = card {z::complex. z ^ n = 1}" by (rule sym, rule bij_betw_same_card, rule bij_betw_nth_root_unity) fact+ also have "\ = n" by (rule card_roots_unity_eq) fact+ finally show ?thesis . qed lemma sum_roots_unity: assumes "n > 1" shows "\{z::complex. z ^ n = 1} = 0" proof - define \ where "\ = cis (2 * pi / real n)" have [simp]: "\ \ 1" proof assume "\ = 1" with assms obtain k :: int where "2 * pi / real n = 2 * pi * of_int k" by (auto simp: \_def complex_eq_iff cos_one_2pi_int) with assms have "real n * of_int k = 1" by (simp add: field_simps) also have "real n * of_int k = of_int (int n * k)" by simp also have "1 = (of_int 1 :: real)" by simp also note of_int_eq_iff finally show False using assms by (auto simp: zmult_eq_1_iff) qed have "(\z | z ^ n = 1. z :: complex) = (\k = (\k ^ k)" by (intro sum.cong refl) (auto simp: \_def DeMoivre mult_ac) also have "\ = (\ ^ n - 1) / (\ - 1)" by (subst geometric_sum) auto also have "\ ^ n - 1 = cis (2 * pi) - 1" using assms by (auto simp: \_def DeMoivre) also have "\ = 0" by (simp add: complex_eq_iff) finally show ?thesis by simp qed lemma sum_nth_roots: assumes "n > 1" shows "\{z::complex. z ^ n = c} = 0" proof (cases "c = 0") case True with assms have "{z::complex. z ^ n = c} = {0}" by auto also have "\\ = 0" by simp finally show ?thesis . next case False - define c' where "c' = root n (norm c) * cis (arg c / n)" + define c' where "c' = root n (norm c) * cis (Arg c / n)" from False and assms have "(\{z. z ^ n = c}) = (\z | z ^ n = 1. c' * z)" by (subst sum.reindex_bij_betw [OF bij_betw_nth_root_unity, symmetric]) (auto simp: sum_distrib_left finite_roots_unity c'_def) also from assms have "\ = 0" by (simp add: sum_distrib_left [symmetric] sum_roots_unity) finally show ?thesis . qed subsection \Square root of complex numbers\ primcorec csqrt :: "complex \ complex" where "Re (csqrt z) = sqrt ((cmod z + Re z) / 2)" | "Im (csqrt z) = (if Im z = 0 then 1 else sgn (Im z)) * sqrt ((cmod z - Re z) / 2)" lemma csqrt_of_real_nonneg [simp]: "Im x = 0 \ Re x \ 0 \ csqrt x = sqrt (Re x)" by (simp add: complex_eq_iff norm_complex_def) lemma csqrt_of_real_nonpos [simp]: "Im x = 0 \ Re x \ 0 \ csqrt x = \ * sqrt \Re x\" by (simp add: complex_eq_iff norm_complex_def) lemma of_real_sqrt: "x \ 0 \ of_real (sqrt x) = csqrt (of_real x)" by (simp add: complex_eq_iff norm_complex_def) lemma csqrt_0 [simp]: "csqrt 0 = 0" by simp lemma csqrt_1 [simp]: "csqrt 1 = 1" by simp lemma csqrt_ii [simp]: "csqrt \ = (1 + \) / sqrt 2" by (simp add: complex_eq_iff Re_divide Im_divide real_sqrt_divide real_div_sqrt) lemma power2_csqrt[simp,algebra]: "(csqrt z)\<^sup>2 = z" proof (cases "Im z = 0") case True then show ?thesis using real_sqrt_pow2[of "Re z"] real_sqrt_pow2[of "- Re z"] by (cases "0::real" "Re z" rule: linorder_cases) (simp_all add: complex_eq_iff Re_power2 Im_power2 power2_eq_square cmod_eq_Re) next case False moreover have "cmod z * cmod z - Re z * Re z = Im z * Im z" by (simp add: norm_complex_def power2_eq_square) moreover have "\Re z\ \ cmod z" by (simp add: norm_complex_def) ultimately show ?thesis by (simp add: Re_power2 Im_power2 complex_eq_iff real_sgn_eq field_simps real_sqrt_mult[symmetric] real_sqrt_divide) qed lemma csqrt_eq_0 [simp]: "csqrt z = 0 \ z = 0" by auto (metis power2_csqrt power_eq_0_iff) lemma csqrt_eq_1 [simp]: "csqrt z = 1 \ z = 1" by auto (metis power2_csqrt power2_eq_1_iff) lemma csqrt_principal: "0 < Re (csqrt z) \ Re (csqrt z) = 0 \ 0 \ Im (csqrt z)" by (auto simp add: not_less cmod_plus_Re_le_0_iff Im_eq_0) lemma Re_csqrt: "0 \ Re (csqrt z)" by (metis csqrt_principal le_less) lemma csqrt_square: assumes "0 < Re b \ (Re b = 0 \ 0 \ Im b)" shows "csqrt (b^2) = b" proof - have "csqrt (b^2) = b \ csqrt (b^2) = - b" by (simp add: power2_eq_iff[symmetric]) moreover have "csqrt (b^2) \ -b \ b = 0" using csqrt_principal[of "b ^ 2"] assms by (intro disjCI notI) (auto simp: complex_eq_iff) ultimately show ?thesis by auto qed lemma csqrt_unique: "w\<^sup>2 = z \ 0 < Re w \ Re w = 0 \ 0 \ Im w \ csqrt z = w" by (auto simp: csqrt_square) lemma csqrt_minus [simp]: assumes "Im x < 0 \ (Im x = 0 \ 0 \ Re x)" shows "csqrt (- x) = \ * csqrt x" proof - have "csqrt ((\ * csqrt x)^2) = \ * csqrt x" proof (rule csqrt_square) have "Im (csqrt x) \ 0" using assms by (auto simp add: cmod_eq_Re mult_le_0_iff field_simps complex_Re_le_cmod) then show "0 < Re (\ * csqrt x) \ Re (\ * csqrt x) = 0 \ 0 \ Im (\ * csqrt x)" by (auto simp add: Re_csqrt simp del: csqrt.simps) qed also have "(\ * csqrt x)^2 = - x" by (simp add: power_mult_distrib) finally show ?thesis . qed text \Legacy theorem names\ lemmas cmod_def = norm_complex_def lemma legacy_Complex_simps: shows Complex_eq_0: "Complex a b = 0 \ a = 0 \ b = 0" and complex_add: "Complex a b + Complex c d = Complex (a + c) (b + d)" and complex_minus: "- (Complex a b) = Complex (- a) (- b)" and complex_diff: "Complex a b - Complex c d = Complex (a - c) (b - d)" and Complex_eq_1: "Complex a b = 1 \ a = 1 \ b = 0" and Complex_eq_neg_1: "Complex a b = - 1 \ a = - 1 \ b = 0" and complex_mult: "Complex a b * Complex c d = Complex (a * c - b * d) (a * d + b * c)" and complex_inverse: "inverse (Complex a b) = Complex (a / (a\<^sup>2 + b\<^sup>2)) (- b / (a\<^sup>2 + b\<^sup>2))" and Complex_eq_numeral: "Complex a b = numeral w \ a = numeral w \ b = 0" and Complex_eq_neg_numeral: "Complex a b = - numeral w \ a = - numeral w \ b = 0" and complex_scaleR: "scaleR r (Complex a b) = Complex (r * a) (r * b)" and Complex_eq_i: "Complex x y = \ \ x = 0 \ y = 1" and i_mult_Complex: "\ * Complex a b = Complex (- b) a" and Complex_mult_i: "Complex a b * \ = Complex (- b) a" and i_complex_of_real: "\ * complex_of_real r = Complex 0 r" and complex_of_real_i: "complex_of_real r * \ = Complex 0 r" and Complex_add_complex_of_real: "Complex x y + complex_of_real r = Complex (x+r) y" and complex_of_real_add_Complex: "complex_of_real r + Complex x y = Complex (r+x) y" and Complex_mult_complex_of_real: "Complex x y * complex_of_real r = Complex (x*r) (y*r)" and complex_of_real_mult_Complex: "complex_of_real r * Complex x y = Complex (r*x) (r*y)" and complex_eq_cancel_iff2: "(Complex x y = complex_of_real xa) = (x = xa \ y = 0)" and complex_cnj: "cnj (Complex a b) = Complex a (- b)" and Complex_sum': "sum (\x. Complex (f x) 0) s = Complex (sum f s) 0" and Complex_sum: "Complex (sum f s) 0 = sum (\x. Complex (f x) 0) s" and complex_of_real_def: "complex_of_real r = Complex r 0" and complex_norm: "cmod (Complex x y) = sqrt (x\<^sup>2 + y\<^sup>2)" by (simp_all add: norm_complex_def field_simps complex_eq_iff Re_divide Im_divide) lemma Complex_in_Reals: "Complex x 0 \ \" by (metis Reals_of_real complex_of_real_def) end diff --git a/src/HOL/Complex_Analysis/Complex_Residues.thy b/src/HOL/Complex_Analysis/Complex_Residues.thy --- a/src/HOL/Complex_Analysis/Complex_Residues.thy +++ b/src/HOL/Complex_Analysis/Complex_Residues.thy @@ -1,545 +1,540 @@ theory Complex_Residues imports Complex_Singularities begin subsection \Definition of residues\ text\Wenda Li and LC Paulson (2016). A Formal Proof of Cauchy's Residue Theorem. Interactive Theorem Proving\ definition\<^marker>\tag important\ residue :: "(complex \ complex) \ complex \ complex" where "residue f z = (SOME int. \e>0. \\>0. \ (f has_contour_integral 2*pi* \ *int) (circlepath z \))" -lemma Eps_cong: - assumes "\x. P x = Q x" - shows "Eps P = Eps Q" - using ext[of P Q, OF assms] by simp - lemma residue_cong: assumes eq: "eventually (\z. f z = g z) (at z)" and "z = z'" shows "residue f z = residue g z'" proof - from assms have eq': "eventually (\z. g z = f z) (at z)" by (simp add: eq_commute) let ?P = "\f c e. (\\>0. \ < e \ (f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \))" have "residue f z = residue g z" unfolding residue_def proof (rule Eps_cong) fix c :: complex have "\e>0. ?P g c e" if "\e>0. ?P f c e" and "eventually (\z. f z = g z) (at z)" for f g proof - from that(1) obtain e where e: "e > 0" "?P f c e" by blast from that(2) obtain e' where e': "e' > 0" "\z'. z' \ z \ dist z' z < e' \ f z' = g z'" unfolding eventually_at by blast have "?P g c (min e e')" proof (intro allI exI impI, goal_cases) case (1 \) hence "(f has_contour_integral of_real (2 * pi) * \ * c) (circlepath z \)" using e(2) by auto thus ?case proof (rule has_contour_integral_eq) fix z' assume "z' \ path_image (circlepath z \)" hence "dist z' z < e'" and "z' \ z" using 1 by (auto simp: dist_commute) with e'(2)[of z'] show "f z' = g z'" by simp qed qed moreover from e and e' have "min e e' > 0" by auto ultimately show ?thesis by blast qed from this[OF _ eq] and this[OF _ eq'] show "(\e>0. ?P f c e) \ (\e>0. ?P g c e)" by blast qed with assms show ?thesis by simp qed lemma contour_integral_circlepath_eq: assumes "open s" and f_holo:"f holomorphic_on (s-{z})" and "0e2" and e2_cball:"cball z e2 \ s" shows "f contour_integrable_on circlepath z e1" "f contour_integrable_on circlepath z e2" "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" proof - define l where "l \ linepath (z+e2) (z+e1)" have [simp]:"valid_path l" "pathstart l=z+e2" "pathfinish l=z+e1" unfolding l_def by auto have "e2>0" using \e1>0\ \e1\e2\ by auto have zl_img:"z\path_image l" proof assume "z \ path_image l" then have "e2 \ cmod (e2 - e1)" using segment_furthest_le[of z "z+e2" "z+e1" "z+e2",simplified] \e1>0\ \e2>0\ unfolding l_def by (auto simp add:closed_segment_commute) thus False using \e2>0\ \e1>0\ \e1\e2\ apply (subst (asm) norm_of_real) by auto qed define g where "g \ circlepath z e2 +++ l +++ reversepath (circlepath z e1) +++ reversepath l" show [simp]: "f contour_integrable_on circlepath z e2" "f contour_integrable_on (circlepath z e1)" proof - show "f contour_integrable_on circlepath z e2" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e2>0\ e2_cball by auto show "f contour_integrable_on (circlepath z e1)" apply (intro contour_integrable_continuous_circlepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) using \e1>0\ \e1\e2\ e2_cball by auto qed have [simp]:"f contour_integrable_on l" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) hence "closed_segment (z + e2) (z + e1) \ s - {z}" using zl_img e2_cball unfolding l_def by auto then show "f contour_integrable_on l" unfolding l_def apply (intro contour_integrable_continuous_linepath[OF continuous_on_subset[OF holomorphic_on_imp_continuous_on[OF f_holo]]]) by auto qed let ?ig="\g. contour_integral g f" have "(f has_contour_integral 0) g" proof (rule Cauchy_theorem_global[OF _ f_holo]) show "open (s - {z})" using \open s\ by auto show "valid_path g" unfolding g_def l_def by auto show "pathfinish g = pathstart g" unfolding g_def l_def by auto next have path_img:"path_image g \ cball z e2" proof - have "closed_segment (z + e2) (z + e1) \ cball z e2" using \e2>0\ \e1>0\ \e1\e2\ by (intro closed_segment_subset,auto simp add:dist_norm) moreover have "sphere z \e1\ \ cball z e2" using \e2>0\ \e1\e2\ \e1>0\ by auto ultimately show ?thesis unfolding g_def l_def using \e2>0\ by (simp add: path_image_join closed_segment_commute) qed show "path_image g \ s - {z}" proof - have "z\path_image g" using zl_img unfolding g_def l_def by (auto simp add: path_image_join closed_segment_commute) moreover note \cball z e2 \ s\ and path_img ultimately show ?thesis by auto qed show "winding_number g w = 0" when"w \ s - {z}" for w proof - have "winding_number g w = 0" when "w\s" using that e2_cball apply (intro winding_number_zero_outside[OF _ _ _ _ path_img]) by (auto simp add:g_def l_def) moreover have "winding_number g z=0" proof - let ?Wz="\g. winding_number g z" have "?Wz g = ?Wz (circlepath z e2) + ?Wz l + ?Wz (reversepath (circlepath z e1)) + ?Wz (reversepath l)" using \e2>0\ \e1>0\ zl_img unfolding g_def l_def by (subst winding_number_join,auto simp add:path_image_join closed_segment_commute)+ also have "... = ?Wz (circlepath z e2) + ?Wz (reversepath (circlepath z e1))" using zl_img apply (subst (2) winding_number_reversepath) by (auto simp add:l_def closed_segment_commute) also have "... = 0" proof - have "?Wz (circlepath z e2) = 1" using \e2>0\ by (auto intro: winding_number_circlepath_centre) moreover have "?Wz (reversepath (circlepath z e1)) = -1" using \e1>0\ apply (subst winding_number_reversepath) by (auto intro: winding_number_circlepath_centre) ultimately show ?thesis by auto qed finally show ?thesis . qed ultimately show ?thesis using that by auto qed qed then have "0 = ?ig g" using contour_integral_unique by simp also have "... = ?ig (circlepath z e2) + ?ig l + ?ig (reversepath (circlepath z e1)) + ?ig (reversepath l)" unfolding g_def by (auto simp add:contour_integrable_reversepath_eq) also have "... = ?ig (circlepath z e2) - ?ig (circlepath z e1)" by (auto simp add:contour_integral_reversepath) finally show "contour_integral (circlepath z e2) f = contour_integral (circlepath z e1) f" by simp qed lemma base_residue: assumes "open s" "z\s" "r>0" and f_holo:"f holomorphic_on (s - {z})" and r_cball:"cball z r \ s" shows "(f has_contour_integral 2 * pi * \ * (residue f z)) (circlepath z r)" proof - obtain e where "e>0" and e_cball:"cball z e \ s" using open_contains_cball[of s] \open s\ \z\s\ by auto define c where "c \ 2 * pi * \" define i where "i \ contour_integral (circlepath z e) f / c" have "(f has_contour_integral c*i) (circlepath z \)" when "\>0" "\ proof - have "contour_integral (circlepath z e) f = contour_integral (circlepath z \) f" "f contour_integrable_on circlepath z \" "f contour_integrable_on circlepath z e" using \\ by (intro contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ _ e_cball],auto)+ then show ?thesis unfolding i_def c_def by (auto intro:has_contour_integral_integral) qed then have "\e>0. \\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" unfolding residue_def c_def apply (rule_tac someI[of _ i],intro exI[where x=e]) by (auto simp add:\e>0\ c_def) then obtain e' where "e'>0" and e'_def:"\\>0. \ (f has_contour_integral c * (residue f z)) (circlepath z \)" by auto let ?int="\e. contour_integral (circlepath z e) f" define \ where "\ \ Min {r,e'} / 2" have "\>0" "\\r" "\r>0\ \e'>0\ unfolding \_def by auto have "(f has_contour_integral c * (residue f z)) (circlepath z \)" using e'_def[rule_format,OF \\>0\ \\] . then show ?thesis unfolding c_def using contour_integral_circlepath_eq[OF \open s\ f_holo \\>0\ \\\r\ r_cball] by (auto elim: has_contour_integral_eqpath[of _ _ "circlepath z \" "circlepath z r"]) qed lemma residue_holo: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s" shows "residue f z = 0" proof - define c where "c \ 2 * pi * \" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f has_contour_integral c*residue f z) (circlepath z e)" using f_holo by (auto intro: base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) moreover have "(f has_contour_integral 0) (circlepath z e)" using f_holo e_cball \e>0\ by (auto intro: Cauchy_theorem_convex_simple[of _ "cball z e"]) ultimately have "c*residue f z =0" using has_contour_integral_unique by blast thus ?thesis unfolding c_def by auto qed lemma residue_const:"residue (\_. c) z = 0" by (intro residue_holo[of "UNIV::complex set"],auto intro:holomorphic_intros) lemma residue_add: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z + g z) z= residue f z + residue g z" proof - define c where "c \ 2 * pi * \" define fg where "fg \ (\z. f z+g z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(fg has_contour_integral c * residue fg z) (circlepath z e)" unfolding fg_def using f_holo g_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro:holomorphic_intros) moreover have "(fg has_contour_integral c*residue f z + c* residue g z) (circlepath z e)" unfolding fg_def using f_holo g_holo by (auto intro: has_contour_integral_add base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) ultimately have "c*(residue f z + residue g z) = c * residue fg z" using has_contour_integral_unique by (auto simp add:distrib_left) thus ?thesis unfolding fg_def by (auto simp add:c_def) qed lemma residue_lmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. c * (f z)) z= c * residue f z" proof (cases "c=0") case True thus ?thesis using residue_const by auto next case False define c' where "c' \ 2 * pi * \" define f' where "f' \ (\z. c * (f z))" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c' * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) by (auto intro:holomorphic_intros) moreover have "(f' has_contour_integral c * (c' * residue f z)) (circlepath z e)" unfolding f'_def using f_holo by (auto intro: has_contour_integral_lmul base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c'_def]) ultimately have "c' * residue f' z = c * (c' * residue f z)" using has_contour_integral_unique by auto thus ?thesis unfolding f'_def c'_def using False by (auto simp add:field_simps) qed lemma residue_rmul: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) * c) z= residue f z * c" using residue_lmul[OF assms,of c] by (auto simp add:algebra_simps) lemma residue_div: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. (f z) / c) z= residue f z / c " using residue_lmul[OF assms,of "1/c"] by (auto simp add:algebra_simps) lemma residue_neg: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" shows "residue (\z. - (f z)) z= - residue f z" using residue_lmul[OF assms,of "-1"] by auto lemma residue_diff: assumes "open s" "z \ s" and f_holo: "f holomorphic_on s - {z}" and g_holo:"g holomorphic_on s - {z}" shows "residue (\z. f z - g z) z= residue f z - residue g z" using residue_add[OF assms(1,2,3),of "\z. - g z"] residue_neg[OF assms(1,2,4)] by (auto intro:holomorphic_intros g_holo) lemma residue_simple: assumes "open s" "z\s" and f_holo:"f holomorphic_on s" shows "residue (\w. f w / (w - z)) z = f z" proof - define c where "c \ 2 * pi * \" define f' where "f' \ \w. f w / (w - z)" obtain e where "e>0" and e_cball:"cball z e \ s" using \open s\ \z\s\ using open_contains_cball_eq by blast have "(f' has_contour_integral c * f z) (circlepath z e)" unfolding f'_def c_def using \e>0\ f_holo e_cball by (auto intro!: Cauchy_integral_circlepath_simple holomorphic_intros) moreover have "(f' has_contour_integral c * residue f' z) (circlepath z e)" unfolding f'_def using f_holo apply (intro base_residue[OF \open s\ \z\s\ \e>0\ _ e_cball,folded c_def]) by (auto intro!:holomorphic_intros) ultimately have "c * f z = c * residue f' z" using has_contour_integral_unique by blast thus ?thesis unfolding c_def f'_def by auto qed lemma residue_simple': assumes s: "open s" "z \ s" and holo: "f holomorphic_on (s - {z})" and lim: "((\w. f w * (w - z)) \ c) (at z)" shows "residue f z = c" proof - define g where "g = (\w. if w = z then c else f w * (w - z))" from holo have "(\w. f w * (w - z)) holomorphic_on (s - {z})" (is "?P") by (force intro: holomorphic_intros) also have "?P \ g holomorphic_on (s - {z})" by (intro holomorphic_cong refl) (simp_all add: g_def) finally have *: "g holomorphic_on (s - {z})" . note lim also have "(\w. f w * (w - z)) \z\ c \ g \z\ g z" by (intro filterlim_cong refl) (simp_all add: g_def [abs_def] eventually_at_filter) finally have **: "g \z\ g z" . have g_holo: "g holomorphic_on s" by (rule no_isolated_singularity'[where K = "{z}"]) (insert assms * **, simp_all add: at_within_open_NO_MATCH) from s and this have "residue (\w. g w / (w - z)) z = g z" by (rule residue_simple) also have "\\<^sub>F za in at z. g za / (za - z) = f za" unfolding eventually_at by (auto intro!: exI[of _ 1] simp: field_simps g_def) hence "residue (\w. g w / (w - z)) z = residue f z" by (intro residue_cong refl) finally show ?thesis by (simp add: g_def) qed lemma residue_holomorphic_over_power: assumes "open A" "z0 \ A" "f holomorphic_on A" shows "residue (\z. f z / (z - z0) ^ Suc n) z0 = (deriv ^^ n) f z0 / fact n" proof - let ?f = "\z. f z / (z - z0) ^ Suc n" from assms(1,2) obtain r where r: "r > 0" "cball z0 r \ A" by (auto simp: open_contains_cball) have "(?f has_contour_integral 2 * pi * \ * residue ?f z0) (circlepath z0 r)" using r assms by (intro base_residue[of A]) (auto intro!: holomorphic_intros) moreover have "(?f has_contour_integral 2 * pi * \ / fact n * (deriv ^^ n) f z0) (circlepath z0 r)" using assms r by (intro Cauchy_has_contour_integral_higher_derivative_circlepath) (auto intro!: holomorphic_on_subset[OF assms(3)] holomorphic_on_imp_continuous_on) ultimately have "2 * pi * \ * residue ?f z0 = 2 * pi * \ / fact n * (deriv ^^ n) f z0" by (rule has_contour_integral_unique) thus ?thesis by (simp add: field_simps) qed lemma residue_holomorphic_over_power': assumes "open A" "0 \ A" "f holomorphic_on A" shows "residue (\z. f z / z ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" using residue_holomorphic_over_power[OF assms] by simp theorem residue_fps_expansion_over_power_at_0: assumes "f has_fps_expansion F" shows "residue (\z. f z / z ^ Suc n) 0 = fps_nth F n" proof - from has_fps_expansion_imp_holomorphic[OF assms] guess s . note s = this have "residue (\z. f z / (z - 0) ^ Suc n) 0 = (deriv ^^ n) f 0 / fact n" using assms s unfolding has_fps_expansion_def by (intro residue_holomorphic_over_power[of s]) (auto simp: zero_ereal_def) also from assms have "\ = fps_nth F n" by (subst fps_nth_fps_expansion) auto finally show ?thesis by simp qed lemma residue_pole_order: fixes f::"complex \ complex" and z::complex defines "n \ nat (- zorder f z)" and "h \ zor_poly f z" assumes f_iso:"isolated_singularity_at f z" and pole:"is_pole f z" shows "residue f z = ((deriv ^^ (n - 1)) h z / fact (n-1))" proof - define g where "g \ \x. if x=z then 0 else inverse (f x)" obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}" using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast obtain r where "0 < n" "0 < r" and r_cball:"cball z r \ ball z e" and h_holo: "h holomorphic_on cball z r" and h_divide:"(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" proof - obtain r where r:"zorder f z < 0" "h z \ 0" "r>0" "cball z r \ ball z e" "h holomorphic_on cball z r" "(\w\cball z r - {z}. f w = h w / (w - z) ^ n \ h w \ 0)" using zorder_exist_pole[OF f_holo,simplified,OF \is_pole f z\,folded n_def h_def] by auto have "n>0" using \zorder f z < 0\ unfolding n_def by simp moreover have "(\w\cball z r. (w\z \ f w = h w / (w - z) ^ n) \ h w \ 0)" using \h z\0\ r(6) by blast ultimately show ?thesis using r(3,4,5) that by blast qed have r_nonzero:"\w. w \ ball z r - {z} \ f w \ 0" using h_divide by simp define c where "c \ 2 * pi * \" define der_f where "der_f \ ((deriv ^^ (n - 1)) h z / fact (n-1))" define h' where "h' \ \u. h u / (u - z) ^ n" have "(h' has_contour_integral c / fact (n - 1) * (deriv ^^ (n - 1)) h z) (circlepath z r)" unfolding h'_def proof (rule Cauchy_has_contour_integral_higher_derivative_circlepath[of z r h z "n-1", folded c_def Suc_pred'[OF \n>0\]]) show "continuous_on (cball z r) h" using holomorphic_on_imp_continuous_on h_holo by simp show "h holomorphic_on ball z r" using h_holo by auto show " z \ ball z r" using \r>0\ by auto qed then have "(h' has_contour_integral c * der_f) (circlepath z r)" unfolding der_f_def by auto then have "(f has_contour_integral c * der_f) (circlepath z r)" proof (elim has_contour_integral_eq) fix x assume "x \ path_image (circlepath z r)" hence "x\cball z r - {z}" using \r>0\ by auto then show "h' x = f x" using h_divide unfolding h'_def by auto qed moreover have "(f has_contour_integral c * residue f z) (circlepath z r)" using base_residue[of \ball z e\ z,simplified,OF \r>0\ f_holo r_cball,folded c_def] unfolding c_def by simp ultimately have "c * der_f = c * residue f z" using has_contour_integral_unique by blast hence "der_f = residue f z" unfolding c_def by auto thus ?thesis unfolding der_f_def by auto qed lemma residue_simple_pole: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" shows "residue f z0 = zor_poly f z0 z0" using assms by (subst residue_pole_order) simp_all lemma residue_simple_pole_limit: assumes "isolated_singularity_at f z0" assumes "is_pole f z0" "zorder f z0 = - 1" assumes "((\x. f (g x) * (g x - z0)) \ c) F" assumes "filterlim g (at z0) F" "F \ bot" shows "residue f z0 = c" proof - have "residue f z0 = zor_poly f z0 z0" by (rule residue_simple_pole assms)+ also have "\ = c" apply (rule zor_poly_pole_eqI) using assms by auto finally show ?thesis . qed lemma assumes f_holo:"f holomorphic_on s" and g_holo:"g holomorphic_on s" and "open s" "connected s" "z \ s" assumes g_deriv:"(g has_field_derivative g') (at z)" assumes "f z \ 0" "g z = 0" "g' \ 0" shows porder_simple_pole_deriv: "zorder (\w. f w / g w) z = - 1" and residue_simple_pole_deriv: "residue (\w. f w / g w) z = f z / g'" proof - have [simp]:"isolated_singularity_at f z" "isolated_singularity_at g z" using isolated_singularity_at_holomorphic[OF _ \open s\ \z\s\] f_holo g_holo by (meson Diff_subset holomorphic_on_subset)+ have [simp]:"not_essential f z" "not_essential g z" unfolding not_essential_def using f_holo g_holo assms(3,5) by (meson continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on)+ have g_nconst:"\\<^sub>F w in at z. g w \0 " proof (rule ccontr) assume "\ (\\<^sub>F w in at z. g w \ 0)" then have "\\<^sub>F w in nhds z. g w = 0" unfolding eventually_at eventually_nhds frequently_at using \g z = 0\ by (metis open_ball UNIV_I centre_in_ball dist_commute mem_ball) then have "deriv g z = deriv (\_. 0) z" by (intro deriv_cong_ev) auto then have "deriv g z = 0" by auto then have "g' = 0" using g_deriv DERIV_imp_deriv by blast then show False using \g'\0\ by auto qed have "zorder (\w. f w / g w) z = zorder f z - zorder g z" proof - have "\\<^sub>F w in at z. f w \0 \ w\s" apply (rule non_zero_neighbour_alt) using assms by auto with g_nconst have "\\<^sub>F w in at z. f w * g w \ 0" by (elim frequently_rev_mp eventually_rev_mp,auto) then show ?thesis using zorder_divide[of f z g] by auto qed moreover have "zorder f z=0" apply (rule zorder_zero_eqI[OF f_holo \open s\ \z\s\]) using \f z\0\ by auto moreover have "zorder g z=1" apply (rule zorder_zero_eqI[OF g_holo \open s\ \z\s\]) subgoal using assms(8) by auto subgoal using DERIV_imp_deriv assms(9) g_deriv by auto subgoal by simp done ultimately show "zorder (\w. f w / g w) z = - 1" by auto show "residue (\w. f w / g w) z = f z / g'" proof (rule residue_simple_pole_limit[where g=id and F="at z",simplified]) show "zorder (\w. f w / g w) z = - 1" by fact show "isolated_singularity_at (\w. f w / g w) z" by (auto intro: singularity_intros) show "is_pole (\w. f w / g w) z" proof (rule is_pole_divide) have "\\<^sub>F x in at z. g x \ 0" apply (rule non_zero_neighbour) using g_nconst by auto moreover have "g \z\ 0" using DERIV_isCont assms(8) continuous_at g_deriv by force ultimately show "filterlim g (at 0) (at z)" unfolding filterlim_at by simp show "isCont f z" using assms(3,5) continuous_on_eq_continuous_at f_holo holomorphic_on_imp_continuous_on by auto show "f z \ 0" by fact qed show "filterlim id (at z) (at z)" by (simp add: filterlim_iff) have "((\w. (f w * (w - z)) / g w) \ f z / g') (at z)" proof (rule lhopital_complex_simple) show "((\w. f w * (w - z)) has_field_derivative f z) (at z)" using assms by (auto intro!: derivative_eq_intros holomorphic_derivI[OF f_holo]) show "(g has_field_derivative g') (at z)" by fact qed (insert assms, auto) then show "((\w. (f w / g w) * (w - z)) \ f z / g') (at z)" by (simp add: field_split_simps) qed qed subsection \Poles and residues of some well-known functions\ (* TODO: add more material here for other functions *) lemma is_pole_Gamma: "is_pole Gamma (-of_nat n)" unfolding is_pole_def using Gamma_poles . lemma Gamma_residue: "residue Gamma (-of_nat n) = (-1) ^ n / fact n" proof (rule residue_simple') show "open (- (\\<^sub>\\<^sub>0 - {-of_nat n}) :: complex set)" by (intro open_Compl closed_subset_Ints) auto show "Gamma holomorphic_on (- (\\<^sub>\\<^sub>0 - {-of_nat n}) - {- of_nat n})" by (rule holomorphic_Gamma) auto show "(\w. Gamma w * (w - (-of_nat n))) \(-of_nat n)\ (- 1) ^ n / fact n" using Gamma_residues[of n] by simp qed auto -end \ No newline at end of file +end diff --git a/src/HOL/Nonstandard_Analysis/NSComplex.thy b/src/HOL/Nonstandard_Analysis/NSComplex.thy --- a/src/HOL/Nonstandard_Analysis/NSComplex.thy +++ b/src/HOL/Nonstandard_Analysis/NSComplex.thy @@ -1,603 +1,603 @@ (* Title: HOL/Nonstandard_Analysis/NSComplex.thy Author: Jacques D. Fleuriot, University of Edinburgh Author: Lawrence C Paulson *) section \Nonstandard Complex Numbers\ theory NSComplex imports NSA begin type_synonym hcomplex = "complex star" abbreviation hcomplex_of_complex :: "complex \ complex star" where "hcomplex_of_complex \ star_of" abbreviation hcmod :: "complex star \ real star" where "hcmod \ hnorm" subsubsection \Real and Imaginary parts\ definition hRe :: "hcomplex \ hypreal" where "hRe = *f* Re" definition hIm :: "hcomplex \ hypreal" where "hIm = *f* Im" subsubsection \Imaginary unit\ definition iii :: hcomplex where "iii = star_of \" subsubsection \Complex conjugate\ definition hcnj :: "hcomplex \ hcomplex" where "hcnj = *f* cnj" subsubsection \Argand\ definition hsgn :: "hcomplex \ hcomplex" where "hsgn = *f* sgn" definition harg :: "hcomplex \ hypreal" - where "harg = *f* arg" + where "harg = *f* Arg" definition \ \abbreviation for \cos a + i sin a\\ hcis :: "hypreal \ hcomplex" where "hcis = *f* cis" subsubsection \Injection from hyperreals\ abbreviation hcomplex_of_hypreal :: "hypreal \ hcomplex" where "hcomplex_of_hypreal \ of_hypreal" definition \ \abbreviation for \r * (cos a + i sin a)\\ hrcis :: "hypreal \ hypreal \ hcomplex" where "hrcis = *f2* rcis" subsubsection \\e ^ (x + iy)\\ definition hExp :: "hcomplex \ hcomplex" where "hExp = *f* exp" definition HComplex :: "hypreal \ hypreal \ hcomplex" where "HComplex = *f2* Complex" lemmas hcomplex_defs [transfer_unfold] = hRe_def hIm_def iii_def hcnj_def hsgn_def harg_def hcis_def hrcis_def hExp_def HComplex_def lemma Standard_hRe [simp]: "x \ Standard \ hRe x \ Standard" by (simp add: hcomplex_defs) lemma Standard_hIm [simp]: "x \ Standard \ hIm x \ Standard" by (simp add: hcomplex_defs) lemma Standard_iii [simp]: "iii \ Standard" by (simp add: hcomplex_defs) lemma Standard_hcnj [simp]: "x \ Standard \ hcnj x \ Standard" by (simp add: hcomplex_defs) lemma Standard_hsgn [simp]: "x \ Standard \ hsgn x \ Standard" by (simp add: hcomplex_defs) lemma Standard_harg [simp]: "x \ Standard \ harg x \ Standard" by (simp add: hcomplex_defs) lemma Standard_hcis [simp]: "r \ Standard \ hcis r \ Standard" by (simp add: hcomplex_defs) lemma Standard_hExp [simp]: "x \ Standard \ hExp x \ Standard" by (simp add: hcomplex_defs) lemma Standard_hrcis [simp]: "r \ Standard \ s \ Standard \ hrcis r s \ Standard" by (simp add: hcomplex_defs) lemma Standard_HComplex [simp]: "r \ Standard \ s \ Standard \ HComplex r s \ Standard" by (simp add: hcomplex_defs) lemma hcmod_def: "hcmod = *f* cmod" by (rule hnorm_def) subsection \Properties of Nonstandard Real and Imaginary Parts\ lemma hcomplex_hRe_hIm_cancel_iff: "\w z. w = z \ hRe w = hRe z \ hIm w = hIm z" by transfer (rule complex_eq_iff) lemma hcomplex_equality [intro?]: "\z w. hRe z = hRe w \ hIm z = hIm w \ z = w" by transfer (rule complex_eqI) lemma hcomplex_hRe_zero [simp]: "hRe 0 = 0" by transfer simp lemma hcomplex_hIm_zero [simp]: "hIm 0 = 0" by transfer simp lemma hcomplex_hRe_one [simp]: "hRe 1 = 1" by transfer simp lemma hcomplex_hIm_one [simp]: "hIm 1 = 0" by transfer simp subsection \Addition for Nonstandard Complex Numbers\ lemma hRe_add: "\x y. hRe (x + y) = hRe x + hRe y" by transfer simp lemma hIm_add: "\x y. hIm (x + y) = hIm x + hIm y" by transfer simp subsection \More Minus Laws\ lemma hRe_minus: "\z. hRe (- z) = - hRe z" by transfer (rule uminus_complex.sel) lemma hIm_minus: "\z. hIm (- z) = - hIm z" by transfer (rule uminus_complex.sel) lemma hcomplex_add_minus_eq_minus: "x + y = 0 \ x = - y" for x y :: hcomplex apply (drule minus_unique) apply (simp add: minus_equation_iff [of x y]) done lemma hcomplex_i_mult_eq [simp]: "iii * iii = - 1" by transfer (rule i_squared) lemma hcomplex_i_mult_left [simp]: "\z. iii * (iii * z) = - z" by transfer (rule complex_i_mult_minus) lemma hcomplex_i_not_zero [simp]: "iii \ 0" by transfer (rule complex_i_not_zero) subsection \More Multiplication Laws\ lemma hcomplex_mult_minus_one: "- 1 * z = - z" for z :: hcomplex by simp lemma hcomplex_mult_minus_one_right: "z * - 1 = - z" for z :: hcomplex by simp lemma hcomplex_mult_left_cancel: "c \ 0 \ c * a = c * b \ a = b" for a b c :: hcomplex by simp lemma hcomplex_mult_right_cancel: "c \ 0 \ a * c = b * c \ a = b" for a b c :: hcomplex by simp subsection \Subtraction and Division\ (* TODO: delete *) lemma hcomplex_diff_eq_eq [simp]: "x - y = z \ x = z + y" for x y z :: hcomplex by (rule diff_eq_eq) subsection \Embedding Properties for \<^term>\hcomplex_of_hypreal\ Map\ lemma hRe_hcomplex_of_hypreal [simp]: "\z. hRe (hcomplex_of_hypreal z) = z" by transfer (rule Re_complex_of_real) lemma hIm_hcomplex_of_hypreal [simp]: "\z. hIm (hcomplex_of_hypreal z) = 0" by transfer (rule Im_complex_of_real) lemma hcomplex_of_epsilon_not_zero [simp]: "hcomplex_of_hypreal \ \ 0" by (simp add: epsilon_not_zero) subsection \\HComplex\ theorems\ lemma hRe_HComplex [simp]: "\x y. hRe (HComplex x y) = x" by transfer simp lemma hIm_HComplex [simp]: "\x y. hIm (HComplex x y) = y" by transfer simp lemma hcomplex_surj [simp]: "\z. HComplex (hRe z) (hIm z) = z" by transfer (rule complex_surj) lemma hcomplex_induct [case_names rect(*, induct type: hcomplex*)]: "(\x y. P (HComplex x y)) \ P z" by (rule hcomplex_surj [THEN subst]) blast subsection \Modulus (Absolute Value) of Nonstandard Complex Number\ lemma hcomplex_of_hypreal_abs: "hcomplex_of_hypreal \x\ = hcomplex_of_hypreal (hcmod (hcomplex_of_hypreal x))" by simp lemma HComplex_inject [simp]: "\x y x' y'. HComplex x y = HComplex x' y' \ x = x' \ y = y'" by transfer (rule complex.inject) lemma HComplex_add [simp]: "\x1 y1 x2 y2. HComplex x1 y1 + HComplex x2 y2 = HComplex (x1 + x2) (y1 + y2)" by transfer (rule complex_add) lemma HComplex_minus [simp]: "\x y. - HComplex x y = HComplex (- x) (- y)" by transfer (rule complex_minus) lemma HComplex_diff [simp]: "\x1 y1 x2 y2. HComplex x1 y1 - HComplex x2 y2 = HComplex (x1 - x2) (y1 - y2)" by transfer (rule complex_diff) lemma HComplex_mult [simp]: "\x1 y1 x2 y2. HComplex x1 y1 * HComplex x2 y2 = HComplex (x1*x2 - y1*y2) (x1*y2 + y1*x2)" by transfer (rule complex_mult) text \\HComplex_inverse\ is proved below.\ lemma hcomplex_of_hypreal_eq: "\r. hcomplex_of_hypreal r = HComplex r 0" by transfer (rule complex_of_real_def) lemma HComplex_add_hcomplex_of_hypreal [simp]: "\x y r. HComplex x y + hcomplex_of_hypreal r = HComplex (x + r) y" by transfer (rule Complex_add_complex_of_real) lemma hcomplex_of_hypreal_add_HComplex [simp]: "\r x y. hcomplex_of_hypreal r + HComplex x y = HComplex (r + x) y" by transfer (rule complex_of_real_add_Complex) lemma HComplex_mult_hcomplex_of_hypreal: "\x y r. HComplex x y * hcomplex_of_hypreal r = HComplex (x * r) (y * r)" by transfer (rule Complex_mult_complex_of_real) lemma hcomplex_of_hypreal_mult_HComplex: "\r x y. hcomplex_of_hypreal r * HComplex x y = HComplex (r * x) (r * y)" by transfer (rule complex_of_real_mult_Complex) lemma i_hcomplex_of_hypreal [simp]: "\r. iii * hcomplex_of_hypreal r = HComplex 0 r" by transfer (rule i_complex_of_real) lemma hcomplex_of_hypreal_i [simp]: "\r. hcomplex_of_hypreal r * iii = HComplex 0 r" by transfer (rule complex_of_real_i) subsection \Conjugation\ lemma hcomplex_hcnj_cancel_iff [iff]: "\x y. hcnj x = hcnj y \ x = y" by transfer (rule complex_cnj_cancel_iff) lemma hcomplex_hcnj_hcnj [simp]: "\z. hcnj (hcnj z) = z" by transfer (rule complex_cnj_cnj) lemma hcomplex_hcnj_hcomplex_of_hypreal [simp]: "\x. hcnj (hcomplex_of_hypreal x) = hcomplex_of_hypreal x" by transfer (rule complex_cnj_complex_of_real) lemma hcomplex_hmod_hcnj [simp]: "\z. hcmod (hcnj z) = hcmod z" by transfer (rule complex_mod_cnj) lemma hcomplex_hcnj_minus: "\z. hcnj (- z) = - hcnj z" by transfer (rule complex_cnj_minus) lemma hcomplex_hcnj_inverse: "\z. hcnj (inverse z) = inverse (hcnj z)" by transfer (rule complex_cnj_inverse) lemma hcomplex_hcnj_add: "\w z. hcnj (w + z) = hcnj w + hcnj z" by transfer (rule complex_cnj_add) lemma hcomplex_hcnj_diff: "\w z. hcnj (w - z) = hcnj w - hcnj z" by transfer (rule complex_cnj_diff) lemma hcomplex_hcnj_mult: "\w z. hcnj (w * z) = hcnj w * hcnj z" by transfer (rule complex_cnj_mult) lemma hcomplex_hcnj_divide: "\w z. hcnj (w / z) = hcnj w / hcnj z" by transfer (rule complex_cnj_divide) lemma hcnj_one [simp]: "hcnj 1 = 1" by transfer (rule complex_cnj_one) lemma hcomplex_hcnj_zero [simp]: "hcnj 0 = 0" by transfer (rule complex_cnj_zero) lemma hcomplex_hcnj_zero_iff [iff]: "\z. hcnj z = 0 \ z = 0" by transfer (rule complex_cnj_zero_iff) lemma hcomplex_mult_hcnj: "\z. z * hcnj z = hcomplex_of_hypreal ((hRe z)\<^sup>2 + (hIm z)\<^sup>2)" by transfer (rule complex_mult_cnj) subsection \More Theorems about the Function \<^term>\hcmod\\ lemma hcmod_hcomplex_of_hypreal_of_nat [simp]: "hcmod (hcomplex_of_hypreal (hypreal_of_nat n)) = hypreal_of_nat n" by simp lemma hcmod_hcomplex_of_hypreal_of_hypnat [simp]: "hcmod (hcomplex_of_hypreal(hypreal_of_hypnat n)) = hypreal_of_hypnat n" by simp lemma hcmod_mult_hcnj: "\z. hcmod (z * hcnj z) = (hcmod z)\<^sup>2" by transfer (rule complex_mod_mult_cnj) lemma hcmod_triangle_ineq2 [simp]: "\a b. hcmod (b + a) - hcmod b \ hcmod a" by transfer (rule complex_mod_triangle_ineq2) lemma hcmod_diff_ineq [simp]: "\a b. hcmod a - hcmod b \ hcmod (a + b)" by transfer (rule norm_diff_ineq) subsection \Exponentiation\ lemma hcomplexpow_0 [simp]: "z ^ 0 = 1" for z :: hcomplex by (rule power_0) lemma hcomplexpow_Suc [simp]: "z ^ (Suc n) = z * (z ^ n)" for z :: hcomplex by (rule power_Suc) lemma hcomplexpow_i_squared [simp]: "iii\<^sup>2 = -1" by transfer (rule power2_i) lemma hcomplex_of_hypreal_pow: "\x. hcomplex_of_hypreal (x ^ n) = hcomplex_of_hypreal x ^ n" by transfer (rule of_real_power) lemma hcomplex_hcnj_pow: "\z. hcnj (z ^ n) = hcnj z ^ n" by transfer (rule complex_cnj_power) lemma hcmod_hcomplexpow: "\x. hcmod (x ^ n) = hcmod x ^ n" by transfer (rule norm_power) lemma hcpow_minus: "\x n. (- x :: hcomplex) pow n = (if ( *p* even) n then (x pow n) else - (x pow n))" by transfer simp lemma hcpow_mult: "(r * s) pow n = (r pow n) * (s pow n)" for r s :: hcomplex by (fact hyperpow_mult) lemma hcpow_zero2 [simp]: "\n. 0 pow (hSuc n) = (0::'a::semiring_1 star)" by transfer (rule power_0_Suc) lemma hcpow_not_zero [simp,intro]: "\r n. r \ 0 \ r pow n \ (0::hcomplex)" by (fact hyperpow_not_zero) lemma hcpow_zero_zero: "r pow n = 0 \ r = 0" for r :: hcomplex by (blast intro: ccontr dest: hcpow_not_zero) subsection \The Function \<^term>\hsgn\\ lemma hsgn_zero [simp]: "hsgn 0 = 0" by transfer (rule sgn_zero) lemma hsgn_one [simp]: "hsgn 1 = 1" by transfer (rule sgn_one) lemma hsgn_minus: "\z. hsgn (- z) = - hsgn z" by transfer (rule sgn_minus) lemma hsgn_eq: "\z. hsgn z = z / hcomplex_of_hypreal (hcmod z)" by transfer (rule sgn_eq) lemma hcmod_i: "\x y. hcmod (HComplex x y) = ( *f* sqrt) (x\<^sup>2 + y\<^sup>2)" by transfer (rule complex_norm) lemma hcomplex_eq_cancel_iff1 [simp]: "hcomplex_of_hypreal xa = HComplex x y \ xa = x \ y = 0" by (simp add: hcomplex_of_hypreal_eq) lemma hcomplex_eq_cancel_iff2 [simp]: "HComplex x y = hcomplex_of_hypreal xa \ x = xa \ y = 0" by (simp add: hcomplex_of_hypreal_eq) lemma HComplex_eq_0 [simp]: "\x y. HComplex x y = 0 \ x = 0 \ y = 0" by transfer (rule Complex_eq_0) lemma HComplex_eq_1 [simp]: "\x y. HComplex x y = 1 \ x = 1 \ y = 0" by transfer (rule Complex_eq_1) lemma i_eq_HComplex_0_1: "iii = HComplex 0 1" by transfer (simp add: complex_eq_iff) lemma HComplex_eq_i [simp]: "\x y. HComplex x y = iii \ x = 0 \ y = 1" by transfer (rule Complex_eq_i) lemma hRe_hsgn [simp]: "\z. hRe (hsgn z) = hRe z / hcmod z" by transfer (rule Re_sgn) lemma hIm_hsgn [simp]: "\z. hIm (hsgn z) = hIm z / hcmod z" by transfer (rule Im_sgn) lemma HComplex_inverse: "\x y. inverse (HComplex x y) = HComplex (x / (x\<^sup>2 + y\<^sup>2)) (- y / (x\<^sup>2 + y\<^sup>2))" by transfer (rule complex_inverse) lemma hRe_mult_i_eq[simp]: "\y. hRe (iii * hcomplex_of_hypreal y) = 0" by transfer simp lemma hIm_mult_i_eq [simp]: "\y. hIm (iii * hcomplex_of_hypreal y) = y" by transfer simp lemma hcmod_mult_i [simp]: "\y. hcmod (iii * hcomplex_of_hypreal y) = \y\" by transfer (simp add: norm_complex_def) lemma hcmod_mult_i2 [simp]: "\y. hcmod (hcomplex_of_hypreal y * iii) = \y\" by transfer (simp add: norm_complex_def) subsubsection \\harg\\ lemma cos_harg_i_mult_zero [simp]: "\y. y \ 0 \ ( *f* cos) (harg (HComplex 0 y)) = 0" by transfer (simp add: Complex_eq) subsection \Polar Form for Nonstandard Complex Numbers\ lemma complex_split_polar2: "\n. \r a. (z n) = complex_of_real r * Complex (cos a) (sin a)" unfolding Complex_eq by (auto intro: complex_split_polar) lemma hcomplex_split_polar: "\z. \r a. z = hcomplex_of_hypreal r * (HComplex (( *f* cos) a) (( *f* sin) a))" by transfer (simp add: Complex_eq complex_split_polar) lemma hcis_eq: "\a. hcis a = hcomplex_of_hypreal (( *f* cos) a) + iii * hcomplex_of_hypreal (( *f* sin) a)" by transfer (simp add: complex_eq_iff) lemma hrcis_Ex: "\z. \r a. z = hrcis r a" by transfer (rule rcis_Ex) lemma hRe_hcomplex_polar [simp]: "\r a. hRe (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* cos) a" by transfer simp lemma hRe_hrcis [simp]: "\r a. hRe (hrcis r a) = r * ( *f* cos) a" by transfer (rule Re_rcis) lemma hIm_hcomplex_polar [simp]: "\r a. hIm (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = r * ( *f* sin) a" by transfer simp lemma hIm_hrcis [simp]: "\r a. hIm (hrcis r a) = r * ( *f* sin) a" by transfer (rule Im_rcis) lemma hcmod_unit_one [simp]: "\a. hcmod (HComplex (( *f* cos) a) (( *f* sin) a)) = 1" by transfer (simp add: cmod_unit_one) lemma hcmod_complex_polar [simp]: "\r a. hcmod (hcomplex_of_hypreal r * HComplex (( *f* cos) a) (( *f* sin) a)) = \r\" by transfer (simp add: Complex_eq cmod_complex_polar) lemma hcmod_hrcis [simp]: "\r a. hcmod(hrcis r a) = \r\" by transfer (rule complex_mod_rcis) text \\(r1 * hrcis a) * (r2 * hrcis b) = r1 * r2 * hrcis (a + b)\\ lemma hcis_hrcis_eq: "\a. hcis a = hrcis 1 a" by transfer (rule cis_rcis_eq) declare hcis_hrcis_eq [symmetric, simp] lemma hrcis_mult: "\a b r1 r2. hrcis r1 a * hrcis r2 b = hrcis (r1 * r2) (a + b)" by transfer (rule rcis_mult) lemma hcis_mult: "\a b. hcis a * hcis b = hcis (a + b)" by transfer (rule cis_mult) lemma hcis_zero [simp]: "hcis 0 = 1" by transfer (rule cis_zero) lemma hrcis_zero_mod [simp]: "\a. hrcis 0 a = 0" by transfer (rule rcis_zero_mod) lemma hrcis_zero_arg [simp]: "\r. hrcis r 0 = hcomplex_of_hypreal r" by transfer (rule rcis_zero_arg) lemma hcomplex_i_mult_minus [simp]: "\x. iii * (iii * x) = - x" by transfer (rule complex_i_mult_minus) lemma hcomplex_i_mult_minus2 [simp]: "iii * iii * x = - x" by simp lemma hcis_hypreal_of_nat_Suc_mult: "\a. hcis (hypreal_of_nat (Suc n) * a) = hcis a * hcis (hypreal_of_nat n * a)" by transfer (simp add: distrib_right cis_mult) lemma NSDeMoivre: "\a. (hcis a) ^ n = hcis (hypreal_of_nat n * a)" by transfer (rule DeMoivre) lemma hcis_hypreal_of_hypnat_Suc_mult: "\a n. hcis (hypreal_of_hypnat (n + 1) * a) = hcis a * hcis (hypreal_of_hypnat n * a)" by transfer (simp add: distrib_right cis_mult) lemma NSDeMoivre_ext: "\a n. (hcis a) pow n = hcis (hypreal_of_hypnat n * a)" by transfer (rule DeMoivre) lemma NSDeMoivre2: "\a r. (hrcis r a) ^ n = hrcis (r ^ n) (hypreal_of_nat n * a)" by transfer (rule DeMoivre2) lemma DeMoivre2_ext: "\a r n. (hrcis r a) pow n = hrcis (r pow n) (hypreal_of_hypnat n * a)" by transfer (rule DeMoivre2) lemma hcis_inverse [simp]: "\a. inverse (hcis a) = hcis (- a)" by transfer (rule cis_inverse) lemma hrcis_inverse: "\a r. inverse (hrcis r a) = hrcis (inverse r) (- a)" by transfer (simp add: rcis_inverse inverse_eq_divide [symmetric]) lemma hRe_hcis [simp]: "\a. hRe (hcis a) = ( *f* cos) a" by transfer simp lemma hIm_hcis [simp]: "\a. hIm (hcis a) = ( *f* sin) a" by transfer simp lemma cos_n_hRe_hcis_pow_n: "( *f* cos) (hypreal_of_nat n * a) = hRe (hcis a ^ n)" by (simp add: NSDeMoivre) lemma sin_n_hIm_hcis_pow_n: "( *f* sin) (hypreal_of_nat n * a) = hIm (hcis a ^ n)" by (simp add: NSDeMoivre) lemma cos_n_hRe_hcis_hcpow_n: "( *f* cos) (hypreal_of_hypnat n * a) = hRe (hcis a pow n)" by (simp add: NSDeMoivre_ext) lemma sin_n_hIm_hcis_hcpow_n: "( *f* sin) (hypreal_of_hypnat n * a) = hIm (hcis a pow n)" by (simp add: NSDeMoivre_ext) lemma hExp_add: "\a b. hExp (a + b) = hExp a * hExp b" by transfer (rule exp_add) subsection \\<^term>\hcomplex_of_complex\: the Injection from type \<^typ>\complex\ to to \<^typ>\hcomplex\\ lemma hcomplex_of_complex_i: "iii = hcomplex_of_complex \" by (rule iii_def) lemma hRe_hcomplex_of_complex: "hRe (hcomplex_of_complex z) = hypreal_of_real (Re z)" by transfer (rule refl) lemma hIm_hcomplex_of_complex: "hIm (hcomplex_of_complex z) = hypreal_of_real (Im z)" by transfer (rule refl) lemma hcmod_hcomplex_of_complex: "hcmod (hcomplex_of_complex x) = hypreal_of_real (cmod x)" by transfer (rule refl) subsection \Numerals and Arithmetic\ lemma hcomplex_of_hypreal_eq_hcomplex_of_complex: "hcomplex_of_hypreal (hypreal_of_real x) = hcomplex_of_complex (complex_of_real x)" by transfer (rule refl) lemma hcomplex_hypreal_numeral: "hcomplex_of_complex (numeral w) = hcomplex_of_hypreal(numeral w)" by transfer (rule of_real_numeral [symmetric]) lemma hcomplex_hypreal_neg_numeral: "hcomplex_of_complex (- numeral w) = hcomplex_of_hypreal(- numeral w)" by transfer (rule of_real_neg_numeral [symmetric]) lemma hcomplex_numeral_hcnj [simp]: "hcnj (numeral v :: hcomplex) = numeral v" by transfer (rule complex_cnj_numeral) lemma hcomplex_numeral_hcmod [simp]: "hcmod (numeral v :: hcomplex) = (numeral v :: hypreal)" by transfer (rule norm_numeral) lemma hcomplex_neg_numeral_hcmod [simp]: "hcmod (- numeral v :: hcomplex) = (numeral v :: hypreal)" by transfer (rule norm_neg_numeral) lemma hcomplex_numeral_hRe [simp]: "hRe (numeral v :: hcomplex) = numeral v" by transfer (rule complex_Re_numeral) lemma hcomplex_numeral_hIm [simp]: "hIm (numeral v :: hcomplex) = 0" by transfer (rule complex_Im_numeral) end diff --git a/src/HOL/Series.thy b/src/HOL/Series.thy --- a/src/HOL/Series.thy +++ b/src/HOL/Series.thy @@ -1,1333 +1,1329 @@ (* Title : Series.thy Author : Jacques D. Fleuriot Copyright : 1998 University of Cambridge Converted to Isar and polished by lcp Converted to sum and polished yet more by TNN Additional contributions by Jeremy Avigad *) section \Infinite Series\ theory Series imports Limits Inequalities begin subsection \Definition of infinite summability\ definition sums :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a \ bool" (infixr "sums" 80) where "f sums s \ (\n. \i s" definition summable :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ bool" where "summable f \ (\s. f sums s)" definition suminf :: "(nat \ 'a::{topological_space, comm_monoid_add}) \ 'a" (binder "\" 10) where "suminf f = (THE s. f sums s)" text\Variants of the definition\ lemma sums_def': "f sums s \ (\n. \i = 0..n. f i) \ s" unfolding sums_def apply (subst filterlim_sequentially_Suc [symmetric]) apply (simp only: lessThan_Suc_atMost atLeast0AtMost) done lemma sums_def_le: "f sums s \ (\n. \i\n. f i) \ s" by (simp add: sums_def' atMost_atLeast0) lemma bounded_imp_summable: fixes a :: "nat \ 'a::{conditionally_complete_linorder,linorder_topology,linordered_comm_semiring_strict}" assumes 0: "\n. a n \ 0" and bounded: "\n. (\k\n. a k) \ B" shows "summable a" proof - have "bdd_above (range(\n. \k\n. a k))" by (meson bdd_aboveI2 bounded) moreover have "incseq (\n. \k\n. a k)" by (simp add: mono_def "0" sum_mono2) ultimately obtain s where "(\n. \k\n. a k) \ s" using LIMSEQ_incseq_SUP by blast then show ?thesis by (auto simp: sums_def_le summable_def) qed subsection \Infinite summability on topological monoids\ lemma sums_subst[trans]: "f = g \ g sums z \ f sums z" by simp lemma sums_cong: "(\n. f n = g n) \ f sums c \ g sums c" - by (drule ext) simp + by presburger lemma sums_summable: "f sums l \ summable f" by (simp add: sums_def summable_def, blast) lemma summable_iff_convergent: "summable f \ convergent (\n. \i convergent (\n. sum f {..n})" - by (simp_all only: summable_iff_convergent convergent_def - lessThan_Suc_atMost [symmetric] filterlim_sequentially_Suc[of "\n. sum f {..n. \in. 0) sums 0" unfolding sums_def by simp lemma summable_zero[simp, intro]: "summable (\n. 0)" by (rule sums_zero [THEN sums_summable]) lemma sums_group: "f sums s \ 0 < k \ (\n. sum f {n * k ..< n * k + k}) sums s" apply (simp only: sums_def sum.nat_group tendsto_def eventually_sequentially) apply (erule all_forward imp_forward exE| assumption)+ - apply (rule_tac x="N" in exI) by (metis le_square mult.commute mult.left_neutral mult_le_cancel2 mult_le_mono) lemma suminf_cong: "(\n. f n = g n) \ suminf f = suminf g" - by (rule arg_cong[of f g], rule ext) simp + by presburger lemma summable_cong: fixes f g :: "nat \ 'a::real_normed_vector" assumes "eventually (\x. f x = g x) sequentially" shows "summable f = summable g" proof - from assms obtain N where N: "\n\N. f n = g n" by (auto simp: eventually_at_top_linorder) define C where "C = (\kn. sum f {.. {N.. {N.. {N..n. n \ N \ f n = 0" shows "f sums (\n\N. f n)" proof - have eq: "sum f {..n. f n = 0) \ (f sums 0)" by (metis (no_types) finite.emptyI sum.empty sums_finite) lemma summable_finite: "finite N \ (\n. n \ N \ f n = 0) \ summable f" by (rule sums_summable) (rule sums_finite) lemma sums_If_finite_set: "finite A \ (\r. if r \ A then f r else 0) sums (\r\A. f r)" using sums_finite[of A "(\r. if r \ A then f r else 0)"] by simp lemma summable_If_finite_set[simp, intro]: "finite A \ summable (\r. if r \ A then f r else 0)" by (rule sums_summable) (rule sums_If_finite_set) lemma sums_If_finite: "finite {r. P r} \ (\r. if P r then f r else 0) sums (\r | P r. f r)" using sums_If_finite_set[of "{r. P r}"] by simp lemma summable_If_finite[simp, intro]: "finite {r. P r} \ summable (\r. if P r then f r else 0)" by (rule sums_summable) (rule sums_If_finite) lemma sums_single: "(\r. if r = i then f r else 0) sums f i" using sums_If_finite[of "\r. r = i"] by simp lemma summable_single[simp, intro]: "summable (\r. if r = i then f r else 0)" by (rule sums_summable) (rule sums_single) context fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" begin lemma summable_sums[intro]: "summable f \ f sums (suminf f)" by (simp add: summable_def sums_def suminf_def) (metis convergent_LIMSEQ_iff convergent_def lim_def) lemma summable_LIMSEQ: "summable f \ (\n. \i suminf f" by (rule summable_sums [unfolded sums_def]) lemma summable_LIMSEQ': "summable f \ (\n. \i\n. f i) \ suminf f" using sums_def_le by blast lemma sums_unique: "f sums s \ s = suminf f" by (metis limI suminf_eq_lim sums_def) lemma sums_iff: "f sums x \ summable f \ suminf f = x" by (metis summable_sums sums_summable sums_unique) lemma summable_sums_iff: "summable f \ f sums suminf f" by (auto simp: sums_iff summable_sums) lemma sums_unique2: "f sums a \ f sums b \ a = b" for a b :: 'a by (simp add: sums_iff) lemma sums_Uniq: "\\<^sub>\\<^sub>1a. f sums a" for a b :: 'a by (simp add: sums_unique2 Uniq_def) lemma suminf_finite: assumes N: "finite N" and f: "\n. n \ N \ f n = 0" shows "suminf f = (\n\N. f n)" using sums_finite[OF assms, THEN sums_unique] by simp end lemma suminf_zero[simp]: "suminf (\n. 0::'a::{t2_space, comm_monoid_add}) = 0" by (rule sums_zero [THEN sums_unique, symmetric]) subsection \Infinite summability on ordered, topological monoids\ lemma sums_le: "(\n. f n \ g n) \ f sums s \ g sums t \ s \ t" for f g :: "nat \ 'a::{ordered_comm_monoid_add,linorder_topology}" by (rule LIMSEQ_le) (auto intro: sum_mono simp: sums_def) context fixes f :: "nat \ 'a::{ordered_comm_monoid_add,linorder_topology}" begin lemma suminf_le: "(\n. f n \ g n) \ summable f \ summable g \ suminf f \ suminf g" using sums_le by blast lemma sum_le_suminf: shows "summable f \ finite I \ (\n. n \- I \ 0 \ f n) \ sum f I \ suminf f" by (rule sums_le[OF _ sums_If_finite_set summable_sums]) auto lemma suminf_nonneg: "summable f \ (\n. 0 \ f n) \ 0 \ suminf f" using sum_le_suminf by force lemma suminf_le_const: "summable f \ (\n. sum f {.. x) \ suminf f \ x" by (metis LIMSEQ_le_const2 summable_LIMSEQ) lemma suminf_eq_zero_iff: assumes "summable f" and pos: "\n. 0 \ f n" shows "suminf f = 0 \ (\n. f n = 0)" proof - assume "suminf f = 0" + assume L: "suminf f = 0" then have f: "(\n. \i 0" using summable_LIMSEQ[of f] assms by simp then have "\i. (\n\{i}. f n) \ 0" - proof (rule LIMSEQ_le_const) - show "\N. \n\N. (\n\{i}. f n) \ sum f {..summable f\ order_refl pos sum.infinite sum_le_suminf) with pos show "\n. f n = 0" - by (auto intro!: antisym) + by (simp add: order.antisym) qed (metis suminf_zero fun_eq_iff) lemma suminf_pos_iff: "summable f \ (\n. 0 \ f n) \ 0 < suminf f \ (\i. 0 < f i)" using sum_le_suminf[of "{}"] suminf_eq_zero_iff by (simp add: less_le) lemma suminf_pos2: assumes "summable f" "\n. 0 \ f n" "0 < f i" shows "0 < suminf f" proof - have "0 < (\n \ suminf f" using assms by (intro sum_le_suminf) auto finally show ?thesis . qed lemma suminf_pos: "summable f \ (\n. 0 < f n) \ 0 < suminf f" by (intro suminf_pos2[where i=0]) (auto intro: less_imp_le) end context fixes f :: "nat \ 'a::{ordered_cancel_comm_monoid_add,linorder_topology}" begin lemma sum_less_suminf2: "summable f \ (\m. m\n \ 0 \ f m) \ n \ i \ 0 < f i \ sum f {.. (\m. m\n \ 0 < f m) \ sum f {.. 'a::{ordered_comm_monoid_add,linorder_topology,conditionally_complete_linorder}" assumes pos[simp]: "\n. 0 \ f n" and le: "\n. (\i x" shows "summable f" unfolding summable_def sums_def [abs_def] proof (rule exI LIMSEQ_incseq_SUP)+ show "bdd_above (range (\n. sum f {..n. sum f {.. 'a::{canonically_ordered_monoid_add,linorder_topology,complete_linorder}" by (intro summableI_nonneg_bounded[where x=top] zero_le top_greatest) lemma suminf_eq_SUP_real: assumes X: "summable X" "\i. 0 \ X i" shows "suminf X = (SUP i. \nInfinite summability on topological monoids\ context fixes f g :: "nat \ 'a::{t2_space,topological_comm_monoid_add}" begin lemma sums_Suc: assumes "(\n. f (Suc n)) sums l" shows "f sums (l + f 0)" proof - have "(\n. (\i l + f 0" using assms by (auto intro!: tendsto_add simp: sums_def) moreover have "(\ii g sums b \ (\n. f n + g n) sums (a + b)" unfolding sums_def by (simp add: sum.distrib tendsto_add) lemma summable_add: "summable f \ summable g \ summable (\n. f n + g n)" unfolding summable_def by (auto intro: sums_add) lemma suminf_add: "summable f \ summable g \ suminf f + suminf g = (\n. f n + g n)" by (intro sums_unique sums_add summable_sums) end context fixes f :: "'i \ nat \ 'a::{t2_space,topological_comm_monoid_add}" and I :: "'i set" begin lemma sums_sum: "(\i. i \ I \ (f i) sums (x i)) \ (\n. \i\I. f i n) sums (\i\I. x i)" by (induct I rule: infinite_finite_induct) (auto intro!: sums_add) lemma suminf_sum: "(\i. i \ I \ summable (f i)) \ (\n. \i\I. f i n) = (\i\I. \n. f i n)" using sums_unique[OF sums_sum, OF summable_sums] by simp lemma summable_sum: "(\i. i \ I \ summable (f i)) \ summable (\n. \i\I. f i n)" using sums_summable[OF sums_sum[OF summable_sums]] . end lemma sums_If_finite_set': fixes f g :: "nat \ 'a::{t2_space,topological_ab_group_add}" assumes "g sums S" and "finite A" and "S' = S + (\n\A. f n - g n)" shows "(\n. if n \ A then f n else g n) sums S'" proof - have "(\n. g n + (if n \ A then f n - g n else 0)) sums (S + (\n\A. f n - g n))" by (intro sums_add assms sums_If_finite_set) also have "(\n. g n + (if n \ A then f n - g n else 0)) = (\n. if n \ A then f n else g n)" by (simp add: fun_eq_iff) finally show ?thesis using assms by simp qed subsection \Infinite summability on real normed vector spaces\ context fixes f :: "nat \ 'a::real_normed_vector" begin lemma sums_Suc_iff: "(\n. f (Suc n)) sums s \ f sums (s + f 0)" proof - have "f sums (s + f 0) \ (\i. \j s + f 0" by (subst filterlim_sequentially_Suc) (simp add: sums_def) also have "\ \ (\i. (\j s + f 0" by (simp add: ac_simps lessThan_Suc_eq_insert_0 image_Suc_lessThan sum.atLeast1_atMost_eq) also have "\ \ (\n. f (Suc n)) sums s" proof assume "(\i. (\j s + f 0" with tendsto_add[OF this tendsto_const, of "- f 0"] show "(\i. f (Suc i)) sums s" by (simp add: sums_def) qed (auto intro: tendsto_add simp: sums_def) finally show ?thesis .. qed lemma summable_Suc_iff: "summable (\n. f (Suc n)) = summable f" proof assume "summable f" then have "f sums suminf f" by (rule summable_sums) then have "(\n. f (Suc n)) sums (suminf f - f 0)" by (simp add: sums_Suc_iff) then show "summable (\n. f (Suc n))" unfolding summable_def by blast qed (auto simp: sums_Suc_iff summable_def) lemma sums_Suc_imp: "f 0 = 0 \ (\n. f (Suc n)) sums s \ (\n. f n) sums s" using sums_Suc_iff by simp end context (* Separate contexts are necessary to allow general use of the results above, here. *) fixes f :: "nat \ 'a::real_normed_vector" begin lemma sums_diff: "f sums a \ g sums b \ (\n. f n - g n) sums (a - b)" unfolding sums_def by (simp add: sum_subtractf tendsto_diff) lemma summable_diff: "summable f \ summable g \ summable (\n. f n - g n)" unfolding summable_def by (auto intro: sums_diff) lemma suminf_diff: "summable f \ summable g \ suminf f - suminf g = (\n. f n - g n)" by (intro sums_unique sums_diff summable_sums) lemma sums_minus: "f sums a \ (\n. - f n) sums (- a)" unfolding sums_def by (simp add: sum_negf tendsto_minus) lemma summable_minus: "summable f \ summable (\n. - f n)" unfolding summable_def by (auto intro: sums_minus) lemma suminf_minus: "summable f \ (\n. - f n) = - (\n. f n)" by (intro sums_unique [symmetric] sums_minus summable_sums) lemma sums_iff_shift: "(\i. f (i + n)) sums s \ f sums (s + (\ii. f (Suc i + n)) sums s \ (\i. f (i + n)) sums (s + f n)" by (subst sums_Suc_iff) simp with Suc show ?case by (simp add: ac_simps) qed corollary sums_iff_shift': "(\i. f (i + n)) sums (s - (\i f sums s" by (simp add: sums_iff_shift) lemma sums_zero_iff_shift: assumes "\i. i < n \ f i = 0" shows "(\i. f (i+n)) sums s \ (\i. f i) sums s" by (simp add: assms sums_iff_shift) lemma summable_iff_shift [simp]: "summable (\n. f (n + k)) \ summable f" by (metis diff_add_cancel summable_def sums_iff_shift [abs_def]) lemma sums_split_initial_segment: "f sums s \ (\i. f (i + n)) sums (s - (\i summable (\n. f(n + k))" by (simp add: summable_iff_shift) lemma suminf_minus_initial_segment: "summable f \ (\n. f (n + k)) = (\n. f n) - (\i suminf f = (\n. f(n + k)) + (\i (\n. f (Suc n)) = suminf f - f 0" using suminf_split_initial_segment[of 1] by simp lemma suminf_exist_split: fixes r :: real assumes "0 < r" and "summable f" shows "\N. \n\N. norm (\i. f (i + n)) < r" proof - from LIMSEQ_D[OF summable_LIMSEQ[OF \summable f\] \0 < r\] obtain N :: nat where "\ n \ N. norm (sum f {..summable f\]) qed lemma summable_LIMSEQ_zero: assumes "summable f" shows "f \ 0" proof - have "Cauchy (\n. sum f {.. convergent f" by (force dest!: summable_LIMSEQ_zero simp: convergent_def) lemma summable_imp_Bseq: "summable f \ Bseq f" by (simp add: convergent_imp_Bseq summable_imp_convergent) end lemma summable_minus_iff: "summable (\n. - f n) \ summable f" for f :: "nat \ 'a::real_normed_vector" by (auto dest: summable_minus) (* used two ways, hence must be outside the context above *) lemma (in bounded_linear) sums: "(\n. X n) sums a \ (\n. f (X n)) sums (f a)" unfolding sums_def by (drule tendsto) (simp only: sum) lemma (in bounded_linear) summable: "summable (\n. X n) \ summable (\n. f (X n))" unfolding summable_def by (auto intro: sums) lemma (in bounded_linear) suminf: "summable (\n. X n) \ f (\n. X n) = (\n. f (X n))" by (intro sums_unique sums summable_sums) lemmas sums_of_real = bounded_linear.sums [OF bounded_linear_of_real] lemmas summable_of_real = bounded_linear.summable [OF bounded_linear_of_real] lemmas suminf_of_real = bounded_linear.suminf [OF bounded_linear_of_real] lemmas sums_scaleR_left = bounded_linear.sums[OF bounded_linear_scaleR_left] lemmas summable_scaleR_left = bounded_linear.summable[OF bounded_linear_scaleR_left] lemmas suminf_scaleR_left = bounded_linear.suminf[OF bounded_linear_scaleR_left] lemmas sums_scaleR_right = bounded_linear.sums[OF bounded_linear_scaleR_right] lemmas summable_scaleR_right = bounded_linear.summable[OF bounded_linear_scaleR_right] lemmas suminf_scaleR_right = bounded_linear.suminf[OF bounded_linear_scaleR_right] lemma summable_const_iff: "summable (\_. c) \ c = 0" for c :: "'a::real_normed_vector" proof - have "\ summable (\_. c)" if "c \ 0" proof - from that have "filterlim (\n. of_nat n * norm c) at_top sequentially" by (subst mult.commute) (auto intro!: filterlim_tendsto_pos_mult_at_top filterlim_real_sequentially) then have "\ convergent (\n. norm (\kInfinite summability on real normed algebras\ context fixes f :: "nat \ 'a::real_normed_algebra" begin lemma sums_mult: "f sums a \ (\n. c * f n) sums (c * a)" by (rule bounded_linear.sums [OF bounded_linear_mult_right]) lemma summable_mult: "summable f \ summable (\n. c * f n)" by (rule bounded_linear.summable [OF bounded_linear_mult_right]) lemma suminf_mult: "summable f \ suminf (\n. c * f n) = c * suminf f" by (rule bounded_linear.suminf [OF bounded_linear_mult_right, symmetric]) lemma sums_mult2: "f sums a \ (\n. f n * c) sums (a * c)" by (rule bounded_linear.sums [OF bounded_linear_mult_left]) lemma summable_mult2: "summable f \ summable (\n. f n * c)" by (rule bounded_linear.summable [OF bounded_linear_mult_left]) lemma suminf_mult2: "summable f \ suminf f * c = (\n. f n * c)" by (rule bounded_linear.suminf [OF bounded_linear_mult_left]) end lemma sums_mult_iff: fixes f :: "nat \ 'a::{real_normed_algebra,field}" assumes "c \ 0" shows "(\n. c * f n) sums (c * d) \ f sums d" using sums_mult[of f d c] sums_mult[of "\n. c * f n" "c * d" "inverse c"] by (force simp: field_simps assms) lemma sums_mult2_iff: fixes f :: "nat \ 'a::{real_normed_algebra,field}" assumes "c \ 0" shows "(\n. f n * c) sums (d * c) \ f sums d" using sums_mult_iff[OF assms, of f d] by (simp add: mult.commute) lemma sums_of_real_iff: "(\n. of_real (f n) :: 'a::real_normed_div_algebra) sums of_real c \ f sums c" by (simp add: sums_def of_real_sum[symmetric] tendsto_of_real_iff del: of_real_sum) subsection \Infinite summability on real normed fields\ context fixes c :: "'a::real_normed_field" begin lemma sums_divide: "f sums a \ (\n. f n / c) sums (a / c)" by (rule bounded_linear.sums [OF bounded_linear_divide]) lemma summable_divide: "summable f \ summable (\n. f n / c)" by (rule bounded_linear.summable [OF bounded_linear_divide]) lemma suminf_divide: "summable f \ suminf (\n. f n / c) = suminf f / c" by (rule bounded_linear.suminf [OF bounded_linear_divide, symmetric]) lemma summable_inverse_divide: "summable (inverse \ f) \ summable (\n. c / f n)" by (auto dest: summable_mult [of _ c] simp: field_simps) lemma sums_mult_D: "(\n. c * f n) sums a \ c \ 0 \ f sums (a/c)" using sums_mult_iff by fastforce lemma summable_mult_D: "summable (\n. c * f n) \ c \ 0 \ summable f" by (auto dest: summable_divide) text \Sum of a geometric progression.\ lemma geometric_sums: assumes "norm c < 1" shows "(\n. c^n) sums (1 / (1 - c))" proof - have neq_0: "c - 1 \ 0" using assms by auto then have "(\n. c ^ n / (c - 1) - 1 / (c - 1)) \ 0 / (c - 1) - 1 / (c - 1)" by (intro tendsto_intros assms) then have "(\n. (c ^ n - 1) / (c - 1)) \ 1 / (1 - c)" by (simp add: nonzero_minus_divide_right [OF neq_0] diff_divide_distrib) with neq_0 show "(\n. c ^ n) sums (1 / (1 - c))" by (simp add: sums_def geometric_sum) qed lemma summable_geometric: "norm c < 1 \ summable (\n. c^n)" by (rule geometric_sums [THEN sums_summable]) lemma suminf_geometric: "norm c < 1 \ suminf (\n. c^n) = 1 / (1 - c)" by (rule sums_unique[symmetric]) (rule geometric_sums) lemma summable_geometric_iff [simp]: "summable (\n. c ^ n) \ norm c < 1" proof assume "summable (\n. c ^ n :: 'a :: real_normed_field)" then have "(\n. norm c ^ n) \ 0" by (simp add: norm_power [symmetric] tendsto_norm_zero_iff summable_LIMSEQ_zero) from order_tendstoD(2)[OF this zero_less_one] obtain n where "norm c ^ n < 1" by (auto simp: eventually_at_top_linorder) then show "norm c < 1" using one_le_power[of "norm c" n] by (cases "norm c \ 1") (linarith, simp) qed (rule summable_geometric) end text \Biconditional versions for constant factors\ context fixes c :: "'a::real_normed_field" begin lemma summable_cmult_iff [simp]: "summable (\n. c * f n) \ c=0 \ summable f" proof - have "\summable (\n. c * f n); c \ 0\ \ summable f" using summable_mult_D by blast then show ?thesis by (auto simp: summable_mult) qed lemma summable_divide_iff [simp]: "summable (\n. f n / c) \ c=0 \ summable f" proof - have "\summable (\n. f n / c); c \ 0\ \ summable f" by (auto dest: summable_divide [where c = "1/c"]) then show ?thesis by (auto simp: summable_divide) qed end lemma power_half_series: "(\n. (1/2::real)^Suc n) sums 1" proof - have 2: "(\n. (1/2::real)^n) sums 2" using geometric_sums [of "1/2::real"] by auto have "(\n. (1/2::real)^Suc n) = (\n. (1 / 2) ^ n / 2)" by (simp add: mult.commute) then show ?thesis using sums_divide [OF 2, of 2] by simp qed subsection \Telescoping\ lemma telescope_sums: fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "(\n. f (Suc n) - f n) sums (c - f 0)" unfolding sums_def proof (subst filterlim_sequentially_Suc [symmetric]) have "(\n. \kn. f (Suc n) - f 0)" by (simp add: lessThan_Suc_atMost atLeast0AtMost [symmetric] sum_Suc_diff) also have "\ \ c - f 0" by (intro tendsto_diff LIMSEQ_Suc[OF assms] tendsto_const) finally show "(\n. \n c - f 0" . qed lemma telescope_sums': fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "(\n. f n - f (Suc n)) sums (f 0 - c)" using sums_minus[OF telescope_sums[OF assms]] by (simp add: algebra_simps) lemma telescope_summable: fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "summable (\n. f (Suc n) - f n)" using telescope_sums[OF assms] by (simp add: sums_iff) lemma telescope_summable': fixes c :: "'a::real_normed_vector" assumes "f \ c" shows "summable (\n. f n - f (Suc n))" using summable_minus[OF telescope_summable[OF assms]] by (simp add: algebra_simps) subsection \Infinite summability on Banach spaces\ text \Cauchy-type criterion for convergence of series (c.f. Harrison).\ lemma summable_Cauchy: "summable f \ (\e>0. \N. \m\N. \n. norm (sum f {m.. 'a::banach" proof assume f: "summable f" show ?rhs proof clarify fix e :: real assume "0 < e" then obtain M where M: "\m n. \m\M; n\M\ \ norm (sum f {.. M" for m n proof (cases m n rule: linorder_class.le_cases) assume "m \ n" then show ?thesis by (metis (mono_tags, hide_lams) M atLeast0LessThan order_trans sum_diff_nat_ivl that zero_le) next assume "n \ m" then show ?thesis by (simp add: \0 < e\) qed then show "\N. \m\N. \n. norm (sum f {m..m n. m \ N \ norm (sum f {m..N" "n\N" for m n proof (cases m n rule: linorder_class.le_cases) assume "m \ n" then show ?thesis - by (metis Groups_Big.sum_diff N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff norm_minus_commute \m\N\) + by (metis N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff norm_minus_commute sum_diff \m\N\) next assume "n \ m" then show ?thesis - by (metis Groups_Big.sum_diff N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff \n\N\) + by (metis N finite_lessThan lessThan_minus_lessThan lessThan_subset_iff sum_diff \n\N\) qed then show "\M. \m\M. \n\M. norm (sum f {.. 'a :: banach" - assumes "eventually (\m. \n\m. norm (sum f {m.. g m) sequentially" - assumes "filterlim g (nhds 0) sequentially" + assumes ev: "eventually (\m. \n\m. norm (sum f {m.. g m) sequentially" + assumes g0: "g \ 0" shows "summable f" proof (subst summable_Cauchy, intro allI impI, goal_cases) case (1 e) - from order_tendstoD(2)[OF assms(2) this] and assms(1) - have "eventually (\m. \n. norm (sum f {m..\<^sub>F x in sequentially. g x < e" + using g0 order_tendstoD(2) by blast + with ev have "eventually (\m. \n. norm (sum f {m.. m") auto qed qed thus ?case by (auto simp: eventually_at_top_linorder) qed context fixes f :: "nat \ 'a::banach" begin text \Absolute convergence imples normal convergence.\ lemma summable_norm_cancel: "summable (\n. norm (f n)) \ summable f" unfolding summable_Cauchy apply (erule all_forward imp_forward ex_forward | assumption)+ apply (fastforce simp add: order_le_less_trans [OF norm_sum] order_le_less_trans [OF abs_ge_self]) done lemma summable_norm: "summable (\n. norm (f n)) \ norm (suminf f) \ (\n. norm (f n))" by (auto intro: LIMSEQ_le tendsto_norm summable_norm_cancel summable_LIMSEQ norm_sum) text \Comparison tests.\ lemma summable_comparison_test: assumes fg: "\N. \n\N. norm (f n) \ g n" and g: "summable g" shows "summable f" proof - obtain N where N: "\n. n\N \ norm (f n) \ g n" using assms by blast show ?thesis proof (clarsimp simp add: summable_Cauchy) fix e :: real assume "0 < e" then obtain Ng where Ng: "\m n. m \ Ng \ norm (sum g {m..max N Ng" for m n proof - have "norm (sum f {m.. sum g {m.. norm (sum g {m..N. \m\N. \n. norm (sum f {m..n. norm (f n) \ g n) sequentially \ summable g \ summable f" by (rule summable_comparison_test) (auto simp: eventually_at_top_linorder) text \A better argument order.\ lemma summable_comparison_test': "summable g \ (\n. n \ N \ norm (f n) \ g n) \ summable f" by (rule summable_comparison_test) auto subsection \The Ratio Test\ lemma summable_ratio_test: assumes "c < 1" "\n. n \ N \ norm (f (Suc n)) \ c * norm (f n)" shows "summable f" proof (cases "0 < c") case True show "summable f" proof (rule summable_comparison_test) show "\N'. \n\N'. norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" proof (intro exI allI impI) fix n assume "N \ n" then show "norm (f n) \ (norm (f N) / (c ^ N)) * c ^ n" proof (induct rule: inc_induct) case base with True show ?case by simp next case (step m) have "norm (f (Suc m)) / c ^ Suc m * c ^ n \ norm (f m) / c ^ m * c ^ n" using \0 < c\ \c < 1\ assms(2)[OF \N \ m\] by (simp add: field_simps) with step show ?case by simp qed qed show "summable (\n. norm (f N) / c ^ N * c ^ n)" using \0 < c\ \c < 1\ by (intro summable_mult summable_geometric) simp qed next case False have "f (Suc n) = 0" if "n \ N" for n proof - from that have "norm (f (Suc n)) \ c * norm (f n)" by (rule assms(2)) also have "\ \ 0" using False by (simp add: not_less mult_nonpos_nonneg) finally show ?thesis by auto qed then show "summable f" by (intro sums_summable[OF sums_finite, of "{.. Suc N}"]) (auto simp: not_le Suc_less_eq2) qed end text \Relations among convergence and absolute convergence for power series.\ lemma Abel_lemma: fixes a :: "nat \ 'a::real_normed_vector" assumes r: "0 \ r" and r0: "r < r0" and M: "\n. norm (a n) * r0^n \ M" shows "summable (\n. norm (a n) * r^n)" proof (rule summable_comparison_test') show "summable (\n. M * (r / r0) ^ n)" using assms by (auto simp add: summable_mult summable_geometric) show "norm (norm (a n) * r ^ n) \ M * (r / r0) ^ n" for n using r r0 M [of n] dual_order.order_iff_strict by (fastforce simp add: abs_mult field_simps) qed text \Summability of geometric series for real algebras.\ lemma complete_algebra_summable_geometric: fixes x :: "'a::{real_normed_algebra_1,banach}" assumes "norm x < 1" shows "summable (\n. x ^ n)" proof (rule summable_comparison_test) show "\N. \n\N. norm (x ^ n) \ norm x ^ n" by (simp add: norm_power_ineq) from assms show "summable (\n. norm x ^ n)" by (simp add: summable_geometric) qed subsection \Cauchy Product Formula\ text \ Proof based on Analysis WebNotes: Chapter 07, Class 41 \<^url>\http://www.math.unl.edu/~webnotes/classes/class41/prp77.htm\ \ lemma Cauchy_product_sums: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes a: "summable (\k. norm (a k))" and b: "summable (\k. norm (b k))" shows "(\k. \i\k. a i * b (k - i)) sums ((\k. a k) * (\k. b k))" proof - let ?S1 = "\n::nat. {.. {..m n. m \ n \ ?S1 m \ ?S1 n" by auto have S2_le_S1: "\n. ?S2 n \ ?S1 n" by auto have S1_le_S2: "\n. ?S1 (n div 2) \ ?S2 n" by auto have finite_S1: "\n. finite (?S1 n)" by simp with S2_le_S1 have finite_S2: "\n. finite (?S2 n)" by (rule finite_subset) let ?g = "\(i,j). a i * b j" let ?f = "\(i,j). norm (a i) * norm (b j)" have f_nonneg: "\x. 0 \ ?f x" by auto then have norm_sum_f: "\A. norm (sum ?f A) = sum ?f A" unfolding real_norm_def by (simp only: abs_of_nonneg sum_nonneg [rule_format]) have "(\n. (\kk (\k. a k) * (\k. b k)" by (intro tendsto_mult summable_LIMSEQ summable_norm_cancel [OF a] summable_norm_cancel [OF b]) then have 1: "(\n. sum ?g (?S1 n)) \ (\k. a k) * (\k. b k)" by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan) have "(\n. (\kk (\k. norm (a k)) * (\k. norm (b k))" using a b by (intro tendsto_mult summable_LIMSEQ) then have "(\n. sum ?f (?S1 n)) \ (\k. norm (a k)) * (\k. norm (b k))" by (simp only: sum_product sum.Sigma [rule_format] finite_lessThan) then have "convergent (\n. sum ?f (?S1 n))" by (rule convergentI) then have Cauchy: "Cauchy (\n. sum ?f (?S1 n))" by (rule convergent_Cauchy) have "Zfun (\n. sum ?f (?S1 n - ?S2 n)) sequentially" proof (rule ZfunI, simp only: eventually_sequentially norm_sum_f) fix r :: real assume r: "0 < r" from CauchyD [OF Cauchy r] obtain N where "\m\N. \n\N. norm (sum ?f (?S1 m) - sum ?f (?S1 n)) < r" .. then have "\m n. N \ n \ n \ m \ norm (sum ?f (?S1 m - ?S1 n)) < r" by (simp only: sum_diff finite_S1 S1_mono) then have N: "\m n. N \ n \ n \ m \ sum ?f (?S1 m - ?S1 n) < r" by (simp only: norm_sum_f) show "\N. \n\N. sum ?f (?S1 n - ?S2 n) < r" proof (intro exI allI impI) fix n assume "2 * N \ n" then have n: "N \ n div 2" by simp have "sum ?f (?S1 n - ?S2 n) \ sum ?f (?S1 n - ?S1 (n div 2))" by (intro sum_mono2 finite_Diff finite_S1 f_nonneg Diff_mono subset_refl S1_le_S2) also have "\ < r" using n div_le_dividend by (rule N) finally show "sum ?f (?S1 n - ?S2 n) < r" . qed qed then have "Zfun (\n. sum ?g (?S1 n - ?S2 n)) sequentially" apply (rule Zfun_le [rule_format]) apply (simp only: norm_sum_f) apply (rule order_trans [OF norm_sum sum_mono]) apply (auto simp add: norm_mult_ineq) done then have 2: "(\n. sum ?g (?S1 n) - sum ?g (?S2 n)) \ 0" unfolding tendsto_Zfun_iff diff_0_right by (simp only: sum_diff finite_S1 S2_le_S1) with 1 have "(\n. sum ?g (?S2 n)) \ (\k. a k) * (\k. b k)" by (rule Lim_transform2) then show ?thesis by (simp only: sums_def sum.triangle_reindex) qed lemma Cauchy_product: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes "summable (\k. norm (a k))" and "summable (\k. norm (b k))" shows "(\k. a k) * (\k. b k) = (\k. \i\k. a i * b (k - i))" using assms by (rule Cauchy_product_sums [THEN sums_unique]) lemma summable_Cauchy_product: fixes a b :: "nat \ 'a::{real_normed_algebra,banach}" assumes "summable (\k. norm (a k))" and "summable (\k. norm (b k))" shows "summable (\k. \i\k. a i * b (k - i))" using Cauchy_product_sums[OF assms] by (simp add: sums_iff) subsection \Series on \<^typ>\real\s\ lemma summable_norm_comparison_test: "\N. \n\N. norm (f n) \ g n \ summable g \ summable (\n. norm (f n))" by (rule summable_comparison_test) auto lemma summable_rabs_comparison_test: "\N. \n\N. \f n\ \ g n \ summable g \ summable (\n. \f n\)" for f :: "nat \ real" by (rule summable_comparison_test) auto lemma summable_rabs_cancel: "summable (\n. \f n\) \ summable f" for f :: "nat \ real" by (rule summable_norm_cancel) simp lemma summable_rabs: "summable (\n. \f n\) \ \suminf f\ \ (\n. \f n\)" for f :: "nat \ real" by (fold real_norm_def) (rule summable_norm) lemma summable_zero_power [simp]: "summable (\n. 0 ^ n :: 'a::{comm_ring_1,topological_space})" proof - have "(\n. 0 ^ n :: 'a) = (\n. if n = 0 then 0^0 else 0)" by (intro ext) (simp add: zero_power) moreover have "summable \" by simp ultimately show ?thesis by simp qed lemma summable_zero_power' [simp]: "summable (\n. f n * 0 ^ n :: 'a::{ring_1,topological_space})" proof - have "(\n. f n * 0 ^ n :: 'a) = (\n. if n = 0 then f 0 * 0^0 else 0)" by (intro ext) (simp add: zero_power) moreover have "summable \" by simp ultimately show ?thesis by simp qed lemma summable_power_series: fixes z :: real assumes le_1: "\i. f i \ 1" and nonneg: "\i. 0 \ f i" and z: "0 \ z" "z < 1" shows "summable (\i. f i * z^i)" proof (rule summable_comparison_test[OF _ summable_geometric]) show "norm z < 1" using z by (auto simp: less_imp_le) show "\n. \N. \na\N. norm (f na * z ^ na) \ z ^ na" using z by (auto intro!: exI[of _ 0] mult_left_le_one_le simp: abs_mult nonneg power_abs less_imp_le le_1) qed lemma summable_0_powser: "summable (\n. f n * 0 ^ n :: 'a::real_normed_div_algebra)" proof - have A: "(\n. f n * 0 ^ n) = (\n. if n = 0 then f n else 0)" by (intro ext) auto then show ?thesis by (subst A) simp_all qed lemma summable_powser_split_head: "summable (\n. f (Suc n) * z ^ n :: 'a::real_normed_div_algebra) = summable (\n. f n * z ^ n)" proof - have "summable (\n. f (Suc n) * z ^ n) \ summable (\n. f (Suc n) * z ^ Suc n)" (is "?lhs \ ?rhs") proof show ?rhs if ?lhs using summable_mult2[OF that, of z] by (simp add: power_commutes algebra_simps) show ?lhs if ?rhs using summable_mult2[OF that, of "inverse z"] by (cases "z \ 0", subst (asm) power_Suc2) (simp_all add: algebra_simps) qed also have "\ \ summable (\n. f n * z ^ n)" by (rule summable_Suc_iff) finally show ?thesis . qed lemma summable_powser_ignore_initial_segment: fixes f :: "nat \ 'a :: real_normed_div_algebra" shows "summable (\n. f (n + m) * z ^ n) \ summable (\n. f n * z ^ n)" proof (induction m) case (Suc m) have "summable (\n. f (n + Suc m) * z ^ n) = summable (\n. f (Suc n + m) * z ^ n)" by simp also have "\ = summable (\n. f (n + m) * z ^ n)" by (rule summable_powser_split_head) also have "\ = summable (\n. f n * z ^ n)" by (rule Suc.IH) finally show ?case . qed simp_all lemma powser_split_head: fixes f :: "nat \ 'a::{real_normed_div_algebra,banach}" assumes "summable (\n. f n * z ^ n)" shows "suminf (\n. f n * z ^ n) = f 0 + suminf (\n. f (Suc n) * z ^ n) * z" and "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" and "summable (\n. f (Suc n) * z ^ n)" proof - from assms show "summable (\n. f (Suc n) * z ^ n)" by (subst summable_powser_split_head) from suminf_mult2[OF this, of z] have "(\n. f (Suc n) * z ^ n) * z = (\n. f (Suc n) * z ^ Suc n)" by (simp add: power_commutes algebra_simps) also from assms have "\ = suminf (\n. f n * z ^ n) - f 0" by (subst suminf_split_head) simp_all finally show "suminf (\n. f n * z ^ n) = f 0 + suminf (\n. f (Suc n) * z ^ n) * z" by simp then show "suminf (\n. f (Suc n) * z ^ n) * z = suminf (\n. f n * z ^ n) - f 0" by simp qed lemma summable_partial_sum_bound: fixes f :: "nat \ 'a :: banach" and e :: real assumes summable: "summable f" and e: "e > 0" obtains N where "\m n. m \ N \ norm (\k=m..n. f k) < e" proof - from summable have "Cauchy (\n. \km n. m \ N \ n \ N \ norm ((\kkk=m..n. f k) < e" if m: "m \ N" for m n proof (cases "n \ m") case True with m have "norm ((\kkkkk=m..n. f k)" by (subst sum_diff [symmetric]) (simp_all add: sum.last_plus) finally show ?thesis . next case False with e show ?thesis by simp_all qed then show ?thesis by (rule that) qed lemma powser_sums_if: "(\n. (if n = m then (1 :: 'a::{ring_1,topological_space}) else 0) * z^n) sums z^m" proof - have "(\n. (if n = m then 1 else 0) * z^n) = (\n. if n = m then z^n else 0)" by (intro ext) auto then show ?thesis by (simp add: sums_single) qed lemma fixes f :: "nat \ real" assumes "summable f" and "inj g" and pos: "\x. 0 \ f x" shows summable_reindex: "summable (f \ g)" and suminf_reindex_mono: "suminf (f \ g) \ suminf f" and suminf_reindex: "(\x. x \ range g \ f x = 0) \ suminf (f \ g) = suminf f" proof - from \inj g\ have [simp]: "\A. inj_on g A" by (rule subset_inj_on) simp have smaller: "\n. (\i g) i) \ suminf f" proof fix n have "\ n' \ (g ` {..n'. n' < n \ g n' < m" by blast have "(\i \ (\i \ suminf f" using \summable f\ by (rule sum_le_suminf) (simp_all add: pos) finally show "(\i g) i) \ suminf f" by simp qed have "incseq (\n. \i g) i)" by (rule incseq_SucI) (auto simp add: pos) then obtain L where L: "(\ n. \i g) i) \ L" using smaller by(rule incseq_convergent) then have "(f \ g) sums L" by (simp add: sums_def) then show "summable (f \ g)" by (auto simp add: sums_iff) then have "(\n. \i g) i) \ suminf (f \ g)" by (rule summable_LIMSEQ) then show le: "suminf (f \ g) \ suminf f" by(rule LIMSEQ_le_const2)(blast intro: smaller[rule_format]) assume f: "\x. x \ range g \ f x = 0" from \summable f\ have "suminf f \ suminf (f \ g)" proof (rule suminf_le_const) fix n have "\ n' \ (g -` {..n'. g n' < n \ n' < m" by blast have "(\ii\{.. range g. f i)" using f by(auto intro: sum.mono_neutral_cong_right) also have "\ = (\i\g -` {.. g) i)" by (rule sum.reindex_cong[where l=g])(auto) also have "\ \ (\i g) i)" by (rule sum_mono2)(auto simp add: pos n) also have "\ \ suminf (f \ g)" using \summable (f \ g)\ by (rule sum_le_suminf) (simp_all add: pos) finally show "sum f {.. suminf (f \ g)" . qed with le show "suminf (f \ g) = suminf f" by (rule antisym) qed lemma sums_mono_reindex: assumes subseq: "strict_mono g" and zero: "\n. n \ range g \ f n = 0" shows "(\n. f (g n)) sums c \ f sums c" unfolding sums_def proof assume lim: "(\n. \k c" have "(\n. \kn. \kkk\g`{.. = (\kkk \ c" by (simp only: o_def) finally show "(\n. \k c" . next assume lim: "(\n. \k c" define g_inv where "g_inv n = (LEAST m. g m \ n)" for n from filterlim_subseq[OF subseq] have g_inv_ex: "\m. g m \ n" for n by (auto simp: filterlim_at_top eventually_at_top_linorder) then have g_inv: "g (g_inv n) \ n" for n unfolding g_inv_def by (rule LeastI_ex) have g_inv_least: "m \ g_inv n" if "g m \ n" for m n using that unfolding g_inv_def by (rule Least_le) have g_inv_least': "g m < n" if "m < g_inv n" for m n using that g_inv_least[of n m] by linarith have "(\n. \kn. \k {.. range g" proof (rule notI, elim imageE) fix l assume l: "k = g l" have "g l < g (g_inv n)" by (rule less_le_trans[OF _ g_inv]) (use k l in simp_all) with subseq have "l < g_inv n" by (simp add: strict_mono_less) with k l show False by simp qed then have "f k = 0" by (rule zero) } with g_inv_least' g_inv have "(\kk\g`{.. = (\kkk n" also have "n \ g (g_inv n)" by (rule g_inv) finally have "K \ g_inv n" using subseq by (simp add: strict_mono_less_eq) } then have "filterlim g_inv at_top sequentially" by (auto simp: filterlim_at_top eventually_at_top_linorder) with lim have "(\n. \k c" by (rule filterlim_compose) finally show "(\n. \k c" . qed lemma summable_mono_reindex: assumes subseq: "strict_mono g" and zero: "\n. n \ range g \ f n = 0" shows "summable (\n. f (g n)) \ summable f" using sums_mono_reindex[of g f, OF assms] by (simp add: summable_def) lemma suminf_mono_reindex: fixes f :: "nat \ 'a::{t2_space,comm_monoid_add}" assumes "strict_mono g" "\n. n \ range g \ f n = 0" shows "suminf (\n. f (g n)) = suminf f" proof (cases "summable f") case True with sums_mono_reindex [of g f, OF assms] and summable_mono_reindex [of g f, OF assms] show ?thesis by (simp add: sums_iff) next case False then have "\(\c. f sums c)" unfolding summable_def by blast then have "suminf f = The (\_. False)" by (simp add: suminf_def) moreover from False have "\ summable (\n. f (g n))" using summable_mono_reindex[of g f, OF assms] by simp then have "\(\c. (\n. f (g n)) sums c)" unfolding summable_def by blast then have "suminf (\n. f (g n)) = The (\_. False)" by (simp add: suminf_def) ultimately show ?thesis by simp qed lemma summable_bounded_partials: fixes f :: "nat \ 'a :: {real_normed_vector,complete_space}" assumes bound: "eventually (\x0. \a\x0. \b>a. norm (sum f {a<..b}) \ g a) sequentially" assumes g: "g \ 0" shows "summable f" unfolding summable_iff_convergent' proof (intro Cauchy_convergent CauchyI', goal_cases) case (1 \) with g have "eventually (\x. \g x\ < \) sequentially" by (auto simp: tendsto_iff) from eventually_conj[OF this bound] obtain x0 where x0: "\x. x \ x0 \ \g x\ < \" "\a b. x0 \ a \ a < b \ norm (sum f {a<..b}) \ g a" unfolding eventually_at_top_linorder by auto show ?case proof (intro exI[of _ x0] allI impI) fix m n assume mn: "x0 \ m" "m < n" have "dist (sum f {..m}) (sum f {..n}) = norm (sum f {..n} - sum f {..m})" by (simp add: dist_norm norm_minus_commute) also have "sum f {..n} - sum f {..m} = sum f ({..n} - {..m})" using mn by (intro Groups_Big.sum_diff [symmetric]) auto also have "{..n} - {..m} = {m<..n}" using mn by auto also have "norm (sum f {m<..n}) \ g m" using mn by (intro x0) auto also have "\ \ \g m\" by simp also have "\ < \" using mn by (intro x0) auto finally show "dist (sum f {..m}) (sum f {..n}) < \" . qed qed end