diff --git a/src/HOL/Algebra/Algebra.thy b/src/HOL/Algebra/Algebra.thy --- a/src/HOL/Algebra/Algebra.thy +++ b/src/HOL/Algebra/Algebra.thy @@ -1,7 +1,8 @@ (* Title: HOL/Algebra/Algebra.thy *) theory Algebra imports Sylow Chinese_Remainder Zassenhaus Galois_Connection Generated_Fields Free_Abelian_Groups Divisibility Embedded_Algebras IntRing Sym_Groups Exact_Sequence Polynomials Algebraic_Closure + Left_Coset begin end diff --git a/src/HOL/Algebra/Left_Coset.thy b/src/HOL/Algebra/Left_Coset.thy new file mode 100644 --- /dev/null +++ b/src/HOL/Algebra/Left_Coset.thy @@ -0,0 +1,137 @@ +theory Left_Coset +imports Coset + +(*This instance of Coset.thy but for left cosets is due to Jonas Rädle and has been imported + from the AFP entry Orbit_Stabiliser. *) + +begin + +definition + LCOSETS :: "[_, 'a set] \ ('a set)set" ("lcosets\ _" [81] 80) + where "lcosets\<^bsub>G\<^esub> H = (\a\carrier G. {a <#\<^bsub>G\<^esub> H})" + +definition + LFactGroup :: "[('a,'b) monoid_scheme, 'a set] \ ('a set) monoid" (infixl "LMod" 65) + \ \Actually defined for groups rather than monoids\ + where "LFactGroup G H = \carrier = lcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\" + +lemma (in group) lcos_self: "[| x \ carrier G; subgroup H G |] ==> x \ x <# H" + by (simp add: group_l_invI subgroup.lcos_module_rev subgroup.one_closed) + +text \Elements of a left coset are in the carrier\ +lemma (in subgroup) elemlcos_carrier: + assumes "group G" "a \ carrier G" "a' \ a <# H" + shows "a' \ carrier G" + by (meson assms group.l_coset_carrier subgroup_axioms) + +text \Step one for lemma \rcos_module\\ +lemma (in subgroup) lcos_module_imp: + assumes "group G" + assumes xcarr: "x \ carrier G" + and x'cos: "x' \ x <# H" + shows "(inv x \ x') \ H" +proof - + interpret group G by fact + obtain h + where hH: "h \ H" and x': "x' = x \ h" and hcarr: "h \ carrier G" + using assms by (auto simp: l_coset_def) + have "(inv x) \ x' = (inv x) \ (x \ h)" + by (simp add: x') + have "\ = (x \ inv x) \ h" + by (metis hcarr inv_closed inv_inv l_inv m_assoc xcarr) + also have "\ = h" + by (simp add: hcarr xcarr) + finally have "(inv x) \ x' = h" + using x' by metis + then show "(inv x) \ x' \ H" + using hH by force +qed + +text \Left cosets are subsets of the carrier.\ +lemma (in subgroup) lcosets_carrier: + assumes "group G" + assumes XH: "X \ lcosets H" + shows "X \ carrier G" +proof - + interpret group G by fact + show "X \ carrier G" + using XH l_coset_subset_G subset by (auto simp: LCOSETS_def) +qed + +lemma (in group) lcosets_part_G: + assumes "subgroup H G" + shows "\(lcosets H) = carrier G" +proof - + interpret subgroup H G by fact + show ?thesis + proof + show "\ (lcosets H) \ carrier G" + by (force simp add: LCOSETS_def l_coset_def) + show "carrier G \ \ (lcosets H)" + by (auto simp add: LCOSETS_def intro: lcos_self assms) + qed +qed + +lemma (in group) lcosets_subset_PowG: + "subgroup H G \ lcosets H \ Pow(carrier G)" + using lcosets_part_G subset_Pow_Union by blast + +lemma (in group) lcos_disjoint: + assumes "subgroup H G" + assumes p: "a \ lcosets H" "b \ lcosets H" "a\b" + shows "a \ b = {}" +proof - + interpret subgroup H G by fact + show ?thesis + using p l_repr_independence subgroup_axioms unfolding LCOSETS_def disjoint_iff + by blast +qed + +text\The next two lemmas support the proof of \card_cosets_equal\.\ +lemma (in group) inj_on_f': + "\H \ carrier G; a \ carrier G\ \ inj_on (\y. y \ inv a) (a <# H)" + by (simp add: inj_on_g l_coset_subset_G) + +lemma (in group) inj_on_f'': + "\H \ carrier G; a \ carrier G\ \ inj_on (\y. inv a \ y) (a <# H)" + by (meson inj_on_cmult inv_closed l_coset_subset_G subset_inj_on) + +lemma (in group) inj_on_g': + "\H \ carrier G; a \ carrier G\ \ inj_on (\y. a \ y) H" + using inj_on_cmult inj_on_subset by blast + +lemma (in group) l_card_cosets_equal: + assumes "c \ lcosets H" and H: "H \ carrier G" and fin: "finite(carrier G)" + shows "card H = card c" +proof - + obtain x where x: "x \ carrier G" and c: "c = x <# H" + using assms by (auto simp add: LCOSETS_def) + have "inj_on ((\) x) H" + by (simp add: H group.inj_on_g' x) + moreover + have "(\) x ` H \ x <# H" + by (force simp add: m_assoc subsetD l_coset_def) + moreover + have "inj_on ((\) (inv x)) (x <# H)" + by (simp add: H group.inj_on_f'' x) + moreover + have "\h. h \ H \ inv x \ (x \ h) \ H" + by (metis H in_mono inv_solve_left m_closed x) + then have "(\) (inv x) ` (x <# H) \ H" + by (auto simp: l_coset_def) + ultimately show ?thesis + by (metis H fin c card_bij_eq finite_imageD finite_subset) +qed + +theorem (in group) l_lagrange: + assumes "finite(carrier G)" "subgroup H G" + shows "card(lcosets H) * card(H) = order(G)" +proof - + have "card H * card (lcosets H) = card (\ (lcosets H))" + using card_partition + by (metis (no_types, lifting) assms finite_UnionD l_card_cosets_equal lcos_disjoint lcosets_part_G subgroup.subset) + then show ?thesis + by (simp add: assms(2) lcosets_part_G mult.commute order_def) +qed + +end 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,4099 +1,4190 @@ 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)" (is "?L=?R") proof assume ?L then have "cos (y-x) = 1" using cos_add [of y "-x"] by simp then show ?R by (metis cos_one_2pi_int add.commute diff_add_cancel mult.assoc mult.commute) next assume ?R then show ?L by (auto simp: sin_add cos_add) 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 = complex_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 = complex_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 = complex_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 = complex_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 add: 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 add: 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 cos_gt_neg1: assumes "(t::real) \ {-pi<.. -1" proof - have "cos t \ -1" by simp moreover have "cos t \ -1" proof assume "cos t = -1" then obtain n where n: "t = (2 * of_int n + 1) * pi" by (subst (asm) cos_eq_minus1) auto from assms have "t / pi \ {-1<..<1}" by (auto simp: field_simps) also from n have "t / pi = of_int (2 * n + 1)" by auto finally show False by auto qed ultimately show ?thesis by linarith 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" using cos_double_sin [of "Re z"] apply (simp add: sin_cos_eq norm_cos_squared exp_minus mult.commute [of _ 2] exp_double) apply (simp add: algebra_simps power2_eq_square) done 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" by (metis (no_types, lifting) Arg2pi Arg2pi_eq Arg2pi_unique assms mult.assoc mult_eq_0_iff mult_pos_pos of_real_mult zero_less_norm_iff) 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_irrefl 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)" (is "?lhs = ?rhs") proof assume ?lhs then have "(cmod w) * (z / cmod z) = w" by (metis Arg2pi_eq assms(2) mult_eq_0_iff nonzero_mult_div_cancel_left) then show ?rhs by (metis assms divide_pos_pos of_real_divide times_divide_eq_left times_divide_eq_right zero_less_norm_iff) qed auto 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"] Arg2pi [of "w * z"] by auto 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: assumes "z \ 0" shows "Arg2pi (cnj z) = Arg2pi (inverse z)" proof - have "\r>0. of_real r / z = cnj z" by (metis assms complex_norm_square nonzero_mult_div_cancel_left zero_less_norm_iff zero_less_power) then show ?thesis by (metis Arg2pi_times_of_real2 divide_inverse_commute) qed lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \ \ then Arg2pi z else 2*pi - Arg2pi z)" by (metis Arg2pi_cnj_eq_inverse Arg2pi_inverse Reals_cnj_iff complex_cnj_zero) lemma Arg2pi_exp: "0 \ Im z \ Im z < 2*pi \ Arg2pi(exp z) = Im z" by (simp add: Arg2pi_unique 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, opaque_lifting) 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, opaque_lifting) 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]: "S \ \\<^sub>\\<^sub>0 = {} \ Ln holomorphic_on S" by (simp add: disjoint_iff 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 lemma norm_Ln_le: fixes z :: complex assumes "norm z < 1/2" shows "norm (Ln(1+z)) \ 2 * norm z" proof - have sums: "(\n. - ((- z) ^ n) / of_nat n) sums ln (1 + z)" by (intro Ln_series') (use assms in auto) have summable: "summable (\n. norm (- ((- z) ^ n / of_nat n)))" using ln_series'[of "-norm z"] assms by (simp add: sums_iff summable_minus_iff norm_power norm_divide) have "norm (ln (1 + z)) = norm (\n. -((-z) ^ n / of_nat n))" using sums by (simp add: sums_iff) also have "\ \ (\n. norm (-((-z) ^ n / of_nat n)))" using summable by (rule summable_norm) also have "\ = (\n. norm (-((-z) ^ Suc n / of_nat (Suc n))))" using summable by (subst suminf_split_head) auto also have "\ \ (\n. norm z * (1 / 2) ^ n)" proof (rule suminf_le) show "summable (\n. norm z * (1 / 2) ^ n)" by (intro summable_mult summable_geometric) auto next show "summable (\n. norm (- ((- z) ^ Suc n / of_nat (Suc n))))" using summable by (subst summable_Suc_iff) next fix n have "norm (-((-z) ^ Suc n / of_nat (Suc n))) = norm z * (norm z ^ n / real (Suc n))" by (simp add: norm_power norm_divide norm_mult del: of_nat_Suc) also have "\ \ norm z * ((1 / 2) ^ n / 1)" using assms by (intro mult_left_mono frac_le power_mono) auto finally show "norm (- ((- z) ^ Suc n / of_nat (Suc n))) \ norm z * (1 / 2) ^ n" by simp qed also have "\ = norm z * (\n. (1 / 2) ^ n)" by (subst suminf_mult) (auto intro: summable_geometric) also have "(\n. (1 / 2 :: real) ^ n) = 2" using geometric_sums[of "1 / 2 :: real"] by (simp add: sums_iff) finally show ?thesis by (simp add: mult_ac) 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)" using Im_Ln_pos_le Im_Ln_pos_lt Re_Ln_pos_lt by fastforce lemma Im_Ln_eq_pi: "z \ 0 \ (Im(Ln z) = pi \ Re(z) < 0 \ Im(z) = 0)" using Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt complex.expand by fastforce 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_of_nat: "\r > 0; z \ 0\ \ Ln(of_nat r * z :: complex) = ln (of_nat r) + Ln(z)" using Ln_times_of_real[of "of_nat r" z] by simp 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 lemma norm_Ln_times_le: assumes "w \ 0" "z \ 0" shows "cmod (Ln(w * z)) \ cmod (Ln(w) + Ln(z))" proof (cases "- pi < Im(Ln w + Ln z) \ Im(Ln w + Ln z) \ pi") case True then show ?thesis by (simp add: Ln_times_simple assms) next case False then show ?thesis by (smt (verit) Im_Ln_le_pi assms cmod_Im_le_iff exp_Ln exp_add ln_unique mpi_less_Im_Ln mult_eq_0_iff norm_exp_eq_Re) qed corollary norm_Ln_prod_le: fixes f :: "'a \ complex" assumes "\x. x \ A \ f x \ 0" shows "cmod (Ln (prod f A)) \ (\x \ A. cmod (Ln (f x)))" using assms proof (induction A rule: infinite_finite_induct) case (insert x A) then show ?case by simp (smt (verit) norm_Ln_times_le norm_triangle_ineq prod_zero_iff) qed auto lemma norm_Ln_exp_le: "norm (Ln (exp z)) \ norm z" by (smt (verit) Im_Ln_le_pi Ln_exp Re_Ln cmod_Im_le_iff exp_not_eq_zero ln_exp mpi_less_Im_Ln norm_exp_eq_Re) subsection\<^marker>\tag unimportant\\Uniform convergence and products\ (* TODO: could be generalised perhaps, but then one would have to do without the ln *) lemma uniformly_convergent_on_prod_aux: fixes f :: "nat \ complex \ complex" assumes norm_f: "\n x. x \ A \ norm (f n x) < 1" assumes cont: "\n. continuous_on A (f n)" assumes conv: "uniformly_convergent_on A (\N x. \nN x. \nN x. \n \\<^sub>\\<^sub>0" if "x \ A" for n x proof assume "f n x + 1 \ \\<^sub>\\<^sub>0" then obtain t where t: "t \ 0" "f n x + 1 = of_real t" by (auto elim!: nonpos_Reals_cases) hence "f n x = of_real (t - 1)" by (simp add: algebra_simps) also have "norm \ \ 1" using t by (subst norm_of_real) auto finally show False using norm_f[of x n] that by auto qed thus "\\<^sub>F n in sequentially. continuous_on A (\x. \n S ` A \ y \ cball 0 1}" have "compact B" unfolding B_def by (intro compact_sums compact_continuous_image cont' A) auto have "uniformly_convergent_on A (\N x. exp ((\ncompact B\ by (auto dest: compact_imp_closed) show "uniformly_continuous_on B exp" by (intro compact_uniformly_continuous continuous_intros \compact B\) have "eventually (\N. \x\A. dist (\nN. \x\A. (\n B) sequentially" proof eventually_elim case (elim N) show "\x\A. (\n B" proof safe fix x assume x: "x \ A" have "(\nnn ball 0 1" using elim x by (auto simp: dist_norm norm_minus_commute) ultimately show "(\n B" unfolding B_def using x by fastforce qed qed qed also have "?this \ uniformly_convergent_on A (\N x. \n A" have "exp (\nn = (\n {.. -1" using norm_f[of x n] x by auto thus "exp (ln (1 + f n x)) = 1 + f n x" by (simp add: add_eq_0_iff) qed auto finally show "exp (\nn complex \ complex" assumes cont: "\n. continuous_on A (f n)" assumes A: "compact A" assumes conv_sum: "uniformly_convergent_on A (\N x. \nN x. \nn x. n \ M \ x \ A \ norm (f n x) < 1 / 2" proof - from conv_sum have "uniformly_Cauchy_on A (\N x. \n 0" by simp ultimately obtain M where M: "\x m n. x \ A \ m \ M \ n \ M \ dist (\kk M" "x \ A" have "dist (\kkkkN x. \n 2 * norm (f (n + M) x)" if "x \ A" for n x by (rule norm_Ln_le) (use M[of "n + M" x] that in auto) have *: "filterlim (\n. n + M) at_top at_top" by (rule filterlim_add_const_nat_at_top) have "uniformly_convergent_on A (\N x. 2 * ((\nnN x. 2 * ((\nnN x. \nnnn\{.. = (\nn. n + M" "\n. n - M"]) auto finally show "?lhs N x = ?rhs N x" by (simp add: sum_distrib_left) qed finally show "uniformly_convergent_on A (\N x. \nN x. \n A" for n x using M[of "n + M" x] that by simp qed (use M assms conv in auto) then obtain S where S: "uniform_limit A (\N x. \nN x. (\nnx. (\nx. \nN x. (\nnN x. (\nnN x. (\nnn\{M..n. n - M" "\n. n + M"]) auto also have "(\n = (\n\{..{M.. {M..nnnN x. \nN x. \n ?thesis" proof (rule uniformly_convergent_cong) show "eventually (\x. \y\A. (\nn complex \ complex" assumes cont: "\n. continuous_on A (f n)" assumes A: "compact A" assumes conv_sum: "uniformly_convergent_on A (\N x. \nN x. \nN x. \nauto intro!: continuous_intros\) thus ?thesis by simp qed text\Prop 17.6 of Bak and Newman, Complex Analysis, p. 243. Only this version is for the reals. Can the two proofs be consolidated?\ lemma uniformly_convergent_on_prod_real: fixes u :: "nat \ real \ real" assumes contu: "\k. continuous_on K (u k)" and uconv: "uniformly_convergent_on K (\n x. \ku k x\)" and K: "compact K" shows "uniformly_convergent_on K (\n x. \k \k. complex_of_real \ u k \ Re" define L where "L \ complex_of_real ` K" have "compact L" by (simp add: \compact K\ L_def compact_continuous_image) have "Re ` complex_of_real ` X = X" for X by (auto simp: image_iff) with contu have contf: "\k. continuous_on L (f k)" unfolding f_def L_def by (intro continuous_intros) auto obtain S where S: "\\. \>0 \ \\<^sub>F n in sequentially. \x\K. dist (\ku k x\) (S x) < \" using uconv unfolding uniformly_convergent_on_def uniform_limit_iff by presburger have "\\<^sub>F n in sequentially. \z\L. dist (\k S \ Re) z) < \" if "\>0" for \ using S [OF that] by eventually_elim (simp add: L_def f_def) then have uconvf: "uniformly_convergent_on L (\n z. \k\. \>0 \ \\<^sub>F n in sequentially. \z\L. dist (\k" using uniformly_convergent_on_prod [OF contf \compact L\ uconvf] unfolding uniformly_convergent_on_def uniform_limit_iff by blast have \
: "\(\k \ cmod ((\kk\N. of_real (1 + u k x)) = (\k\N. 1 + of_real (u k x))" for N by force then show ?thesis by (metis Re_complex_of_real abs_Re_le_cmod minus_complex.sel(1) of_real_prod) qed have "\\<^sub>F n in sequentially. \x\K. dist (\k P \ of_real) x) < \" if "\>0" for \ using P [OF that] by eventually_elim (simp add: L_def f_def dist_norm le_less_trans [OF \
]) then show ?thesis unfolding uniformly_convergent_on_def uniform_limit_iff by blast 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)" proof (rule cis_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\<^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 [simp]: "Arg (of_real r) = (if r<0 then pi else 0)" by (simp add: Im_Ln_eq_pi Arg_def) lemma mpi_less_Arg: "-pi < Arg z" and Arg_le_pi: "Arg z \ 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" using Arg_def Ln_times_of_real assms by auto 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_irrefl 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 Arg_def Im_Ln_eq_0 complex_eq_iff complex_is_Real_iff by auto 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" using Arg_eq_pi complex_is_Real_iff by blast 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 (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse) next case False then show ?thesis by (simp add: Arg_def Ln_inverse complex_is_Real_iff complex_nonpos_Reals_iff) qed 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 = (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 "(\z. Im (Ln z)) \z\ Arg z" using Arg_def assms continuous_at by fastforce then show ?thesis unfolding continuous_at by (smt (verit, del_insts) Arg_eq_Im_Ln Lim_transform_away_at assms nonpos_Reals_zero_I) 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) lemma arg_conv_arctan: assumes "Re z > 0" shows "Arg z = arctan (Im z / Re z)" proof (rule cis_Arg_unique) show "sgn z = cis (arctan (Im z / Re z))" proof (rule complex_eqI) have "Re (cis (arctan (Im z / Re z))) = 1 / sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2)" by (simp add: cos_arctan power_divide) also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2" using assms by (simp add: cmod_def field_simps) also have "1 / sqrt \ = Re z / norm z" using assms by (simp add: real_sqrt_divide) finally show "Re (sgn z) = Re (cis (arctan (Im z / Re z)))" by simp next have "Im (cis (arctan (Im z / Re z))) = Im z / (Re z * sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2))" by (simp add: sin_arctan field_simps) also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2" using assms by (simp add: cmod_def field_simps) also have "Im z / (Re z * sqrt \) = Im z / norm z" using assms by (simp add: real_sqrt_divide) finally show "Im (sgn z) = Im (cis (arctan (Im z / Re z)))" by simp qed next show "arctan (Im z / Re z) > -pi" by (smt (verit, ccfv_SIG) arctan field_sum_of_halves) next show "arctan (Im z / Re z) \ pi" by (smt (verit, best) arctan field_sum_of_halves) qed + +subsection \Characterisation of @{term "Im (Ln z)"} (Wenda Li)\ + +lemma Im_Ln_eq_pi_half: + "z \ 0 \ (Im(Ln z) = pi/2 \ 0 < Im(z) \ Re(z) = 0)" + "z \ 0 \ (Im(Ln z) = -pi/2 \ Im(z) < 0 \ Re(z) = 0)" +proof - + show "z \ 0 \ (Im(Ln z) = pi/2 \ 0 < Im(z) \ Re(z) = 0)" + by (metis Im_Ln_eq_pi Im_Ln_le_pi Im_Ln_pos_lt Re_Ln_pos_le Re_Ln_pos_lt + abs_of_nonneg less_eq_real_def order_less_irrefl pi_half_gt_zero) +next + assume "z\0" + have "Im (Ln z) = - pi / 2 \ Im z < 0 \ Re z = 0" + by (metis Im_Ln_pos_le Re_Ln_pos_le Re_Ln_pos_lt_imp \z \ 0\ abs_if + add.inverse_inverse divide_minus_left less_eq_real_def linorder_not_le minus_pi_half_less_zero) + moreover have "Im (Ln z) = - pi / 2" when "Im z < 0" "Re z = 0" + proof - + obtain r::real where "r>0" "z=r * (-\)" + by (smt (verit) \Im z < 0\ \Re z = 0\ add_0 complex_eq mult.commute mult_minus_right of_real_0 of_real_minus) + then have "Im (Ln z) = Im (Ln (r*(-\)))" by auto + also have "... = Im (Ln (complex_of_real r) + Ln (- \))" + by (metis Ln_times_of_real \0 < r\ add.inverse_inverse add.inverse_neutral complex_i_not_zero) + also have "... = - pi/2" + using \r>0\ by simp + finally show "Im (Ln z) = - pi / 2" . + qed + ultimately show "(Im(Ln z) = -pi/2 \ Im(z) < 0 \ Re(z) = 0)" by auto +qed + +lemma Im_Ln_eq: + assumes "z\0" + shows "Im (Ln z) = (if Re z\0 then + if Re z>0 then + arctan (Im z/Re z) + else if Im z\0 then + arctan (Im z/Re z) + pi + else + arctan (Im z/Re z) - pi + else + if Im z>0 then pi/2 else -pi/2)" +proof - + have eq_arctan_pos:"Im (Ln z) = arctan (Im z/Re z)" when "Re z>0" for z + proof - + define wR where "wR \ Re (Ln z)" + define \ where "\ \ arctan (Im z/Re z)" + have "z\0" using that by auto + have "exp (Complex wR \) = z" + proof (rule complex_eqI) + have "Im (exp (Complex wR \)) =exp wR * sin \ " + unfolding Im_exp by simp + also have "... = Im z" + unfolding wR_def Re_Ln[OF \z\0\] \_def using \z\0\ \Re z>0\ + by (auto simp add:sin_arctan divide_simps complex_neq_0 cmod_def real_sqrt_divide) + finally show "Im (exp (Complex wR \)) = Im z" . + next + have "Re (exp (Complex wR \)) = exp wR * cos \ " + unfolding Re_exp by simp + also have "... = Re z" + by (metis Arg_eq_Im_Ln Re_exp \z \ 0\ \_def arg_conv_arctan exp_Ln that wR_def) + finally show "Re (exp (Complex wR \)) = Re z" . + qed + moreover have "-pi<\" "\\pi" + using arctan_lbound [of \Im z / Re z\] arctan_ubound [of \Im z / Re z\] + by (simp_all add: \_def) + ultimately have "Ln z = Complex wR \" using Ln_unique by auto + then show ?thesis using that unfolding \_def by auto + qed + have ?thesis when "Re z=0" + using Im_Ln_eq_pi_half[OF \z\0\] that + using assms complex_eq_iff by auto + moreover have ?thesis when "Re z>0" + using eq_arctan_pos[OF that] that by auto + moreover have ?thesis when "Re z<0" "Im z\0" + proof - + have "Im (Ln (- z)) = arctan (Im (- z) / Re (- z))" + by (simp add: eq_arctan_pos that(1)) + moreover have "Ln (- z) = Ln z - \ * complex_of_real pi" + using Ln_minus assms that by fastforce + ultimately show ?thesis using that by auto + qed + moreover have ?thesis when "Re z<0" "Im z<0" + proof - + have "Im (Ln (- z)) = arctan (Im (- z) / Re (- z))" + by (simp add: eq_arctan_pos that(1)) + moreover have "Ln (- z) = Ln z + \ * complex_of_real pi" + using Ln_minus assms that by auto + ultimately show ?thesis using that by auto + qed + ultimately show ?thesis by linarith +qed + subsection\<^marker>\tag unimportant\\Relation between Ln and Arg2pi, and hence continuity of Arg2pi\ lemma Arg2pi_Ln: "0 < Arg2pi z \ Arg2pi z = Im(Ln(-z)) + pi" by (smt (verit, best) Arg2pi_0 Arg2pi_exp Arg2pi_minus Arg_exp Arg_minus Im_Ln_le_pi exp_Ln mpi_less_Im_Ln neg_equal_0_iff_equal) 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)+ consider "Re z < 0" | "Im z \ 0" using assms using complex_nonneg_Reals_iff not_le by blast then have "(\z. Im (Ln (- z)) + pi) \z\ Arg2pi z" using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within) then show ?thesis unfolding continuous_at by (metis (mono_tags, lifting) Arg2pi_Ln Arg2pi_gt_0 Compl_iff Lim_transform_within_open assms closed_nonneg_Reals_complex open_Compl) 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)" by (smt (verit, ccfv_SIG) Arg2pi_exp Im_Ln_pos_le assms exp_Ln pi_neq_zero zero_complex.simps(1)) 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 powr (numeral n) = x ^ (numeral n)" by (metis of_nat_numeral power_zero_numeral powr_nat) 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 by (metis e_dist \e>0\ dist_commute powr_of_int that) 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 by (metis \e>0\ e_dist dist_commute linorder_not_le powr_of_int that) 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 holomorphic_on S" "g holomorphic_on S" and "\z z'. \z \ S; z' \ S\ \ g z = 0 \ g z' = 0" shows "(\z. f z / g z) holomorphic_on S" by (metis (no_types, lifting) assms division_ring_divide_zero holomorphic_on_divide holomorphic_transform) 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" using tendsto_neg_powr_complex_of_real [of "real o f" F s] 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: "z \ \\<^sub>\\<^sub>0 \ csqrt (inverse z) = inverse (csqrt z)" by (metis Ln_inverse csqrt_eq_0 csqrt_exp_Ln divide_minus_left exp_minus inverse_nonzero_iff_nonzero) lemma cnj_csqrt: "z \ \\<^sub>\\<^sub>0 \ cnj(csqrt z) = csqrt(cnj z)" by (metis cnj_Ln complex_cnj_divide complex_cnj_numeral complex_cnj_zero_iff csqrt_eq_0 csqrt_exp_Ln exp_cnj) 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 then have [simp]: "Im z = 0" and 0: "Re z < 0 \ z = 0" using complex_nonpos_Reals_iff complex_eq_iff by force+ 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))" by (simp add: Im_complex_div_lemma Re_complex_div_lemma assms cmod_eq_Im) 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 (smt (verit, best) Im_complex_of_real imaginary_unit.sel(2) of_real_minus power2_eq_iff power2_i) 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: "0 < x \arctan(inverse x) = pi/2 - arctan x" by (smt (verit, del_insts) arctan arctan_unique tan_cot zero_less_arctan_iff) 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_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: right_diff_distrib flip: power2_eq_square) 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 (simp add: Arcsin_unique) lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2" using Arcsin_unique sin_of_real_pi_half by fastforce lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)" by (simp add: Arcsin_unique) 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" by (metis Arcsin_body_lemma complex_i_mult_minus diff_0 diff_eq_eq power2_minus) 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))" by (smt (verit) Im_Arccos_bound Re_Arccos_bound cmod_le cos_Arccos) then show "cmod (Arccos w) \ pi + cmod w" by auto qed subsection\<^marker>\tag unimportant\\Interrelations between Arcsin and Arccos\ lemma cos_Arcsin_nonzero: "z\<^sup>2 \ 1 \cos(Arcsin z) \ 0" by (metis diff_0_right power_zero_numeral sin_Arcsin sin_squared_eq) lemma sin_Arccos_nonzero: "z\<^sup>2 \ 1 \sin(Arccos z) \ 0" by (metis add.right_neutral cos_Arccos power2_eq_square power_zero_numeral sin_cos_squared_add3) 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 of_real_arcsin: "\x\ \ 1 \ of_real(arcsin x) = Arcsin(of_real x)" by (smt (verit, best) Arcsin_sin Im_complex_of_real Re_complex_of_real arcsin sin_of_real) lemma Im_Arcsin_of_real: "\x\ \ 1 \ Im (Arcsin (of_real x)) = 0" by (metis Im_complex_of_real of_real_arcsin) 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: "\x\ \ 1 \ arcsin x = Re (Arcsin (of_real x))" by (metis Re_complex_of_real of_real_arcsin) subsection\<^marker>\tag unimportant\\Relationship with Arccos on the Real Numbers\ lemma of_real_arccos: "\x\ \ 1 \ of_real(arccos x) = Arccos(of_real x)" by (smt (verit, del_insts) Arccos_unique Im_complex_of_real Re_complex_of_real arccos_lbound arccos_ubound cos_arccos_abs cos_of_real) lemma Im_Arccos_of_real: "\x\ \ 1 \ Im (Arccos (of_real x)) = 0" by (metis Im_complex_of_real of_real_arccos) corollary\<^marker>\tag unimportant\ Arccos_in_Reals [simp]: "z \ \ \ \Re z\ \ 1 \ Arccos z \ \" by (metis Im_Arccos_of_real complex_is_Real_iff of_real_Re) lemma arccos_eq_Re_Arccos: "\x\ \ 1 \ arccos x = Re (Arccos (of_real x))" by (metis Re_complex_of_real of_real_arccos) 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" by (metis assms bot_nat_0.not_eq_extremum exp_divide_power_eq exp_of_nat2_mult exp_two_pi_i power_one) 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, opaque_lifting) 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/Library/Periodic_Fun.thy b/src/HOL/Library/Periodic_Fun.thy --- a/src/HOL/Library/Periodic_Fun.thy +++ b/src/HOL/Library/Periodic_Fun.thy @@ -1,132 +1,152 @@ (* Title: HOL/Library/Periodic_Fun.thy Author: Manuel Eberl, TU München *) section \Periodic Functions\ theory Periodic_Fun imports Complex_Main begin text \ A locale for periodic functions. The idea is that one proves $f(x + p) = f(x)$ for some period $p$ and gets derived results like $f(x - p) = f(x)$ and $f(x + 2p) = f(x)$ for free. \<^term>\g\ and \<^term>\gm\ are ``plus/minus k periods'' functions. \<^term>\g1\ and \<^term>\gn1\ are ``plus/minus one period'' functions. This is useful e.g. if the period is one; the lemmas one gets are then \<^term>\f (x + 1) = f x\ instead of \<^term>\f (x + 1 * 1) = f x\ etc. \ locale periodic_fun = fixes f :: "('a :: {ring_1}) \ 'b" and g gm :: "'a \ 'a \ 'a" and g1 gn1 :: "'a \ 'a" assumes plus_1: "f (g1 x) = f x" assumes periodic_arg_plus_0: "g x 0 = x" assumes periodic_arg_plus_distrib: "g x (of_int (m + n)) = g (g x (of_int n)) (of_int m)" assumes plus_1_eq: "g x 1 = g1 x" and minus_1_eq: "g x (-1) = gn1 x" and minus_eq: "g x (-y) = gm x y" begin lemma plus_of_nat: "f (g x (of_nat n)) = f x" by (induction n) (insert periodic_arg_plus_distrib[of _ 1 "int n" for n], simp_all add: plus_1 periodic_arg_plus_0 plus_1_eq) lemma minus_of_nat: "f (gm x (of_nat n)) = f x" proof - have "f (g x (- of_nat n)) = f (g (g x (- of_nat n)) (of_nat n))" by (rule plus_of_nat[symmetric]) also have "\ = f (g (g x (of_int (- of_nat n))) (of_int (of_nat n)))" by simp also have "\ = f x" by (subst periodic_arg_plus_distrib [symmetric]) (simp add: periodic_arg_plus_0) finally show ?thesis by (simp add: minus_eq) qed lemma plus_of_int: "f (g x (of_int n)) = f x" by (induction n) (simp_all add: plus_of_nat minus_of_nat minus_eq del: of_nat_Suc) lemma minus_of_int: "f (gm x (of_int n)) = f x" using plus_of_int[of x "of_int (-n)"] by (simp add: minus_eq) lemma plus_numeral: "f (g x (numeral n)) = f x" by (subst of_nat_numeral[symmetric], subst plus_of_nat) (rule refl) lemma minus_numeral: "f (gm x (numeral n)) = f x" by (subst of_nat_numeral[symmetric], subst minus_of_nat) (rule refl) lemma minus_1: "f (gn1 x) = f x" using minus_of_nat[of x 1] by (simp flip: minus_1_eq minus_eq) lemmas periodic_simps = plus_of_nat minus_of_nat plus_of_int minus_of_int plus_numeral minus_numeral plus_1 minus_1 end text \ Specialised case of the \<^term>\periodic_fun\ locale for periods that are not 1. Gives lemmas \<^term>\f (x - period) = f x\ etc. \ locale periodic_fun_simple = fixes f :: "('a :: {ring_1}) \ 'b" and period :: 'a assumes plus_period: "f (x + period) = f x" begin sublocale periodic_fun f "\z x. z + x * period" "\z x. z - x * period" "\z. z + period" "\z. z - period" by standard (simp_all add: ring_distribs plus_period) end text \ Specialised case of the \<^term>\periodic_fun\ locale for period 1. Gives lemmas \<^term>\f (x - 1) = f x\ etc. \ locale periodic_fun_simple' = fixes f :: "('a :: {ring_1}) \ 'b" assumes plus_period: "f (x + 1) = f x" begin sublocale periodic_fun f "\z x. z + x" "\z x. z - x" "\z. z + 1" "\z. z - 1" by standard (simp_all add: ring_distribs plus_period) lemma of_nat: "f (of_nat n) = f 0" using plus_of_nat[of 0 n] by simp lemma uminus_of_nat: "f (-of_nat n) = f 0" using minus_of_nat[of 0 n] by simp lemma of_int: "f (of_int n) = f 0" using plus_of_int[of 0 n] by simp lemma uminus_of_int: "f (-of_int n) = f 0" using minus_of_int[of 0 n] by simp lemma of_numeral: "f (numeral n) = f 0" using plus_numeral[of 0 n] by simp lemma of_neg_numeral: "f (-numeral n) = f 0" using minus_numeral[of 0 n] by simp lemma of_1: "f 1 = f 0" using plus_of_nat[of 0 1] by simp lemma of_neg_1: "f (-1) = f 0" using minus_of_nat[of 0 1] by simp lemmas periodic_simps' = of_nat uminus_of_nat of_int uminus_of_int of_numeral of_neg_numeral of_1 of_neg_1 end lemma sin_plus_pi: "sin ((z :: 'a :: {real_normed_field,banach}) + of_real pi) = - sin z" by (simp add: sin_add) lemma cos_plus_pi: "cos ((z :: 'a :: {real_normed_field,banach}) + of_real pi) = - cos z" by (simp add: cos_add) interpretation sin: periodic_fun_simple sin "2 * of_real pi :: 'a :: {real_normed_field,banach}" proof fix z :: 'a have "sin (z + 2 * of_real pi) = sin (z + of_real pi + of_real pi)" by (simp add: ac_simps) also have "\ = sin z" by (simp only: sin_plus_pi) simp finally show "sin (z + 2 * of_real pi) = sin z" . qed interpretation cos: periodic_fun_simple cos "2 * of_real pi :: 'a :: {real_normed_field,banach}" proof fix z :: 'a have "cos (z + 2 * of_real pi) = cos (z + of_real pi + of_real pi)" by (simp add: ac_simps) also have "\ = cos z" by (simp only: cos_plus_pi) simp finally show "cos (z + 2 * of_real pi) = cos z" . qed interpretation tan: periodic_fun_simple tan "2 * of_real pi :: 'a :: {real_normed_field,banach}" by standard (simp only: tan_def [abs_def] sin.plus_1 cos.plus_1) interpretation cot: periodic_fun_simple cot "2 * of_real pi :: 'a :: {real_normed_field,banach}" by standard (simp only: cot_def [abs_def] sin.plus_1 cos.plus_1) - + +lemma cos_eq_neg_periodic_intro: + assumes "x - y = 2*(of_int k)*pi + pi \ x + y = 2*(of_int k)*pi + pi" + shows "cos x = - cos y" using assms +proof + assume "x - y = 2 * (of_int k) * pi + pi" + then show ?thesis + using cos.periodic_simps[of "y+pi"] + by (auto simp add:algebra_simps) +next + assume "x + y = 2 * real_of_int k * pi + pi " + then show ?thesis + using cos.periodic_simps[of "-y+pi"] + by (clarsimp simp add: algebra_simps) (smt (verit)) +qed + +lemma cos_eq_periodic_intro: + assumes "x - y = 2*(of_int k)*pi \ x + y = 2*(of_int k)*pi" + shows "cos x = cos y" + by (smt (verit, ccfv_SIG) assms cos_eq_neg_periodic_intro cos_minus_pi cos_periodic_pi) + end diff --git a/src/HOL/Transcendental.thy b/src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy +++ b/src/HOL/Transcendental.thy @@ -1,7421 +1,7428 @@ (* Title: HOL/Transcendental.thy Author: Jacques D. Fleuriot, University of Cambridge, University of Edinburgh Author: Lawrence C Paulson Author: Jeremy Avigad *) section \Power Series, Transcendental Functions etc.\ theory Transcendental imports Series Deriv NthRoot begin text \A theorem about the factcorial function on the reals.\ lemma square_fact_le_2_fact: "fact n * fact n \ (fact (2 * n) :: real)" proof (induct n) case 0 then show ?case by simp next case (Suc n) have "(fact (Suc n)) * (fact (Suc n)) = of_nat (Suc n) * of_nat (Suc n) * (fact n * fact n :: real)" by (simp add: field_simps) also have "\ \ of_nat (Suc n) * of_nat (Suc n) * fact (2 * n)" by (rule mult_left_mono [OF Suc]) simp also have "\ \ of_nat (Suc (Suc (2 * n))) * of_nat (Suc (2 * n)) * fact (2 * n)" by (rule mult_right_mono)+ (auto simp: field_simps) also have "\ = fact (2 * Suc n)" by (simp add: field_simps) finally show ?case . qed lemma fact_in_Reals: "fact n \ \" by (induction n) auto lemma of_real_fact [simp]: "of_real (fact n) = fact n" by (metis of_nat_fact of_real_of_nat_eq) lemma pochhammer_of_real: "pochhammer (of_real x) n = of_real (pochhammer x n)" by (simp add: pochhammer_prod) lemma norm_fact [simp]: "norm (fact n :: 'a::real_normed_algebra_1) = fact n" proof - have "(fact n :: 'a) = of_real (fact n)" by simp also have "norm \ = fact n" by (subst norm_of_real) simp finally show ?thesis . qed lemma root_test_convergence: fixes f :: "nat \ 'a::banach" assumes f: "(\n. root n (norm (f n))) \ x" \ \could be weakened to lim sup\ and "x < 1" shows "summable f" proof - have "0 \ x" by (rule LIMSEQ_le[OF tendsto_const f]) (auto intro!: exI[of _ 1]) from \x < 1\ obtain z where z: "x < z" "z < 1" by (metis dense) from f \x < z\ have "eventually (\n. root n (norm (f n)) < z) sequentially" by (rule order_tendstoD) then have "eventually (\n. norm (f n) \ z^n) sequentially" using eventually_ge_at_top proof eventually_elim fix n assume less: "root n (norm (f n)) < z" and n: "1 \ n" from power_strict_mono[OF less, of n] n show "norm (f n) \ z ^ n" by simp qed then show "summable f" unfolding eventually_sequentially using z \0 \ x\ by (auto intro!: summable_comparison_test[OF _ summable_geometric]) qed subsection \More facts about binomial coefficients\ text \ These facts could have been proven before, but having real numbers makes the proofs a lot easier. \ lemma central_binomial_odd: "odd n \ n choose (Suc (n div 2)) = n choose (n div 2)" proof - assume "odd n" hence "Suc (n div 2) \ n" by presburger hence "n choose (Suc (n div 2)) = n choose (n - Suc (n div 2))" by (rule binomial_symmetric) also from \odd n\ have "n - Suc (n div 2) = n div 2" by presburger finally show ?thesis . qed lemma binomial_less_binomial_Suc: assumes k: "k < n div 2" shows "n choose k < n choose (Suc k)" proof - from k have k': "k \ n" "Suc k \ n" by simp_all from k' have "real (n choose k) = fact n / (fact k * fact (n - k))" by (simp add: binomial_fact) also from k' have "n - k = Suc (n - Suc k)" by simp also from k' have "fact \ = (real n - real k) * fact (n - Suc k)" by (subst fact_Suc) (simp_all add: of_nat_diff) also from k have "fact k = fact (Suc k) / (real k + 1)" by (simp add: field_simps) also have "fact n / (fact (Suc k) / (real k + 1) * ((real n - real k) * fact (n - Suc k))) = (n choose (Suc k)) * ((real k + 1) / (real n - real k))" using k by (simp add: field_split_simps binomial_fact) also from assms have "(real k + 1) / (real n - real k) < 1" by simp finally show ?thesis using k by (simp add: mult_less_cancel_left) qed lemma binomial_strict_mono: assumes "k < k'" "2*k' \ n" shows "n choose k < n choose k'" proof - from assms have "k \ k' - 1" by simp thus ?thesis proof (induction rule: inc_induct) case base with assms binomial_less_binomial_Suc[of "k' - 1" n] show ?case by simp next case (step k) from step.prems step.hyps assms have "n choose k < n choose (Suc k)" by (intro binomial_less_binomial_Suc) simp_all also have "\ < n choose k'" by (rule step.IH) finally show ?case . qed qed lemma binomial_mono: assumes "k \ k'" "2*k' \ n" shows "n choose k \ n choose k'" using assms binomial_strict_mono[of k k' n] by (cases "k = k'") simp_all lemma binomial_strict_antimono: assumes "k < k'" "2 * k \ n" "k' \ n" shows "n choose k > n choose k'" proof - from assms have "n choose (n - k) > n choose (n - k')" by (intro binomial_strict_mono) (simp_all add: algebra_simps) with assms show ?thesis by (simp add: binomial_symmetric [symmetric]) qed lemma binomial_antimono: assumes "k \ k'" "k \ n div 2" "k' \ n" shows "n choose k \ n choose k'" proof (cases "k = k'") case False note not_eq = False show ?thesis proof (cases "k = n div 2 \ odd n") case False with assms(2) have "2*k \ n" by presburger with not_eq assms binomial_strict_antimono[of k k' n] show ?thesis by simp next case True have "n choose k' \ n choose (Suc (n div 2))" proof (cases "k' = Suc (n div 2)") case False with assms True not_eq have "Suc (n div 2) < k'" by simp with assms binomial_strict_antimono[of "Suc (n div 2)" k' n] True show ?thesis by auto qed simp_all also from True have "\ = n choose k" by (simp add: central_binomial_odd) finally show ?thesis . qed qed simp_all lemma binomial_maximum: "n choose k \ n choose (n div 2)" proof - have "k \ n div 2 \ 2*k \ n" by linarith consider "2*k \ n" | "2*k \ n" "k \ n" | "k > n" by linarith thus ?thesis proof cases case 1 thus ?thesis by (intro binomial_mono) linarith+ next case 2 thus ?thesis by (intro binomial_antimono) simp_all qed (simp_all add: binomial_eq_0) qed lemma binomial_maximum': "(2*n) choose k \ (2*n) choose n" using binomial_maximum[of "2*n"] by simp lemma central_binomial_lower_bound: assumes "n > 0" shows "4^n / (2*real n) \ real ((2*n) choose n)" proof - from binomial[of 1 1 "2*n"] have "4 ^ n = (\k\2*n. (2*n) choose k)" by (simp add: power_mult power2_eq_square One_nat_def [symmetric] del: One_nat_def) also have "{..2*n} = {0<..<2*n} \ {0,2*n}" by auto also have "(\k\\. (2*n) choose k) = (\k\{0<..<2*n}. (2*n) choose k) + (\k\{0,2*n}. (2*n) choose k)" by (subst sum.union_disjoint) auto also have "(\k\{0,2*n}. (2*n) choose k) \ (\k\1. (n choose k)\<^sup>2)" by (cases n) simp_all also from assms have "\ \ (\k\n. (n choose k)\<^sup>2)" by (intro sum_mono2) auto also have "\ = (2*n) choose n" by (rule choose_square_sum) also have "(\k\{0<..<2*n}. (2*n) choose k) \ (\k\{0<..<2*n}. (2*n) choose n)" by (intro sum_mono binomial_maximum') also have "\ = card {0<..<2*n} * ((2*n) choose n)" by simp also have "card {0<..<2*n} \ 2*n - 1" by (cases n) simp_all also have "(2 * n - 1) * (2 * n choose n) + (2 * n choose n) = ((2*n) choose n) * (2*n)" using assms by (simp add: algebra_simps) finally have "4 ^ n \ (2 * n choose n) * (2 * n)" by simp_all hence "real (4 ^ n) \ real ((2 * n choose n) * (2 * n))" by (subst of_nat_le_iff) with assms show ?thesis by (simp add: field_simps) qed subsection \Properties of Power Series\ lemma powser_zero [simp]: "(\n. f n * 0 ^ n) = f 0" for f :: "nat \ 'a::real_normed_algebra_1" proof - have "(\n<1. f n * 0 ^ n) = (\n. f n * 0 ^ n)" by (subst suminf_finite[where N="{0}"]) (auto simp: power_0_left) then show ?thesis by simp qed lemma powser_sums_zero: "(\n. a n * 0^n) sums a 0" for a :: "nat \ 'a::real_normed_div_algebra" using sums_finite [of "{0}" "\n. a n * 0 ^ n"] by simp lemma powser_sums_zero_iff [simp]: "(\n. a n * 0^n) sums x \ a 0 = x" for a :: "nat \ 'a::real_normed_div_algebra" using powser_sums_zero sums_unique2 by blast text \ Power series has a circle or radius of convergence: if it sums for \x\, then it sums absolutely for \z\ with \<^term>\\z\ < \x\\.\ lemma powser_insidea: fixes x z :: "'a::real_normed_div_algebra" assumes 1: "summable (\n. f n * x^n)" and 2: "norm z < norm x" shows "summable (\n. norm (f n * z ^ n))" proof - from 2 have x_neq_0: "x \ 0" by clarsimp from 1 have "(\n. f n * x^n) \ 0" by (rule summable_LIMSEQ_zero) then have "convergent (\n. f n * x^n)" by (rule convergentI) then have "Cauchy (\n. f n * x^n)" by (rule convergent_Cauchy) then have "Bseq (\n. f n * x^n)" by (rule Cauchy_Bseq) then obtain K where 3: "0 < K" and 4: "\n. norm (f n * x^n) \ K" by (auto simp: Bseq_def) have "\N. \n\N. norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" proof (intro exI allI impI) fix n :: nat assume "0 \ n" have "norm (norm (f n * z ^ n)) * norm (x^n) = norm (f n * x^n) * norm (z ^ n)" by (simp add: norm_mult abs_mult) also have "\ \ K * norm (z ^ n)" by (simp only: mult_right_mono 4 norm_ge_zero) also have "\ = K * norm (z ^ n) * (inverse (norm (x^n)) * norm (x^n))" by (simp add: x_neq_0) also have "\ = K * norm (z ^ n) * inverse (norm (x^n)) * norm (x^n)" by (simp only: mult.assoc) finally show "norm (norm (f n * z ^ n)) \ K * norm (z ^ n) * inverse (norm (x^n))" by (simp add: mult_le_cancel_right x_neq_0) qed moreover have "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" proof - from 2 have "norm (norm (z * inverse x)) < 1" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse divide_inverse [where 'a=real, symmetric]) then have "summable (\n. norm (z * inverse x) ^ n)" by (rule summable_geometric) then have "summable (\n. K * norm (z * inverse x) ^ n)" by (rule summable_mult) then show "summable (\n. K * norm (z ^ n) * inverse (norm (x^n)))" using x_neq_0 by (simp add: norm_mult nonzero_norm_inverse power_mult_distrib power_inverse norm_power mult.assoc) qed ultimately show "summable (\n. norm (f n * z ^ n))" by (rule summable_comparison_test) qed lemma powser_inside: fixes f :: "nat \ 'a::{real_normed_div_algebra,banach}" shows "summable (\n. f n * (x^n)) \ norm z < norm x \ summable (\n. f n * (z ^ n))" by (rule powser_insidea [THEN summable_norm_cancel]) lemma powser_times_n_limit_0: fixes x :: "'a::{real_normed_div_algebra,banach}" assumes "norm x < 1" shows "(\n. of_nat n * x ^ n) \ 0" proof - have "norm x / (1 - norm x) \ 0" using assms by (auto simp: field_split_simps) moreover obtain N where N: "norm x / (1 - norm x) < of_int N" using ex_le_of_int by (meson ex_less_of_int) ultimately have N0: "N>0" by auto then have *: "real_of_int (N + 1) * norm x / real_of_int N < 1" using N assms by (auto simp: field_simps) have **: "real_of_int N * (norm x * (real_of_nat (Suc n) * norm (x ^ n))) \ real_of_nat n * (norm x * ((1 + N) * norm (x ^ n)))" if "N \ int n" for n :: nat proof - from that have "real_of_int N * real_of_nat (Suc n) \ real_of_nat n * real_of_int (1 + N)" by (simp add: algebra_simps) then have "(real_of_int N * real_of_nat (Suc n)) * (norm x * norm (x ^ n)) \ (real_of_nat n * (1 + N)) * (norm x * norm (x ^ n))" using N0 mult_mono by fastforce then show ?thesis by (simp add: algebra_simps) qed show ?thesis using * by (rule summable_LIMSEQ_zero [OF summable_ratio_test, where N1="nat N"]) (simp add: N0 norm_mult field_simps ** del: of_nat_Suc of_int_add) qed corollary lim_n_over_pown: fixes x :: "'a::{real_normed_field,banach}" shows "1 < norm x \ ((\n. of_nat n / x^n) \ 0) sequentially" using powser_times_n_limit_0 [of "inverse x"] by (simp add: norm_divide field_split_simps) lemma sum_split_even_odd: fixes f :: "nat \ real" shows "(\i<2 * n. if even i then f i else g i) = (\iii<2 * Suc n. if even i then f i else g i) = (\ii = (\ii real" assumes "g sums x" shows "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" unfolding sums_def proof (rule LIMSEQ_I) fix r :: real assume "0 < r" from \g sums x\[unfolded sums_def, THEN LIMSEQ_D, OF this] obtain no where no_eq: "\n. n \ no \ (norm (sum g {.. 2 * no" for m proof - from that have "m div 2 \ no" by auto have sum_eq: "?SUM (2 * (m div 2)) = sum g {..< m div 2}" using sum_split_even_odd by auto then have "(norm (?SUM (2 * (m div 2)) - x) < r)" using no_eq unfolding sum_eq using \m div 2 \ no\ by auto moreover have "?SUM (2 * (m div 2)) = ?SUM m" proof (cases "even m") case True then show ?thesis by (auto simp: even_two_times_div_two) next case False then have eq: "Suc (2 * (m div 2)) = m" by simp then have "even (2 * (m div 2))" using \odd m\ by auto have "?SUM m = ?SUM (Suc (2 * (m div 2)))" unfolding eq .. also have "\ = ?SUM (2 * (m div 2))" using \even (2 * (m div 2))\ by auto finally show ?thesis by auto qed ultimately show ?thesis by auto qed then show "\no. \ m \ no. norm (?SUM m - x) < r" by blast qed lemma sums_if: fixes g :: "nat \ real" assumes "g sums x" and "f sums y" shows "(\ n. if even n then f (n div 2) else g ((n - 1) div 2)) sums (x + y)" proof - let ?s = "\ n. if even n then 0 else f ((n - 1) div 2)" have if_sum: "(if B then (0 :: real) else E) + (if B then T else 0) = (if B then T else E)" for B T E by (cases B) auto have g_sums: "(\ n. if even n then 0 else g ((n - 1) div 2)) sums x" using sums_if'[OF \g sums x\] . have if_eq: "\B T E. (if \ B then T else E) = (if B then E else T)" by auto have "?s sums y" using sums_if'[OF \f sums y\] . from this[unfolded sums_def, THEN LIMSEQ_Suc] have "(\n. if even n then f (n div 2) else 0) sums y" by (simp add: lessThan_Suc_eq_insert_0 sum.atLeast1_atMost_eq image_Suc_lessThan if_eq sums_def cong del: if_weak_cong) from sums_add[OF g_sums this] show ?thesis by (simp only: if_sum) qed subsection \Alternating series test / Leibniz formula\ (* FIXME: generalise these results from the reals via type classes? *) lemma sums_alternating_upper_lower: fixes a :: "nat \ real" assumes mono: "\n. a (Suc n) \ a n" and a_pos: "\n. 0 \ a n" and "a \ 0" shows "\l. ((\n. (\i<2*n. (- 1)^i*a i) \ l) \ (\ n. \i<2*n. (- 1)^i*a i) \ l) \ ((\n. l \ (\i<2*n + 1. (- 1)^i*a i)) \ (\ n. \i<2*n + 1. (- 1)^i*a i) \ l)" (is "\l. ((\n. ?f n \ l) \ _) \ ((\n. l \ ?g n) \ _)") proof (rule nested_sequence_unique) have fg_diff: "\n. ?f n - ?g n = - a (2 * n)" by auto show "\n. ?f n \ ?f (Suc n)" proof show "?f n \ ?f (Suc n)" for n using mono[of "2*n"] by auto qed show "\n. ?g (Suc n) \ ?g n" proof show "?g (Suc n) \ ?g n" for n using mono[of "Suc (2*n)"] by auto qed show "\n. ?f n \ ?g n" proof show "?f n \ ?g n" for n using fg_diff a_pos by auto qed show "(\n. ?f n - ?g n) \ 0" unfolding fg_diff proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \a \ 0\[THEN LIMSEQ_D] obtain N where "\ n. n \ N \ norm (a n - 0) < r" by auto then have "\n \ N. norm (- a (2 * n) - 0) < r" by auto then show "\N. \n \ N. norm (- a (2 * n) - 0) < r" by auto qed qed lemma summable_Leibniz': fixes a :: "nat \ real" assumes a_zero: "a \ 0" and a_pos: "\n. 0 \ a n" and a_monotone: "\n. a (Suc n) \ a n" shows summable: "summable (\ n. (-1)^n * a n)" and "\n. (\i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "(\n. \i<2*n. (-1)^i*a i) \ (\i. (-1)^i*a i)" and "\n. (\i. (-1)^i*a i) \ (\i<2*n+1. (-1)^i*a i)" and "(\n. \i<2*n+1. (-1)^i*a i) \ (\i. (-1)^i*a i)" proof - let ?S = "\n. (-1)^n * a n" let ?P = "\n. \i n. ?f n \ l" and "?f \ l" and above_l: "\ n. l \ ?g n" and "?g \ l" using sums_alternating_upper_lower[OF a_monotone a_pos a_zero] by blast let ?Sa = "\m. \n l" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" with \?f \ l\[THEN LIMSEQ_D] obtain f_no where f: "\n. n \ f_no \ norm (?f n - l) < r" by auto from \0 < r\ \?g \ l\[THEN LIMSEQ_D] obtain g_no where g: "\n. n \ g_no \ norm (?g n - l) < r" by auto have "norm (?Sa n - l) < r" if "n \ (max (2 * f_no) (2 * g_no))" for n proof - from that have "n \ 2 * f_no" and "n \ 2 * g_no" by auto show ?thesis proof (cases "even n") case True then have n_eq: "2 * (n div 2) = n" by (simp add: even_two_times_div_two) with \n \ 2 * f_no\ have "n div 2 \ f_no" by auto from f[OF this] show ?thesis unfolding n_eq atLeastLessThanSuc_atLeastAtMost . next case False then have "even (n - 1)" by simp then have n_eq: "2 * ((n - 1) div 2) = n - 1" by (simp add: even_two_times_div_two) then have range_eq: "n - 1 + 1 = n" using odd_pos[OF False] by auto from n_eq \n \ 2 * g_no\ have "(n - 1) div 2 \ g_no" by auto from g[OF this] show ?thesis by (simp only: n_eq range_eq) qed qed then show "\no. \n \ no. norm (?Sa n - l) < r" by blast qed then have sums_l: "(\i. (-1)^i * a i) sums l" by (simp only: sums_def) then show "summable ?S" by (auto simp: summable_def) have "l = suminf ?S" by (rule sums_unique[OF sums_l]) fix n show "suminf ?S \ ?g n" unfolding sums_unique[OF sums_l, symmetric] using above_l by auto show "?f n \ suminf ?S" unfolding sums_unique[OF sums_l, symmetric] using below_l by auto show "?g \ suminf ?S" using \?g \ l\ \l = suminf ?S\ by auto show "?f \ suminf ?S" using \?f \ l\ \l = suminf ?S\ by auto qed theorem summable_Leibniz: fixes a :: "nat \ real" assumes a_zero: "a \ 0" and "monoseq a" shows "summable (\ n. (-1)^n * a n)" (is "?summable") and "0 < a 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n. (- 1)^i * a i .. \i<2*n+1. (- 1)^i * a i})" (is "?pos") and "a 0 < 0 \ (\n. (\i. (- 1)^i*a i) \ { \i<2*n+1. (- 1)^i * a i .. \i<2*n. (- 1)^i * a i})" (is "?neg") and "(\n. \i<2*n. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?f") and "(\n. \i<2*n+1. (- 1)^i*a i) \ (\i. (- 1)^i*a i)" (is "?g") proof - have "?summable \ ?pos \ ?neg \ ?f \ ?g" proof (cases "(\n. 0 \ a n) \ (\m. \n\m. a n \ a m)") case True then have ord: "\n m. m \ n \ a n \ a m" and ge0: "\n. 0 \ a n" by auto have mono: "a (Suc n) \ a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF \a \ 0\ ge0] from leibniz[OF mono] show ?thesis using \0 \ a 0\ by auto next let ?a = "\n. - a n" case False with monoseq_le[OF \monoseq a\ \a \ 0\] have "(\ n. a n \ 0) \ (\m. \n\m. a m \ a n)" by auto then have ord: "\n m. m \ n \ ?a n \ ?a m" and ge0: "\ n. 0 \ ?a n" by auto have monotone: "?a (Suc n) \ ?a n" for n using ord[where n="Suc n" and m=n] by auto note leibniz = summable_Leibniz'[OF _ ge0, of "\x. x", OF tendsto_minus[OF \a \ 0\, unfolded minus_zero] monotone] have "summable (\ n. (-1)^n * ?a n)" using leibniz(1) by auto then obtain l where "(\ n. (-1)^n * ?a n) sums l" unfolding summable_def by auto from this[THEN sums_minus] have "(\ n. (-1)^n * a n) sums -l" by auto then have ?summable by (auto simp: summable_def) moreover have "\- a - - b\ = \a - b\" for a b :: real unfolding minus_diff_minus by auto from suminf_minus[OF leibniz(1), unfolded mult_minus_right minus_minus] have move_minus: "(\n. - ((- 1) ^ n * a n)) = - (\n. (- 1) ^ n * a n)" by auto have ?pos using \0 \ ?a 0\ by auto moreover have ?neg using leibniz(2,4) unfolding mult_minus_right sum_negf move_minus neg_le_iff_le by auto moreover have ?f and ?g using leibniz(3,5)[unfolded mult_minus_right sum_negf move_minus, THEN tendsto_minus_cancel] by auto ultimately show ?thesis by auto qed then show ?summable and ?pos and ?neg and ?f and ?g by safe qed subsection \Term-by-Term Differentiability of Power Series\ definition diffs :: "(nat \ 'a::ring_1) \ nat \ 'a" where "diffs c = (\n. of_nat (Suc n) * c (Suc n))" text \Lemma about distributing negation over it.\ lemma diffs_minus: "diffs (\n. - c n) = (\n. - diffs c n)" by (simp add: diffs_def) lemma diffs_equiv: fixes x :: "'a::{real_normed_vector,ring_1}" shows "summable (\n. diffs c n * x^n) \ (\n. of_nat n * c n * x^(n - Suc 0)) sums (\n. diffs c n * x^n)" unfolding diffs_def by (simp add: summable_sums sums_Suc_imp) lemma lemma_termdiff1: fixes z :: "'a :: {monoid_mult,comm_ring}" shows "(\ppi 0" shows "((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0) = h * (\p< n - Suc 0. \q< n - Suc 0 - p. (z + h) ^ q * z ^ (n - 2 - q))" (is "?lhs = ?rhs") proof (cases n) case (Suc m) have 0: "\x k. (\njiijppip::nat. p < n \ f p \ K" and K: "0 \ K" shows "sum f {.. of_nat n * K" proof - have "sum f {.. (\i of_nat n * K" by (auto simp: mult_right_mono K) finally show ?thesis . qed lemma lemma_termdiff3: fixes h z :: "'a::real_normed_field" assumes 1: "h \ 0" and 2: "norm z \ K" and 3: "norm (z + h) \ K" shows "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) \ of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" proof - have "norm (((z + h) ^ n - z ^ n) / h - of_nat n * z ^ (n - Suc 0)) = norm (\pq \ of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2)) * norm h" proof (rule mult_right_mono [OF _ norm_ge_zero]) from norm_ge_zero 2 have K: "0 \ K" by (rule order_trans) have le_Kn: "norm ((z + h) ^ i * z ^ j) \ K ^ n" if "i + j = n" for i j n proof - have "norm (z + h) ^ i * norm z ^ j \ K ^ i * K ^ j" by (intro mult_mono power_mono 2 3 norm_ge_zero zero_le_power K) also have "... = K^n" by (metis power_add that) finally show ?thesis by (simp add: norm_mult norm_power) qed then have "\p q. \p < n; q < n - Suc 0\ \ norm ((z + h) ^ q * z ^ (n - 2 - q)) \ K ^ (n - 2)" by (simp del: subst_all) then show "norm (\pq of_nat n * (of_nat (n - Suc 0) * K ^ (n - 2))" by (intro order_trans [OF norm_sum] real_sum_nat_ivl_bounded2 mult_nonneg_nonneg of_nat_0_le_iff zero_le_power K) qed also have "\ = of_nat n * of_nat (n - Suc 0) * K ^ (n - 2) * norm h" by (simp only: mult.assoc) finally show ?thesis . qed lemma lemma_termdiff4: fixes f :: "'a::real_normed_vector \ 'b::real_normed_vector" and k :: real assumes k: "0 < k" and le: "\h. h \ 0 \ norm h < k \ norm (f h) \ K * norm h" shows "f \0\ 0" proof (rule tendsto_norm_zero_cancel) show "(\h. norm (f h)) \0\ 0" proof (rule real_tendsto_sandwich) show "eventually (\h. 0 \ norm (f h)) (at 0)" by simp show "eventually (\h. norm (f h) \ K * norm h) (at 0)" using k by (auto simp: eventually_at dist_norm le) show "(\h. 0) \(0::'a)\ (0::real)" by (rule tendsto_const) have "(\h. K * norm h) \(0::'a)\ K * norm (0::'a)" by (intro tendsto_intros) then show "(\h. K * norm h) \(0::'a)\ 0" by simp qed qed lemma lemma_termdiff5: fixes g :: "'a::real_normed_vector \ nat \ 'b::banach" and k :: real assumes k: "0 < k" and f: "summable f" and le: "\h n. h \ 0 \ norm h < k \ norm (g h n) \ f n * norm h" shows "(\h. suminf (g h)) \0\ 0" proof (rule lemma_termdiff4 [OF k]) fix h :: 'a assume "h \ 0" and "norm h < k" then have 1: "\n. norm (g h n) \ f n * norm h" by (simp add: le) then have "\N. \n\N. norm (norm (g h n)) \ f n * norm h" by simp moreover from f have 2: "summable (\n. f n * norm h)" by (rule summable_mult2) ultimately have 3: "summable (\n. norm (g h n))" by (rule summable_comparison_test) then have "norm (suminf (g h)) \ (\n. norm (g h n))" by (rule summable_norm) also from 1 3 2 have "(\n. norm (g h n)) \ (\n. f n * norm h)" by (simp add: suminf_le) also from f have "(\n. f n * norm h) = suminf f * norm h" by (rule suminf_mult2 [symmetric]) finally show "norm (suminf (g h)) \ suminf f * norm h" . qed (* FIXME: Long proofs *) lemma termdiffs_aux: fixes x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. diffs (diffs c) n * K ^ n)" and 2: "norm x < norm K" shows "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" proof - from dense [OF 2] obtain r where r1: "norm x < r" and r2: "r < norm K" by fast from norm_ge_zero r1 have r: "0 < r" by (rule order_le_less_trans) then have r_neq_0: "r \ 0" by simp show ?thesis proof (rule lemma_termdiff5) show "0 < r - norm x" using r1 by simp from r r2 have "norm (of_real r::'a) < norm K" by simp with 1 have "summable (\n. norm (diffs (diffs c) n * (of_real r ^ n)))" by (rule powser_insidea) then have "summable (\n. diffs (diffs (\n. norm (c n))) n * r ^ n)" using r by (simp add: diffs_def norm_mult norm_power del: of_nat_Suc) then have "summable (\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * diffs (\n. norm (c n)) n * r ^ (n - Suc 0)) = (\n. diffs (\m. of_nat (m - Suc 0) * norm (c m) * inverse r) n * (r ^ n))" by (simp add: diffs_def r_neq_0 fun_eq_iff split: nat_diff_split) finally have "summable (\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0))" by (rule diffs_equiv [THEN sums_summable]) also have "(\n. of_nat n * (of_nat (n - Suc 0) * norm (c n) * inverse r) * r ^ (n - Suc 0)) = (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" by (rule ext) (simp add: r_neq_0 split: nat_diff_split) finally show "summable (\n. norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2))" . next fix h :: 'a and n assume h: "h \ 0" assume "norm h < r - norm x" then have "norm x + norm h < r" by simp with norm_triangle_ineq have xh: "norm (x + h) < r" by (rule order_le_less_trans) have "norm (((x + h) ^ n - x ^ n) / h - of_nat n * x ^ (n - Suc 0)) \ real n * (real (n - Suc 0) * (r ^ (n - 2) * norm h))" by (metis (mono_tags, lifting) h mult.assoc lemma_termdiff3 less_eq_real_def r1 xh) then show "norm (c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \ norm (c n) * of_nat n * of_nat (n - Suc 0) * r ^ (n - 2) * norm h" by (simp only: norm_mult mult.assoc mult_left_mono [OF _ norm_ge_zero]) qed qed lemma termdiffs: fixes K x :: "'a::{real_normed_field,banach}" assumes 1: "summable (\n. c n * K ^ n)" and 2: "summable (\n. (diffs c) n * K ^ n)" and 3: "summable (\n. (diffs (diffs c)) n * K ^ n)" and 4: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. (diffs c) n * x^n)" unfolding DERIV_def proof (rule LIM_zero_cancel) show "(\h. (suminf (\n. c n * (x + h) ^ n) - suminf (\n. c n * x^n)) / h - suminf (\n. diffs c n * x^n)) \0\ 0" proof (rule LIM_equal2) show "0 < norm K - norm x" using 4 by (simp add: less_diff_eq) next fix h :: 'a assume "norm (h - 0) < norm K - norm x" then have "norm x + norm h < norm K" by simp then have 5: "norm (x + h) < norm K" by (rule norm_triangle_ineq [THEN order_le_less_trans]) have "summable (\n. c n * x^n)" and "summable (\n. c n * (x + h) ^ n)" and "summable (\n. diffs c n * x^n)" using 1 2 4 5 by (auto elim: powser_inside) then have "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. (c n * (x + h) ^ n - c n * x^n) / h - of_nat n * c n * x ^ (n - Suc 0))" by (intro sums_unique sums_diff sums_divide diffs_equiv summable_sums) then show "((\n. c n * (x + h) ^ n) - (\n. c n * x^n)) / h - (\n. diffs c n * x^n) = (\n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0)))" by (simp add: algebra_simps) next show "(\h. \n. c n * (((x + h) ^ n - x^n) / h - of_nat n * x ^ (n - Suc 0))) \0\ 0" by (rule termdiffs_aux [OF 3 4]) qed qed subsection \The Derivative of a Power Series Has the Same Radius of Convergence\ lemma termdiff_converges: fixes x :: "'a::{real_normed_field,banach}" assumes K: "norm x < K" and sm: "\x. norm x < K \ summable(\n. c n * x ^ n)" shows "summable (\n. diffs c n * x ^ n)" proof (cases "x = 0") case True then show ?thesis using powser_sums_zero sums_summable by auto next case False then have "K > 0" using K less_trans zero_less_norm_iff by blast then obtain r :: real where r: "norm x < norm r" "norm r < K" "r > 0" using K False by (auto simp: field_simps abs_less_iff add_pos_pos intro: that [of "(norm x + K) / 2"]) have to0: "(\n. of_nat n * (x / of_real r) ^ n) \ 0" using r by (simp add: norm_divide powser_times_n_limit_0 [of "x / of_real r"]) obtain N where N: "\n. n\N \ real_of_nat n * norm x ^ n < r ^ n" using r LIMSEQ_D [OF to0, of 1] by (auto simp: norm_divide norm_mult norm_power field_simps) have "summable (\n. (of_nat n * c n) * x ^ n)" proof (rule summable_comparison_test') show "summable (\n. norm (c n * of_real r ^ n))" apply (rule powser_insidea [OF sm [of "of_real ((r+K)/2)"]]) using N r norm_of_real [of "r + K", where 'a = 'a] by auto show "\n. N \ n \ norm (of_nat n * c n * x ^ n) \ norm (c n * of_real r ^ n)" using N r by (fastforce simp add: norm_mult norm_power less_eq_real_def) qed then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ Suc n)" using summable_iff_shift [of "\n. of_nat n * c n * x ^ n" 1] by simp then have "summable (\n. (of_nat (Suc n) * c(Suc n)) * x ^ n)" using False summable_mult2 [of "\n. (of_nat (Suc n) * c(Suc n) * x ^ n) * x" "inverse x"] by (simp add: mult.assoc) (auto simp: ac_simps) then show ?thesis by (simp add: diffs_def) qed lemma termdiff_converges_all: fixes x :: "'a::{real_normed_field,banach}" assumes "\x. summable (\n. c n * x^n)" shows "summable (\n. diffs c n * x^n)" by (rule termdiff_converges [where K = "1 + norm x"]) (use assms in auto) lemma termdiffs_strong: fixes K x :: "'a::{real_normed_field,banach}" assumes sm: "summable (\n. c n * K ^ n)" and K: "norm x < norm K" shows "DERIV (\x. \n. c n * x^n) x :> (\n. diffs c n * x^n)" proof - have "norm K + norm x < norm K + norm K" using K by force then have K2: "norm ((of_real (norm K) + of_real (norm x)) / 2 :: 'a) < norm K" by (auto simp: norm_triangle_lt norm_divide field_simps) then have [simp]: "norm ((of_real (norm K) + of_real (norm x)) :: 'a) < norm K * 2" by simp have "summable (\n. c n * (of_real (norm x + norm K) / 2) ^ n)" by (metis K2 summable_norm_cancel [OF powser_insidea [OF sm]] add.commute of_real_add) moreover have "\x. norm x < norm K \ summable (\n. diffs c n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) moreover have "\x. norm x < norm K \ summable (\n. diffs(diffs c) n * x ^ n)" by (blast intro: sm termdiff_converges powser_inside) ultimately show ?thesis by (rule termdiffs [where K = "of_real (norm x + norm K) / 2"]) (use K in \auto simp: field_simps simp flip: of_real_add\) qed lemma termdiffs_strong_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "((\x. \n. c n * x^n) has_field_derivative (\n. diffs c n * x^n)) (at x)" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force simp del: of_real_add) lemma termdiffs_strong': fixes z :: "'a :: {real_normed_field,banach}" assumes "\z. norm z < K \ summable (\n. c n * z ^ n)" assumes "norm z < K" shows "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" proof (rule termdiffs_strong) define L :: real where "L = (norm z + K) / 2" have "0 \ norm z" by simp also note \norm z < K\ finally have K: "K \ 0" by simp from assms K have L: "L \ 0" "norm z < L" "L < K" by (simp_all add: L_def) from L show "norm z < norm (of_real L :: 'a)" by simp from L show "summable (\n. c n * of_real L ^ n)" by (intro assms(1)) simp_all qed lemma termdiffs_sums_strong: fixes z :: "'a :: {banach,real_normed_field}" assumes sums: "\z. norm z < K \ (\n. c n * z ^ n) sums f z" assumes deriv: "(f has_field_derivative f') (at z)" assumes norm: "norm z < K" shows "(\n. diffs c n * z ^ n) sums f'" proof - have summable: "summable (\n. diffs c n * z^n)" by (intro termdiff_converges[OF norm] sums_summable[OF sums]) from norm have "eventually (\z. z \ norm -` {..z. (\n. c n * z^n) = f z) (nhds z)" by eventually_elim (insert sums, simp add: sums_iff) have "((\z. \n. c n * z^n) has_field_derivative (\n. diffs c n * z^n)) (at z)" by (intro termdiffs_strong'[OF _ norm] sums_summable[OF sums]) hence "(f has_field_derivative (\n. diffs c n * z^n)) (at z)" by (subst (asm) DERIV_cong_ev[OF refl eq refl]) from this and deriv have "(\n. diffs c n * z^n) = f'" by (rule DERIV_unique) with summable show ?thesis by (simp add: sums_iff) qed lemma isCont_powser: fixes K x :: "'a::{real_normed_field,banach}" assumes "summable (\n. c n * K ^ n)" assumes "norm x < norm K" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms] by (blast intro!: DERIV_isCont) lemmas isCont_powser' = isCont_o2[OF _ isCont_powser] lemma isCont_powser_converges_everywhere: fixes K x :: "'a::{real_normed_field,banach}" assumes "\y. summable (\n. c n * y ^ n)" shows "isCont (\x. \n. c n * x^n) x" using termdiffs_strong[OF assms[of "of_real (norm x + 1)"], of x] by (force intro!: DERIV_isCont simp del: of_real_add) lemma powser_limit_0: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have "norm (of_real s / 2 :: 'a) < s" using s by (auto simp: norm_divide) then have "summable (\n. a n * (of_real s / 2) ^ n)" by (rule sums_summable [OF sm]) then have "((\x. \n. a n * x ^ n) has_field_derivative (\n. diffs a n * 0 ^ n)) (at 0)" by (rule termdiffs_strong) (use s in \auto simp: norm_divide\) then have "isCont (\x. \n. a n * x ^ n) 0" by (blast intro: DERIV_continuous) then have "((\x. \n. a n * x ^ n) \ a 0) (at 0)" by (simp add: continuous_within) moreover have "(\x. f x - (\n. a n * x ^ n)) \0\ 0" apply (clarsimp simp: LIM_eq) apply (rule_tac x=s in exI) using s sm sums_unique by fastforce ultimately show ?thesis by (rule Lim_transform) qed lemma powser_limit_0_strong: fixes a :: "nat \ 'a::{real_normed_field,banach}" assumes s: "0 < s" and sm: "\x. x \ 0 \ norm x < s \ (\n. a n * x ^ n) sums (f x)" shows "(f \ a 0) (at 0)" proof - have *: "((\x. if x = 0 then a 0 else f x) \ a 0) (at 0)" by (rule powser_limit_0 [OF s]) (auto simp: powser_sums_zero sm) show ?thesis using "*" by (auto cong: Lim_cong_within) qed subsection \Derivability of power series\ lemma DERIV_series': fixes f :: "real \ nat \ real" assumes DERIV_f: "\ n. DERIV (\ x. f x n) x0 :> (f' x0 n)" and allf_summable: "\ x. x \ {a <..< b} \ summable (f x)" and x0_in_I: "x0 \ {a <..< b}" and "summable (f' x0)" and "summable L" and L_def: "\n x y. x \ {a <..< b} \ y \ {a <..< b} \ \f x n - f y n\ \ L n * \x - y\" shows "DERIV (\ x. suminf (f x)) x0 :> (suminf (f' x0))" unfolding DERIV_def proof (rule LIM_I) fix r :: real assume "0 < r" then have "0 < r/3" by auto obtain N_L where N_L: "\ n. N_L \ n \ \ \ i. L (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable L\] by auto obtain N_f' where N_f': "\ n. N_f' \ n \ \ \ i. f' x0 (i + n) \ < r/3" using suminf_exist_split[OF \0 < r/3\ \summable (f' x0)\] by auto let ?N = "Suc (max N_L N_f')" have "\ \ i. f' x0 (i + ?N) \ < r/3" (is "?f'_part < r/3") and L_estimate: "\ \ i. L (i + ?N) \ < r/3" using N_L[of "?N"] and N_f' [of "?N"] by auto let ?diff = "\i x. (f (x0 + x) i - f x0 i) / x" let ?r = "r / (3 * real ?N)" from \0 < r\ have "0 < ?r" by simp let ?s = "\n. SOME s. 0 < s \ (\ x. x \ 0 \ \ x \ < s \ \ ?diff n x - f' x0 n \ < ?r)" define S' where "S' = Min (?s ` {..< ?N })" have "0 < S'" unfolding S'_def proof (rule iffD2[OF Min_gr_iff]) show "\x \ (?s ` {..< ?N }). 0 < x" proof fix x assume "x \ ?s ` {.. {..0 < ?r\, unfolded real_norm_def] obtain s where s_bound: "0 < s \ (\x. x \ 0 \ \x\ < s \ \?diff n x - f' x0 n\ < ?r)" by auto have "0 < ?s n" by (rule someI2[where a=s]) (auto simp: s_bound simp del: of_nat_Suc) then show "0 < x" by (simp only: \x = ?s n\) qed qed auto define S where "S = min (min (x0 - a) (b - x0)) S'" then have "0 < S" and S_a: "S \ x0 - a" and S_b: "S \ b - x0" and "S \ S'" using x0_in_I and \0 < S'\ by auto have "\(suminf (f (x0 + x)) - suminf (f x0)) / x - suminf (f' x0)\ < r" if "x \ 0" and "\x\ < S" for x proof - from that have x_in_I: "x0 + x \ {a <..< b}" using S_a S_b by auto note diff_smbl = summable_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] note div_smbl = summable_divide[OF diff_smbl] note all_smbl = summable_diff[OF div_smbl \summable (f' x0)\] note ign = summable_ignore_initial_segment[where k="?N"] note diff_shft_smbl = summable_diff[OF ign[OF allf_summable[OF x_in_I]] ign[OF allf_summable[OF x0_in_I]]] note div_shft_smbl = summable_divide[OF diff_shft_smbl] note all_shft_smbl = summable_diff[OF div_smbl ign[OF \summable (f' x0)\]] have 1: "\(\?diff (n + ?N) x\)\ \ L (n + ?N)" for n proof - have "\?diff (n + ?N) x\ \ L (n + ?N) * \(x0 + x) - x0\ / \x\" using divide_right_mono[OF L_def[OF x_in_I x0_in_I] abs_ge_zero] by (simp only: abs_divide) with \x \ 0\ show ?thesis by auto qed note 2 = summable_rabs_comparison_test[OF _ ign[OF \summable L\]] from 1 have "\ \ i. ?diff (i + ?N) x \ \ (\ i. L (i + ?N))" by (metis (lifting) abs_idempotent order_trans[OF summable_rabs[OF 2] suminf_le[OF _ 2 ign[OF \summable L\]]]) then have "\\i. ?diff (i + ?N) x\ \ r / 3" (is "?L_part \ r/3") using L_estimate by auto have "\\n \ (\n?diff n x - f' x0 n\)" .. also have "\ < (\n {..< ?N}" have "\x\ < S" using \\x\ < S\ . also have "S \ S'" using \S \ S'\ . also have "S' \ ?s n" unfolding S'_def proof (rule Min_le_iff[THEN iffD2]) have "?s n \ (?s ` {.. ?s n \ ?s n" using \n \ {..< ?N}\ by auto then show "\ a \ (?s ` {.. ?s n" by blast qed auto finally have "\x\ < ?s n" . from DERIV_D[OF DERIV_f[where n=n], THEN LIM_D, OF \0 < ?r\, unfolded real_norm_def diff_0_right, unfolded some_eq_ex[symmetric], THEN conjunct2] have "\x. x \ 0 \ \x\ < ?s n \ \?diff n x - f' x0 n\ < ?r" . with \x \ 0\ and \\x\ < ?s n\ show "\?diff n x - f' x0 n\ < ?r" by blast qed auto also have "\ = of_nat (card {.. = real ?N * ?r" by simp also have "\ = r/3" by (auto simp del: of_nat_Suc) finally have "\\n < r / 3" (is "?diff_part < r / 3") . from suminf_diff[OF allf_summable[OF x_in_I] allf_summable[OF x0_in_I]] have "\(suminf (f (x0 + x)) - (suminf (f x0))) / x - suminf (f' x0)\ = \\n. ?diff n x - f' x0 n\" unfolding suminf_diff[OF div_smbl \summable (f' x0)\, symmetric] using suminf_divide[OF diff_smbl, symmetric] by auto also have "\ \ ?diff_part + \(\n. ?diff (n + ?N) x) - (\ n. f' x0 (n + ?N))\" unfolding suminf_split_initial_segment[OF all_smbl, where k="?N"] unfolding suminf_diff[OF div_shft_smbl ign[OF \summable (f' x0)\]] apply (simp only: add.commute) using abs_triangle_ineq by blast also have "\ \ ?diff_part + ?L_part + ?f'_part" using abs_triangle_ineq4 by auto also have "\ < r /3 + r/3 + r/3" using \?diff_part < r/3\ \?L_part \ r/3\ and \?f'_part < r/3\ by (rule add_strict_mono [OF add_less_le_mono]) finally show ?thesis by auto qed then show "\s > 0. \ x. x \ 0 \ norm (x - 0) < s \ norm (((\n. f (x0 + x) n) - (\n. f x0 n)) / x - (\n. f' x0 n)) < r" using \0 < S\ by auto qed lemma DERIV_power_series': fixes f :: "nat \ real" assumes converges: "\x. x \ {-R <..< R} \ summable (\n. f n * real (Suc n) * x^n)" and x0_in_I: "x0 \ {-R <..< R}" and "0 < R" shows "DERIV (\x. (\n. f n * x^(Suc n))) x0 :> (\n. f n * real (Suc n) * x0^n)" (is "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)") proof - have for_subinterval: "DERIV (\x. suminf (?f x)) x0 :> suminf (?f' x0)" if "0 < R'" and "R' < R" and "-R' < x0" and "x0 < R'" for R' proof - from that have "x0 \ {-R' <..< R'}" and "R' \ {-R <..< R}" and "x0 \ {-R <..< R}" by auto show ?thesis proof (rule DERIV_series') show "summable (\ n. \f n * real (Suc n) * R'^n\)" proof - have "(R' + R) / 2 < R" and "0 < (R' + R) / 2" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) then have in_Rball: "(R' + R) / 2 \ {-R <..< R}" using \R' < R\ by auto have "norm R' < norm ((R' + R) / 2)" using \0 < R'\ \0 < R\ \R' < R\ by (auto simp: field_simps) from powser_insidea[OF converges[OF in_Rball] this] show ?thesis by auto qed next fix n x y assume "x \ {-R' <..< R'}" and "y \ {-R' <..< R'}" show "\?f x n - ?f y n\ \ \f n * real (Suc n) * R'^n\ * \x-y\" proof - have "\f n * x ^ (Suc n) - f n * y ^ (Suc n)\ = (\f n\ * \x-y\) * \\p" unfolding right_diff_distrib[symmetric] diff_power_eq_sum abs_mult by auto also have "\ \ (\f n\ * \x-y\) * (\real (Suc n)\ * \R' ^ n\)" proof (rule mult_left_mono) have "\\p \ (\px ^ p * y ^ (n - p)\)" by (rule sum_abs) also have "\ \ (\p {.. n" by auto have "\x^n\ \ R'^n" if "x \ {-R'<..x\ \ R'" by auto then show ?thesis unfolding power_abs by (rule power_mono) auto qed from mult_mono[OF this[OF \x \ {-R'<.., of p] this[OF \y \ {-R'<.., of "n-p"]] and \0 < R'\ have "\x^p * y^(n - p)\ \ R'^p * R'^(n - p)" unfolding abs_mult by auto then show "\x^p * y^(n - p)\ \ R'^n" unfolding power_add[symmetric] using \p \ n\ by auto qed also have "\ = real (Suc n) * R' ^ n" unfolding sum_constant card_atLeastLessThan by auto finally show "\\p \ \real (Suc n)\ * \R' ^ n\" unfolding abs_of_nonneg[OF zero_le_power[OF less_imp_le[OF \0 < R'\]]] by linarith show "0 \ \f n\ * \x - y\" unfolding abs_mult[symmetric] by auto qed also have "\ = \f n * real (Suc n) * R' ^ n\ * \x - y\" unfolding abs_mult mult.assoc[symmetric] by algebra finally show ?thesis . qed next show "DERIV (\x. ?f x n) x0 :> ?f' x0 n" for n by (auto intro!: derivative_eq_intros simp del: power_Suc) next fix x assume "x \ {-R' <..< R'}" then have "R' \ {-R <..< R}" and "norm x < norm R'" using assms \R' < R\ by auto have "summable (\n. f n * x^n)" proof (rule summable_comparison_test, intro exI allI impI) fix n have le: "\f n\ * 1 \ \f n\ * real (Suc n)" by (rule mult_left_mono) auto show "norm (f n * x^n) \ norm (f n * real (Suc n) * x^n)" unfolding real_norm_def abs_mult using le mult_right_mono by fastforce qed (rule powser_insidea[OF converges[OF \R' \ {-R <..< R}\] \norm x < norm R'\]) from this[THEN summable_mult2[where c=x], simplified mult.assoc, simplified mult.commute] show "summable (?f x)" by auto next show "summable (?f' x0)" using converges[OF \x0 \ {-R <..< R}\] . show "x0 \ {-R' <..< R'}" using \x0 \ {-R' <..< R'}\ . qed qed let ?R = "(R + \x0\) / 2" have "\x0\ < ?R" using assms by (auto simp: field_simps) then have "- ?R < x0" proof (cases "x0 < 0") case True then have "- x0 < ?R" using \\x0\ < ?R\ by auto then show ?thesis unfolding neg_less_iff_less[symmetric, of "- x0"] by auto next case False have "- ?R < 0" using assms by auto also have "\ \ x0" using False by auto finally show ?thesis . qed then have "0 < ?R" "?R < R" "- ?R < x0" and "x0 < ?R" using assms by (auto simp: field_simps) from for_subinterval[OF this] show ?thesis . qed lemma geometric_deriv_sums: fixes z :: "'a :: {real_normed_field,banach}" assumes "norm z < 1" shows "(\n. of_nat (Suc n) * z ^ n) sums (1 / (1 - z)^2)" proof - have "(\n. diffs (\n. 1) n * z^n) sums (1 / (1 - z)^2)" proof (rule termdiffs_sums_strong) fix z :: 'a assume "norm z < 1" thus "(\n. 1 * z^n) sums (1 / (1 - z))" by (simp add: geometric_sums) qed (insert assms, auto intro!: derivative_eq_intros simp: power2_eq_square) thus ?thesis unfolding diffs_def by simp qed lemma isCont_pochhammer [continuous_intros]: "isCont (\z. pochhammer z n) z" for z :: "'a::real_normed_field" by (induct n) (auto simp: pochhammer_rec') lemma continuous_on_pochhammer [continuous_intros]: "continuous_on A (\z. pochhammer z n)" for A :: "'a::real_normed_field set" by (intro continuous_at_imp_continuous_on ballI isCont_pochhammer) lemmas continuous_on_pochhammer' [continuous_intros] = continuous_on_compose2[OF continuous_on_pochhammer _ subset_UNIV] subsection \Exponential Function\ definition exp :: "'a \ 'a::{real_normed_algebra_1,banach}" where "exp = (\x. \n. x^n /\<^sub>R fact n)" lemma summable_exp_generic: fixes x :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \n. x^n /\<^sub>R fact n" shows "summable S" proof - have S_Suc: "\n. S (Suc n) = (x * S n) /\<^sub>R (Suc n)" unfolding S_def by (simp del: mult_Suc) obtain r :: real where r0: "0 < r" and r1: "r < 1" using dense [OF zero_less_one] by fast obtain N :: nat where N: "norm x < real N * r" using ex_less_of_nat_mult r0 by auto from r1 show ?thesis proof (rule summable_ratio_test [rule_format]) fix n :: nat assume n: "N \ n" have "norm x \ real N * r" using N by (rule order_less_imp_le) also have "real N * r \ real (Suc n) * r" using r0 n by (simp add: mult_right_mono) finally have "norm x * norm (S n) \ real (Suc n) * r * norm (S n)" using norm_ge_zero by (rule mult_right_mono) then have "norm (x * S n) \ real (Suc n) * r * norm (S n)" by (rule order_trans [OF norm_mult_ineq]) then have "norm (x * S n) / real (Suc n) \ r * norm (S n)" by (simp add: pos_divide_le_eq ac_simps) then show "norm (S (Suc n)) \ r * norm (S n)" by (simp add: S_Suc inverse_eq_divide) qed qed lemma summable_norm_exp: "summable (\n. norm (x^n /\<^sub>R fact n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_norm_comparison_test [OF exI, rule_format]) show "summable (\n. norm x^n /\<^sub>R fact n)" by (rule summable_exp_generic) show "norm (x^n /\<^sub>R fact n) \ norm x^n /\<^sub>R fact n" for n by (simp add: norm_power_ineq) qed lemma summable_exp: "summable (\n. inverse (fact n) * x^n)" for x :: "'a::{real_normed_field,banach}" using summable_exp_generic [where x=x] by (simp add: scaleR_conv_of_real nonzero_of_real_inverse) lemma exp_converges: "(\n. x^n /\<^sub>R fact n) sums exp x" unfolding exp_def by (rule summable_exp_generic [THEN summable_sums]) lemma exp_fdiffs: "diffs (\n. inverse (fact n)) = (\n. inverse (fact n :: 'a::{real_normed_field,banach}))" by (simp add: diffs_def mult_ac nonzero_inverse_mult_distrib nonzero_of_real_inverse del: mult_Suc of_nat_Suc) lemma diffs_of_real: "diffs (\n. of_real (f n)) = (\n. of_real (diffs f n))" by (simp add: diffs_def) lemma DERIV_exp [simp]: "DERIV exp x :> exp x" unfolding exp_def scaleR_conv_of_real proof (rule DERIV_cong) have sinv: "summable (\n. of_real (inverse (fact n)) * x ^ n)" for x::'a by (rule exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real]) note xx = exp_converges [THEN sums_summable, unfolded scaleR_conv_of_real] show "((\x. \n. of_real (inverse (fact n)) * x ^ n) has_field_derivative (\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n)) (at x)" by (rule termdiffs [where K="of_real (1 + norm x)"]) (simp_all only: diffs_of_real exp_fdiffs sinv norm_of_real) show "(\n. diffs (\n. of_real (inverse (fact n))) n * x ^ n) = (\n. of_real (inverse (fact n)) * x ^ n)" by (simp add: diffs_of_real exp_fdiffs) qed declare DERIV_exp[THEN DERIV_chain2, derivative_intros] and DERIV_exp[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_exp[derivative_intros] = DERIV_exp[THEN DERIV_compose_FDERIV] lemma norm_exp: "norm (exp x) \ exp (norm x)" proof - from summable_norm[OF summable_norm_exp, of x] have "norm (exp x) \ (\n. inverse (fact n) * norm (x^n))" by (simp add: exp_def) also have "\ \ exp (norm x)" using summable_exp_generic[of "norm x"] summable_norm_exp[of x] by (auto simp: exp_def intro!: suminf_le norm_power_ineq) finally show ?thesis . qed lemma isCont_exp: "isCont exp x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_exp [THEN DERIV_isCont]) lemma isCont_exp' [simp]: "isCont f a \ isCont (\x. exp (f x)) a" for f :: "_ \'a::{real_normed_field,banach}" by (rule isCont_o2 [OF _ isCont_exp]) lemma tendsto_exp [tendsto_intros]: "(f \ a) F \ ((\x. exp (f x)) \ exp a) F" for f:: "_ \'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_exp]) lemma continuous_exp [continuous_intros]: "continuous F f \ continuous F (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_exp) lemma continuous_on_exp [continuous_intros]: "continuous_on s f \ continuous_on s (\x. exp (f x))" for f :: "_ \'a::{real_normed_field,banach}" unfolding continuous_on_def by (auto intro: tendsto_exp) subsubsection \Properties of the Exponential Function\ lemma exp_zero [simp]: "exp 0 = 1" unfolding exp_def by (simp add: scaleR_conv_of_real) lemma exp_series_add_commuting: fixes x y :: "'a::{real_normed_algebra_1,banach}" defines S_def: "S \ \x n. x^n /\<^sub>R fact n" assumes comm: "x * y = y * x" shows "S (x + y) n = (\i\n. S x i * S y (n - i))" proof (induct n) case 0 show ?case unfolding S_def by simp next case (Suc n) have S_Suc: "\x n. S x (Suc n) = (x * S x n) /\<^sub>R real (Suc n)" unfolding S_def by (simp del: mult_Suc) then have times_S: "\x n. x * S x n = real (Suc n) *\<^sub>R S x (Suc n)" by simp have S_comm: "\n. S x n * y = y * S x n" by (simp add: power_commuting_commutes comm S_def) have "real (Suc n) *\<^sub>R S (x + y) (Suc n) = (x + y) * (\i\n. S x i * S y (n - i))" by (metis Suc.hyps times_S) also have "\ = x * (\i\n. S x i * S y (n - i)) + y * (\i\n. S x i * S y (n - i))" by (rule distrib_right) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * y * S y (n - i))" by (simp add: sum_distrib_left ac_simps S_comm) also have "\ = (\i\n. x * S x i * S y (n - i)) + (\i\n. S x i * (y * S y (n - i)))" by (simp add: ac_simps) also have "\ = (\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) + (\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp add: times_S Suc_diff_le) also have "(\i\n. real (Suc i) *\<^sub>R (S x (Suc i) * S y (n - i))) = (\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i)))" by (subst sum.atMost_Suc_shift) simp also have "(\i\n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i)))" by simp also have "(\i\Suc n. real i *\<^sub>R (S x i * S y (Suc n - i))) + (\i\Suc n. real (Suc n - i) *\<^sub>R (S x i * S y (Suc n - i))) = (\i\Suc n. real (Suc n) *\<^sub>R (S x i * S y (Suc n - i)))" by (simp flip: sum.distrib scaleR_add_left of_nat_add) also have "\ = real (Suc n) *\<^sub>R (\i\Suc n. S x i * S y (Suc n - i))" by (simp only: scaleR_right.sum) finally show "S (x + y) (Suc n) = (\i\Suc n. S x i * S y (Suc n - i))" by (simp del: sum.cl_ivl_Suc) qed lemma exp_add_commuting: "x * y = y * x \ exp (x + y) = exp x * exp y" by (simp only: exp_def Cauchy_product summable_norm_exp exp_series_add_commuting) lemma exp_times_arg_commute: "exp A * A = A * exp A" by (simp add: exp_def suminf_mult[symmetric] summable_exp_generic power_commutes suminf_mult2) lemma exp_add: "exp (x + y) = exp x * exp y" for x y :: "'a::{real_normed_field,banach}" by (rule exp_add_commuting) (simp add: ac_simps) lemma exp_double: "exp(2 * z) = exp z ^ 2" by (simp add: exp_add_commuting mult_2 power2_eq_square) lemmas mult_exp_exp = exp_add [symmetric] lemma exp_of_real: "exp (of_real x) = of_real (exp x)" unfolding exp_def apply (subst suminf_of_real [OF summable_exp_generic]) apply (simp add: scaleR_conv_of_real) done lemmas of_real_exp = exp_of_real[symmetric] corollary exp_in_Reals [simp]: "z \ \ \ exp z \ \" by (metis Reals_cases Reals_of_real exp_of_real) lemma exp_not_eq_zero [simp]: "exp x \ 0" proof have "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) also assume "exp x = 0" finally show False by simp qed lemma exp_minus_inverse: "exp x * exp (- x) = 1" by (simp add: exp_add_commuting[symmetric]) lemma exp_minus: "exp (- x) = inverse (exp x)" for x :: "'a::{real_normed_field,banach}" by (intro inverse_unique [symmetric] exp_minus_inverse) lemma exp_diff: "exp (x - y) = exp x / exp y" for x :: "'a::{real_normed_field,banach}" using exp_add [of x "- y"] by (simp add: exp_minus divide_inverse) lemma exp_of_nat_mult: "exp (of_nat n * x) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (induct n) (auto simp: distrib_left exp_add mult.commute) corollary exp_of_nat2_mult: "exp (x * of_nat n) = exp x ^ n" for x :: "'a::{real_normed_field,banach}" by (metis exp_of_nat_mult mult_of_nat_commute) lemma exp_sum: "finite I \ exp (sum f I) = prod (\x. exp (f x)) I" by (induct I rule: finite_induct) (auto simp: exp_add_commuting mult.commute) lemma exp_divide_power_eq: fixes x :: "'a::{real_normed_field,banach}" assumes "n > 0" shows "exp (x / of_nat n) ^ n = exp x" using assms proof (induction n arbitrary: x) case (Suc n) show ?case proof (cases "n = 0") case True then show ?thesis by simp next case False have [simp]: "1 + (of_nat n * of_nat n + of_nat n * 2) \ (0::'a)" using of_nat_eq_iff [of "1 + n * n + n * 2" "0"] by simp from False have [simp]: "x * of_nat n / (1 + of_nat n) / of_nat n = x / (1 + of_nat n)" by simp have [simp]: "x / (1 + of_nat n) + x * of_nat n / (1 + of_nat n) = x" using of_nat_neq_0 by (auto simp add: field_split_simps) show ?thesis using Suc.IH [of "x * of_nat n / (1 + of_nat n)"] False by (simp add: exp_add [symmetric]) qed qed simp subsubsection \Properties of the Exponential Function on Reals\ text \Comparisons of \<^term>\exp x\ with zero.\ text \Proof: because every exponential can be seen as a square.\ lemma exp_ge_zero [simp]: "0 \ exp x" for x :: real proof - have "0 \ exp (x/2) * exp (x/2)" by simp then show ?thesis by (simp add: exp_add [symmetric]) qed lemma exp_gt_zero [simp]: "0 < exp x" for x :: real by (simp add: order_less_le) lemma not_exp_less_zero [simp]: "\ exp x < 0" for x :: real by (simp add: not_less) lemma not_exp_le_zero [simp]: "\ exp x \ 0" for x :: real by (simp add: not_le) lemma abs_exp_cancel [simp]: "\exp x\ = exp x" for x :: real by simp text \Strict monotonicity of exponential.\ lemma exp_ge_add_one_self_aux: fixes x :: real assumes "0 \ x" shows "1 + x \ exp x" using order_le_imp_less_or_eq [OF assms] proof assume "0 < x" have "1 + x \ (\n<2. inverse (fact n) * x^n)" by (auto simp: numeral_2_eq_2) also have "\ \ (\n. inverse (fact n) * x^n)" using \0 < x\ by (auto simp add: zero_le_mult_iff intro: sum_le_suminf [OF summable_exp]) finally show "1 + x \ exp x" by (simp add: exp_def) qed auto lemma exp_gt_one: "0 < x \ 1 < exp x" for x :: real proof - assume x: "0 < x" then have "1 < 1 + x" by simp also from x have "1 + x \ exp x" by (simp add: exp_ge_add_one_self_aux) finally show ?thesis . qed lemma exp_less_mono: fixes x y :: real assumes "x < y" shows "exp x < exp y" proof - from \x < y\ have "0 < y - x" by simp then have "1 < exp (y - x)" by (rule exp_gt_one) then have "1 < exp y / exp x" by (simp only: exp_diff) then show "exp x < exp y" by simp qed lemma exp_less_cancel: "exp x < exp y \ x < y" for x y :: real unfolding linorder_not_le [symmetric] by (auto simp: order_le_less exp_less_mono) lemma exp_less_cancel_iff [iff]: "exp x < exp y \ x < y" for x y :: real by (auto intro: exp_less_mono exp_less_cancel) lemma exp_le_cancel_iff [iff]: "exp x \ exp y \ x \ y" for x y :: real by (auto simp: linorder_not_less [symmetric]) lemma exp_inj_iff [iff]: "exp x = exp y \ x = y" for x y :: real by (simp add: order_eq_iff) text \Comparisons of \<^term>\exp x\ with one.\ lemma one_less_exp_iff [simp]: "1 < exp x \ 0 < x" for x :: real using exp_less_cancel_iff [where x = 0 and y = x] by simp lemma exp_less_one_iff [simp]: "exp x < 1 \ x < 0" for x :: real using exp_less_cancel_iff [where x = x and y = 0] by simp lemma one_le_exp_iff [simp]: "1 \ exp x \ 0 \ x" for x :: real using exp_le_cancel_iff [where x = 0 and y = x] by simp lemma exp_le_one_iff [simp]: "exp x \ 1 \ x \ 0" for x :: real using exp_le_cancel_iff [where x = x and y = 0] by simp lemma exp_eq_one_iff [simp]: "exp x = 1 \ x = 0" for x :: real using exp_inj_iff [where x = x and y = 0] by simp lemma lemma_exp_total: "1 \ y \ \x. 0 \ x \ x \ y - 1 \ exp x = y" for y :: real proof (rule IVT) assume "1 \ y" then have "0 \ y - 1" by simp then have "1 + (y - 1) \ exp (y - 1)" by (rule exp_ge_add_one_self_aux) then show "y \ exp (y - 1)" by simp qed (simp_all add: le_diff_eq) lemma exp_total: "0 < y \ \x. exp x = y" for y :: real proof (rule linorder_le_cases [of 1 y]) assume "1 \ y" then show "\x. exp x = y" by (fast dest: lemma_exp_total) next assume "0 < y" and "y \ 1" then have "1 \ inverse y" by (simp add: one_le_inverse_iff) then obtain x where "exp x = inverse y" by (fast dest: lemma_exp_total) then have "exp (- x) = y" by (simp add: exp_minus) then show "\x. exp x = y" .. qed subsection \Natural Logarithm\ class ln = real_normed_algebra_1 + banach + fixes ln :: "'a \ 'a" assumes ln_one [simp]: "ln 1 = 0" definition powr :: "'a \ 'a \ 'a::ln" (infixr "powr" 80) \ \exponentation via ln and exp\ where "x powr a \ if x = 0 then 0 else exp (a * ln x)" lemma powr_0 [simp]: "0 powr z = 0" by (simp add: powr_def) instantiation real :: ln begin definition ln_real :: "real \ real" where "ln_real x = (THE u. exp u = x)" instance by intro_classes (simp add: ln_real_def) end lemma powr_eq_0_iff [simp]: "w powr z = 0 \ w = 0" by (simp add: powr_def) lemma ln_exp [simp]: "ln (exp x) = x" for x :: real by (simp add: ln_real_def) lemma exp_ln [simp]: "0 < x \ exp (ln x) = x" for x :: real by (auto dest: exp_total) lemma exp_ln_iff [simp]: "exp (ln x) = x \ 0 < x" for x :: real by (metis exp_gt_zero exp_ln) lemma ln_unique: "exp y = x \ ln x = y" for x :: real by (erule subst) (rule ln_exp) lemma ln_mult: "0 < x \ 0 < y \ ln (x * y) = ln x + ln y" for x :: real by (rule ln_unique) (simp add: exp_add) lemma ln_prod: "finite I \ (\i. i \ I \ f i > 0) \ ln (prod f I) = sum (\x. ln(f x)) I" for f :: "'a \ real" by (induct I rule: finite_induct) (auto simp: ln_mult prod_pos) lemma ln_inverse: "0 < x \ ln (inverse x) = - ln x" for x :: real by (rule ln_unique) (simp add: exp_minus) lemma ln_div: "0 < x \ 0 < y \ ln (x / y) = ln x - ln y" for x :: real by (rule ln_unique) (simp add: exp_diff) lemma ln_realpow: "0 < x \ ln (x^n) = real n * ln x" by (rule ln_unique) (simp add: exp_of_nat_mult) lemma ln_less_cancel_iff [simp]: "0 < x \ 0 < y \ ln x < ln y \ x < y" for x :: real by (subst exp_less_cancel_iff [symmetric]) simp lemma ln_le_cancel_iff [simp]: "0 < x \ 0 < y \ ln x \ ln y \ x \ y" for x :: real by (simp add: linorder_not_less [symmetric]) lemma ln_inj_iff [simp]: "0 < x \ 0 < y \ ln x = ln y \ x = y" for x :: real by (simp add: order_eq_iff) lemma ln_add_one_self_le_self: "0 \ x \ ln (1 + x) \ x" for x :: real by (rule exp_le_cancel_iff [THEN iffD1]) (simp add: exp_ge_add_one_self_aux) lemma ln_less_self [simp]: "0 < x \ ln x < x" for x :: real by (rule order_less_le_trans [where y = "ln (1 + x)"]) (simp_all add: ln_add_one_self_le_self) lemma ln_ge_iff: "\x::real. 0 < x \ y \ ln x \ exp y \ x" using exp_le_cancel_iff exp_total by force lemma ln_ge_zero [simp]: "1 \ x \ 0 \ ln x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_imp_ge_one: "0 \ ln x \ 0 < x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_ge_zero_iff [simp]: "0 < x \ 0 \ ln x \ 1 \ x" for x :: real using ln_le_cancel_iff [of 1 x] by simp lemma ln_less_zero_iff [simp]: "0 < x \ ln x < 0 \ x < 1" for x :: real using ln_less_cancel_iff [of x 1] by simp lemma ln_le_zero_iff [simp]: "0 < x \ ln x \ 0 \ x \ 1" for x :: real by (metis less_numeral_extra(1) ln_le_cancel_iff ln_one) lemma ln_gt_zero: "1 < x \ 0 < ln x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_imp_gt_one: "0 < ln x \ 0 < x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_gt_zero_iff [simp]: "0 < x \ 0 < ln x \ 1 < x" for x :: real using ln_less_cancel_iff [of 1 x] by simp lemma ln_eq_zero_iff [simp]: "0 < x \ ln x = 0 \ x = 1" for x :: real using ln_inj_iff [of x 1] by simp lemma ln_less_zero: "0 < x \ x < 1 \ ln x < 0" for x :: real by simp lemma ln_neg_is_const: "x \ 0 \ ln x = (THE x. False)" for x :: real by (auto simp: ln_real_def intro!: arg_cong[where f = The]) lemma powr_eq_one_iff [simp]: "a powr x = 1 \ x = 0" if "a > 1" for a x :: real using that by (auto simp: powr_def split: if_splits) lemma isCont_ln: fixes x :: real assumes "x \ 0" shows "isCont ln x" proof (cases "0 < x") case True then have "isCont ln (exp (ln x))" by (intro isCont_inverse_function[where d = "\x\" and f = exp]) auto with True show ?thesis by simp next case False with \x \ 0\ show "isCont ln x" unfolding isCont_def by (subst filterlim_cong[OF _ refl, of _ "nhds (ln 0)" _ "\_. ln 0"]) (auto simp: ln_neg_is_const not_less eventually_at dist_real_def intro!: exI[of _ "\x\"]) qed lemma tendsto_ln [tendsto_intros]: "(f \ a) F \ a \ 0 \ ((\x. ln (f x)) \ ln a) F" for a :: real by (rule isCont_tendsto_compose [OF isCont_ln]) lemma continuous_ln: "continuous F f \ f (Lim F (\x. x)) \ 0 \ continuous F (\x. ln (f x :: real))" unfolding continuous_def by (rule tendsto_ln) lemma isCont_ln' [continuous_intros]: "continuous (at x) f \ f x \ 0 \ continuous (at x) (\x. ln (f x :: real))" unfolding continuous_at by (rule tendsto_ln) lemma continuous_within_ln [continuous_intros]: "continuous (at x within s) f \ f x \ 0 \ continuous (at x within s) (\x. ln (f x :: real))" unfolding continuous_within by (rule tendsto_ln) lemma continuous_on_ln [continuous_intros]: "continuous_on s f \ (\x\s. f x \ 0) \ continuous_on s (\x. ln (f x :: real))" unfolding continuous_on_def by (auto intro: tendsto_ln) lemma DERIV_ln: "0 < x \ DERIV ln x :> inverse x" for x :: real by (rule DERIV_inverse_function [where f=exp and a=0 and b="x+1"]) (auto intro: DERIV_cong [OF DERIV_exp exp_ln] isCont_ln) lemma DERIV_ln_divide: "0 < x \ DERIV ln x :> 1 / x" for x :: real by (rule DERIV_ln[THEN DERIV_cong]) (simp_all add: divide_inverse) declare DERIV_ln_divide[THEN DERIV_chain2, derivative_intros] and DERIV_ln_divide[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_ln[derivative_intros] = DERIV_ln[THEN DERIV_compose_FDERIV] lemma ln_series: assumes "0 < x" and "x < 2" shows "ln x = (\ n. (-1)^n * (1 / real (n + 1)) * (x - 1)^(Suc n))" (is "ln x = suminf (?f (x - 1))") proof - let ?f' = "\x n. (-1)^n * (x - 1)^n" have "ln x - suminf (?f (x - 1)) = ln 1 - suminf (?f (1 - 1))" proof (rule DERIV_isconst3 [where x = x]) fix x :: real assume "x \ {0 <..< 2}" then have "0 < x" and "x < 2" by auto have "norm (1 - x) < 1" using \0 < x\ and \x < 2\ by auto have "1 / x = 1 / (1 - (1 - x))" by auto also have "\ = (\ n. (1 - x)^n)" using geometric_sums[OF \norm (1 - x) < 1\] by (rule sums_unique) also have "\ = suminf (?f' x)" unfolding power_mult_distrib[symmetric] by (rule arg_cong[where f=suminf], rule arg_cong[where f="(^)"], auto) finally have "DERIV ln x :> suminf (?f' x)" using DERIV_ln[OF \0 < x\] unfolding divide_inverse by auto moreover have repos: "\ h x :: real. h - 1 + x = h + x - 1" by auto have "DERIV (\x. suminf (?f x)) (x - 1) :> (\n. (-1)^n * (1 / real (n + 1)) * real (Suc n) * (x - 1) ^ n)" proof (rule DERIV_power_series') show "x - 1 \ {- 1<..<1}" and "(0 :: real) < 1" using \0 < x\ \x < 2\ by auto next fix x :: real assume "x \ {- 1<..<1}" then show "summable (\n. (- 1) ^ n * (1 / real (n + 1)) * real (Suc n) * x^n)" by (simp add: abs_if flip: power_mult_distrib) qed then have "DERIV (\x. suminf (?f x)) (x - 1) :> suminf (?f' x)" unfolding One_nat_def by auto then have "DERIV (\x. suminf (?f (x - 1))) x :> suminf (?f' x)" unfolding DERIV_def repos . ultimately have "DERIV (\x. ln x - suminf (?f (x - 1))) x :> suminf (?f' x) - suminf (?f' x)" by (rule DERIV_diff) then show "DERIV (\x. ln x - suminf (?f (x - 1))) x :> 0" by auto qed (auto simp: assms) then show ?thesis by auto qed lemma exp_first_terms: fixes x :: "'a::{real_normed_algebra_1,banach}" shows "exp x = (\nR (x ^ n)) + (\n. inverse(fact (n + k)) *\<^sub>R (x ^ (n + k)))" proof - have "exp x = suminf (\n. inverse(fact n) *\<^sub>R (x^n))" by (simp add: exp_def) also from summable_exp_generic have "\ = (\ n. inverse(fact(n+k)) *\<^sub>R (x ^ (n + k))) + (\ n::natR (x^n))" (is "_ = _ + ?a") by (rule suminf_split_initial_segment) finally show ?thesis by simp qed lemma exp_first_term: "exp x = 1 + (\n. inverse (fact (Suc n)) *\<^sub>R (x ^ Suc n))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 1] by simp lemma exp_first_two_terms: "exp x = 1 + x + (\n. inverse (fact (n + 2)) *\<^sub>R (x ^ (n + 2)))" for x :: "'a::{real_normed_algebra_1,banach}" using exp_first_terms[of x 2] by (simp add: eval_nat_numeral) lemma exp_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "exp x \ 1 + x + x\<^sup>2" proof - have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ x\<^sup>2" proof - have "(\n. x\<^sup>2 / 2 * (1 / 2) ^ n) sums (x\<^sup>2 / 2 * (1 / (1 - 1 / 2)))" by (intro sums_mult geometric_sums) simp then have sumsx: "(\n. x\<^sup>2 / 2 * (1 / 2) ^ n) sums x\<^sup>2" by simp have "suminf (\n. inverse(fact (n+2)) * (x ^ (n + 2))) \ suminf (\n. (x\<^sup>2/2) * ((1/2)^n))" proof (intro suminf_le allI) show "inverse (fact (n + 2)) * x ^ (n + 2) \ (x\<^sup>2/2) * ((1/2)^n)" for n :: nat proof - have "(2::nat) * 2 ^ n \ fact (n + 2)" by (induct n) simp_all then have "real ((2::nat) * 2 ^ n) \ real_of_nat (fact (n + 2))" by (simp only: of_nat_le_iff) then have "((2::real) * 2 ^ n) \ fact (n + 2)" unfolding of_nat_fact by simp then have "inverse (fact (n + 2)) \ inverse ((2::real) * 2 ^ n)" by (rule le_imp_inverse_le) simp then have "inverse (fact (n + 2)) \ 1/(2::real) * (1/2)^n" by (simp add: power_inverse [symmetric]) then have "inverse (fact (n + 2)) * (x^n * x\<^sup>2) \ 1/2 * (1/2)^n * (1 * x\<^sup>2)" by (rule mult_mono) (rule mult_mono, simp_all add: power_le_one a b) then show ?thesis unfolding power_add by (simp add: ac_simps del: fact_Suc) qed show "summable (\n. inverse (fact (n + 2)) * x ^ (n + 2))" by (rule summable_exp [THEN summable_ignore_initial_segment]) show "summable (\n. x\<^sup>2 / 2 * (1 / 2) ^ n)" by (rule sums_summable [OF sumsx]) qed also have "\ = x\<^sup>2" by (rule sums_unique [THEN sym]) (rule sumsx) finally show ?thesis . qed then show ?thesis unfolding exp_first_two_terms by auto qed corollary exp_half_le2: "exp(1/2) \ (2::real)" using exp_bound [of "1/2"] by (simp add: field_simps) corollary exp_le: "exp 1 \ (3::real)" using exp_bound [of 1] by (simp add: field_simps) lemma exp_bound_half: "norm z \ 1/2 \ norm (exp z) \ 2" by (blast intro: order_trans intro!: exp_half_le2 norm_exp) lemma exp_bound_lemma: assumes "norm z \ 1/2" shows "norm (exp z) \ 1 + 2 * norm z" proof - have *: "(norm z)\<^sup>2 \ norm z * 1" unfolding power2_eq_square by (rule mult_left_mono) (use assms in auto) have "norm (exp z) \ exp (norm z)" by (rule norm_exp) also have "\ \ 1 + (norm z) + (norm z)\<^sup>2" using assms exp_bound by auto also have "\ \ 1 + 2 * norm z" using * by auto finally show ?thesis . qed lemma real_exp_bound_lemma: "0 \ x \ x \ 1/2 \ exp x \ 1 + 2 * x" for x :: real using exp_bound_lemma [of x] by simp lemma ln_one_minus_pos_upper_bound: fixes x :: real assumes a: "0 \ x" and b: "x < 1" shows "ln (1 - x) \ - x" proof - have "(1 - x) * (1 + x + x\<^sup>2) = 1 - x^3" by (simp add: algebra_simps power2_eq_square power3_eq_cube) also have "\ \ 1" by (auto simp: a) finally have "(1 - x) * (1 + x + x\<^sup>2) \ 1" . moreover have c: "0 < 1 + x + x\<^sup>2" by (simp add: add_pos_nonneg a) ultimately have "1 - x \ 1 / (1 + x + x\<^sup>2)" by (elim mult_imp_le_div_pos) also have "\ \ 1 / exp x" by (metis a abs_one b exp_bound exp_gt_zero frac_le less_eq_real_def real_sqrt_abs real_sqrt_pow2_iff real_sqrt_power) also have "\ = exp (- x)" by (auto simp: exp_minus divide_inverse) finally have "1 - x \ exp (- x)" . also have "1 - x = exp (ln (1 - x))" by (metis b diff_0 exp_ln_iff less_iff_diff_less_0 minus_diff_eq) finally have "exp (ln (1 - x)) \ exp (- x)" . then show ?thesis by (auto simp only: exp_le_cancel_iff) qed lemma exp_ge_add_one_self [simp]: "1 + x \ exp x" for x :: real proof (cases "0 \ x \ x \ -1") case True then show ?thesis by (meson exp_ge_add_one_self_aux exp_ge_zero order.trans real_add_le_0_iff) next case False then have ln1: "ln (1 + x) \ x" using ln_one_minus_pos_upper_bound [of "-x"] by simp have "1 + x = exp (ln (1 + x))" using False by auto also have "\ \ exp x" by (simp add: ln1) finally show ?thesis . qed lemma ln_one_plus_pos_lower_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1" shows "x - x\<^sup>2 \ ln (1 + x)" proof - have "exp (x - x\<^sup>2) = exp x / exp (x\<^sup>2)" by (rule exp_diff) also have "\ \ (1 + x + x\<^sup>2) / exp (x \<^sup>2)" by (metis a b divide_right_mono exp_bound exp_ge_zero) also have "\ \ (1 + x + x\<^sup>2) / (1 + x\<^sup>2)" by (simp add: a divide_left_mono add_pos_nonneg) also from a have "\ \ 1 + x" by (simp add: field_simps add_strict_increasing zero_le_mult_iff) finally have "exp (x - x\<^sup>2) \ 1 + x" . also have "\ = exp (ln (1 + x))" proof - from a have "0 < 1 + x" by auto then show ?thesis by (auto simp only: exp_ln_iff [THEN sym]) qed finally have "exp (x - x\<^sup>2) \ exp (ln (1 + x))" . then show ?thesis by (metis exp_le_cancel_iff) qed lemma ln_one_minus_pos_lower_bound: fixes x :: real assumes a: "0 \ x" and b: "x \ 1 / 2" shows "- x - 2 * x\<^sup>2 \ ln (1 - x)" proof - from b have c: "x < 1" by auto then have "ln (1 - x) = - ln (1 + x / (1 - x))" by (auto simp: ln_inverse [symmetric] field_simps intro: arg_cong [where f=ln]) also have "- (x / (1 - x)) \ \" proof - have "ln (1 + x / (1 - x)) \ x / (1 - x)" using a c by (intro ln_add_one_self_le_self) auto then show ?thesis by auto qed also have "- (x / (1 - x)) = - x / (1 - x)" by auto finally have d: "- x / (1 - x) \ ln (1 - x)" . have "0 < 1 - x" using a b by simp then have e: "- x - 2 * x\<^sup>2 \ - x / (1 - x)" using mult_right_le_one_le[of "x * x" "2 * x"] a b by (simp add: field_simps power2_eq_square) from e d show "- x - 2 * x\<^sup>2 \ ln (1 - x)" by (rule order_trans) qed lemma ln_add_one_self_le_self2: fixes x :: real shows "-1 < x \ ln (1 + x) \ x" by (metis diff_gt_0_iff_gt diff_minus_eq_add exp_ge_add_one_self exp_le_cancel_iff exp_ln minus_less_iff) lemma abs_ln_one_plus_x_minus_x_bound_nonneg: fixes x :: real assumes x: "0 \ x" and x1: "x \ 1" shows "\ln (1 + x) - x\ \ x\<^sup>2" proof - from x have "ln (1 + x) \ x" by (rule ln_add_one_self_le_self) then have "ln (1 + x) - x \ 0" by simp then have "\ln(1 + x) - x\ = - (ln(1 + x) - x)" by (rule abs_of_nonpos) also have "\ = x - ln (1 + x)" by simp also have "\ \ x\<^sup>2" proof - from x x1 have "x - x\<^sup>2 \ ln (1 + x)" by (intro ln_one_plus_pos_lower_bound) then show ?thesis by simp qed finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound_nonpos: fixes x :: real assumes a: "-(1 / 2) \ x" and b: "x \ 0" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof - have *: "- (-x) - 2 * (-x)\<^sup>2 \ ln (1 - (- x))" by (metis a b diff_zero ln_one_minus_pos_lower_bound minus_diff_eq neg_le_iff_le) have "\ln (1 + x) - x\ = x - ln (1 - (- x))" using a ln_add_one_self_le_self2 [of x] by (simp add: abs_if) also have "\ \ 2 * x\<^sup>2" using * by (simp add: algebra_simps) finally show ?thesis . qed lemma abs_ln_one_plus_x_minus_x_bound: fixes x :: real assumes "\x\ \ 1 / 2" shows "\ln (1 + x) - x\ \ 2 * x\<^sup>2" proof (cases "0 \ x") case True then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonneg assms by fastforce next case False then show ?thesis using abs_ln_one_plus_x_minus_x_bound_nonpos assms by auto qed lemma ln_x_over_x_mono: fixes x :: real assumes x: "exp 1 \ x" "x \ y" shows "ln y / y \ ln x / x" proof - note x moreover have "0 < exp (1::real)" by simp ultimately have a: "0 < x" and b: "0 < y" by (fast intro: less_le_trans order_trans)+ have "x * ln y - x * ln x = x * (ln y - ln x)" by (simp add: algebra_simps) also have "\ = x * ln (y / x)" by (simp only: ln_div a b) also have "y / x = (x + (y - x)) / x" by simp also have "\ = 1 + (y - x) / x" using x a by (simp add: field_simps) also have "x * ln (1 + (y - x) / x) \ x * ((y - x) / x)" using x a by (intro mult_left_mono ln_add_one_self_le_self) simp_all also have "\ = y - x" using a by simp also have "\ = (y - x) * ln (exp 1)" by simp also have "\ \ (y - x) * ln x" using a x exp_total of_nat_1 x(1) by (fastforce intro: mult_left_mono) also have "\ = y * ln x - x * ln x" by (rule left_diff_distrib) finally have "x * ln y \ y * ln x" by arith then have "ln y \ (y * ln x) / x" using a by (simp add: field_simps) also have "\ = y * (ln x / x)" by simp finally show ?thesis using b by (simp add: field_simps) qed lemma ln_le_minus_one: "0 < x \ ln x \ x - 1" for x :: real using exp_ge_add_one_self[of "ln x"] by simp corollary ln_diff_le: "0 < x \ 0 < y \ ln x - ln y \ (x - y) / y" for x :: real by (simp add: ln_div [symmetric] diff_divide_distrib ln_le_minus_one) lemma ln_eq_minus_one: fixes x :: real assumes "0 < x" "ln x = x - 1" shows "x = 1" proof - let ?l = "\y. ln y - y + 1" have D: "\x::real. 0 < x \ DERIV ?l x :> (1 / x - 1)" by (auto intro!: derivative_eq_intros) show ?thesis proof (cases rule: linorder_cases) assume "x < 1" from dense[OF \x < 1\] obtain a where "x < a" "a < 1" by blast from \x < a\ have "?l x < ?l a" proof (rule DERIV_pos_imp_increasing) fix y assume "x \ y" "y \ a" with \0 < x\ \a < 1\ have "0 < 1 / y - 1" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ 0 < z" by blast qed also have "\ \ 0" using ln_le_minus_one \0 < x\ \x < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "1 < x" from dense[OF this] obtain a where "1 < a" "a < x" by blast from \a < x\ have "?l x < ?l a" proof (rule DERIV_neg_imp_decreasing) fix y assume "a \ y" "y \ x" with \1 < a\ have "1 / y - 1 < 0" "0 < y" by (auto simp: field_simps) with D show "\z. DERIV ?l y :> z \ z < 0" by blast qed also have "\ \ 0" using ln_le_minus_one \1 < a\ by (auto simp: field_simps) finally show "x = 1" using assms by auto next assume "x = 1" then show ?thesis by simp qed qed lemma ln_x_over_x_tendsto_0: "((\x::real. ln x / x) \ 0) at_top" proof (rule lhospital_at_top_at_top[where f' = inverse and g' = "\_. 1"]) from eventually_gt_at_top[of "0::real"] show "\\<^sub>F x in at_top. (ln has_real_derivative inverse x) (at x)" by eventually_elim (auto intro!: derivative_eq_intros simp: field_simps) qed (use tendsto_inverse_0 in \auto simp: filterlim_ident dest!: tendsto_mono[OF at_top_le_at_infinity]\) lemma exp_ge_one_plus_x_over_n_power_n: assumes "x \ - real n" "n > 0" shows "(1 + x / of_nat n) ^ n \ exp x" proof (cases "x = - of_nat n") case False from assms False have "(1 + x / of_nat n) ^ n = exp (of_nat n * ln (1 + x / of_nat n))" by (subst exp_of_nat_mult, subst exp_ln) (simp_all add: field_simps) also from assms False have "ln (1 + x / real n) \ x / real n" by (intro ln_add_one_self_le_self2) (simp_all add: field_simps) with assms have "exp (of_nat n * ln (1 + x / of_nat n)) \ exp x" by (simp add: field_simps) finally show ?thesis . next case True then show ?thesis by (simp add: zero_power) qed lemma exp_ge_one_minus_x_over_n_power_n: assumes "x \ real n" "n > 0" shows "(1 - x / of_nat n) ^ n \ exp (-x)" using exp_ge_one_plus_x_over_n_power_n[of n "-x"] assms by simp lemma exp_at_bot: "(exp \ (0::real)) at_bot" unfolding tendsto_Zfun_iff proof (rule ZfunI, simp add: eventually_at_bot_dense) fix r :: real assume "0 < r" have "exp x < r" if "x < ln r" for x by (metis \0 < r\ exp_less_mono exp_ln that) then show "\k. \n at_top" by (rule filterlim_at_top_at_top[where Q="\x. True" and P="\x. 0 < x" and g=ln]) (auto intro: eventually_gt_at_top) lemma lim_exp_minus_1: "((\z::'a. (exp(z) - 1) / z) \ 1) (at 0)" for x :: "'a::{real_normed_field,banach}" proof - have "((\z::'a. exp(z) - 1) has_field_derivative 1) (at 0)" by (intro derivative_eq_intros | simp)+ then show ?thesis by (simp add: Deriv.has_field_derivative_iff) qed lemma ln_at_0: "LIM x at_right 0. ln (x::real) :> at_bot" by (rule filterlim_at_bot_at_right[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto simp: eventually_at_filter) lemma ln_at_top: "LIM x at_top. ln (x::real) :> at_top" by (rule filterlim_at_top_at_top[where Q="\x. 0 < x" and P="\x. True" and g=exp]) (auto intro: eventually_gt_at_top) lemma filtermap_ln_at_top: "filtermap (ln::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of exp] exp_at_top ln_at_top) auto lemma filtermap_exp_at_top: "filtermap (exp::real \ real) at_top = at_top" by (intro filtermap_fun_inverse[of ln] exp_at_top ln_at_top) (auto simp: eventually_at_top_dense) lemma filtermap_ln_at_right: "filtermap ln (at_right (0::real)) = at_bot" by (auto intro!: filtermap_fun_inverse[where g="\x. exp x"] ln_at_0 simp: filterlim_at exp_at_bot) lemma tendsto_power_div_exp_0: "((\x. x ^ k / exp x) \ (0::real)) at_top" proof (induct k) case 0 show "((\x. x ^ 0 / exp x) \ (0::real)) at_top" by (simp add: inverse_eq_divide[symmetric]) (metis filterlim_compose[OF tendsto_inverse_0] exp_at_top filterlim_mono at_top_le_at_infinity order_refl) next case (Suc k) show ?case proof (rule lhospital_at_top_at_top) show "eventually (\x. DERIV (\x. x ^ Suc k) x :> (real (Suc k) * x^k)) at_top" by eventually_elim (intro derivative_eq_intros, auto) show "eventually (\x. DERIV exp x :> exp x) at_top" by eventually_elim auto show "eventually (\x. exp x \ 0) at_top" by auto from tendsto_mult[OF tendsto_const Suc, of "real (Suc k)"] show "((\x. real (Suc k) * x ^ k / exp x) \ 0) at_top" by simp qed (rule exp_at_top) qed subsubsection\ A couple of simple bounds\ lemma exp_plus_inverse_exp: fixes x::real shows "2 \ exp x + inverse (exp x)" proof - have "2 \ exp x + exp (-x)" using exp_ge_add_one_self [of x] exp_ge_add_one_self [of "-x"] by linarith then show ?thesis by (simp add: exp_minus) qed lemma real_le_x_sinh: fixes x::real assumes "0 \ x" shows "x \ (exp x - inverse(exp x)) / 2" proof - have *: "exp a - inverse(exp a) - 2*a \ exp b - inverse(exp b) - 2*b" if "a \ b" for a b::real using exp_plus_inverse_exp by (fastforce intro: derivative_eq_intros DERIV_nonneg_imp_nondecreasing [OF that]) show ?thesis using*[OF assms] by simp qed lemma real_le_abs_sinh: fixes x::real shows "abs x \ abs((exp x - inverse(exp x)) / 2)" proof (cases "0 \ x") case True show ?thesis using real_le_x_sinh [OF True] True by (simp add: abs_if) next case False have "-x \ (exp(-x) - inverse(exp(-x))) / 2" by (meson False linear neg_le_0_iff_le real_le_x_sinh) also have "\ \ \(exp x - inverse (exp x)) / 2\" by (metis (no_types, opaque_lifting) abs_divide abs_le_iff abs_minus_cancel add.inverse_inverse exp_minus minus_diff_eq order_refl) finally show ?thesis using False by linarith qed subsection\The general logarithm\ definition log :: "real \ real \ real" \ \logarithm of \<^term>\x\ to base \<^term>\a\\ where "log a x = ln x / ln a" lemma tendsto_log [tendsto_intros]: "(f \ a) F \ (g \ b) F \ 0 < a \ a \ 1 \ 0 < b \ ((\x. log (f x) (g x)) \ log a b) F" unfolding log_def by (intro tendsto_intros) auto lemma continuous_log: assumes "continuous F f" and "continuous F g" and "0 < f (Lim F (\x. x))" and "f (Lim F (\x. x)) \ 1" and "0 < g (Lim F (\x. x))" shows "continuous F (\x. log (f x) (g x))" using assms unfolding continuous_def by (rule tendsto_log) lemma continuous_at_within_log[continuous_intros]: assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "0 < f a" and "f a \ 1" and "0 < g a" shows "continuous (at a within s) (\x. log (f x) (g x))" using assms unfolding continuous_within by (rule tendsto_log) lemma isCont_log[continuous_intros, simp]: assumes "isCont f a" "isCont g a" "0 < f a" "f a \ 1" "0 < g a" shows "isCont (\x. log (f x) (g x)) a" using assms unfolding continuous_at by (rule tendsto_log) lemma continuous_on_log[continuous_intros]: assumes "continuous_on s f" "continuous_on s g" and "\x\s. 0 < f x" "\x\s. f x \ 1" "\x\s. 0 < g x" shows "continuous_on s (\x. log (f x) (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_log) lemma powr_one_eq_one [simp]: "1 powr a = 1" by (simp add: powr_def) lemma powr_zero_eq_one [simp]: "x powr 0 = (if x = 0 then 0 else 1)" by (simp add: powr_def) lemma powr_one_gt_zero_iff [simp]: "x powr 1 = x \ 0 \ x" for x :: real by (auto simp: powr_def) declare powr_one_gt_zero_iff [THEN iffD2, simp] lemma powr_diff: fixes w:: "'a::{ln,real_normed_field}" shows "w powr (z1 - z2) = w powr z1 / w powr z2" by (simp add: powr_def algebra_simps exp_diff) lemma powr_mult: "0 \ x \ 0 \ y \ (x * y) powr a = (x powr a) * (y powr a)" for a x y :: real by (simp add: powr_def exp_add [symmetric] ln_mult distrib_left) lemma powr_ge_pzero [simp]: "0 \ x powr y" for x y :: real by (simp add: powr_def) lemma powr_non_neg[simp]: "\a powr x < 0" for a x::real using powr_ge_pzero[of a x] by arith lemma inverse_powr: "\y::real. 0 \ y \ inverse y powr a = inverse (y powr a)" by (simp add: exp_minus ln_inverse powr_def) lemma powr_divide: "\0 \ x; 0 \ y\ \ (x / y) powr a = (x powr a) / (y powr a)" for a b x :: real by (simp add: divide_inverse powr_mult inverse_powr) lemma powr_add: "x powr (a + b) = (x powr a) * (x powr b)" for a b x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_add [symmetric] distrib_right) lemma powr_mult_base: "0 \ x \x * x powr y = x powr (1 + y)" for x :: real by (auto simp: powr_add) lemma powr_powr: "(x powr a) powr b = x powr (a * b)" for a b x :: real by (simp add: powr_def) lemma powr_powr_swap: "(x powr a) powr b = (x powr b) powr a" for a b x :: real by (simp add: powr_powr mult.commute) lemma powr_minus: "x powr (- a) = inverse (x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: powr_def exp_minus [symmetric]) lemma powr_minus_divide: "x powr (- a) = 1/(x powr a)" for a x :: "'a::{ln,real_normed_field}" by (simp add: divide_inverse powr_minus) lemma divide_powr_uminus: "a / b powr c = a * b powr (- c)" for a b c :: real by (simp add: powr_minus_divide) lemma powr_less_mono: "a < b \ 1 < x \ x powr a < x powr b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel: "x powr a < x powr b \ 1 < x \ a < b" for a b x :: real by (simp add: powr_def) lemma powr_less_cancel_iff [simp]: "1 < x \ x powr a < x powr b \ a < b" for a b x :: real by (blast intro: powr_less_cancel powr_less_mono) lemma powr_le_cancel_iff [simp]: "1 < x \ x powr a \ x powr b \ a \ b" for a b x :: real by (simp add: linorder_not_less [symmetric]) lemma powr_realpow: "0 < x \ x powr (real n) = x^n" by (induction n) (simp_all add: ac_simps powr_add) lemma powr_real_of_int': assumes "x \ 0" "x \ 0 \ n > 0" shows "x powr real_of_int n = power_int x n" proof (cases "x = 0") case False with assms have "x > 0" by simp show ?thesis proof (cases n rule: int_cases4) case (nonneg n) thus ?thesis using \x > 0\ by (auto simp add: powr_realpow) next case (neg n) thus ?thesis using \x > 0\ by (auto simp add: powr_realpow powr_minus power_int_minus) qed qed (use assms in auto) lemma log_ln: "ln x = log (exp(1)) x" by (simp add: log_def) lemma DERIV_log: assumes "x > 0" shows "DERIV (\y. log b y) x :> 1 / (ln b * x)" proof - define lb where "lb = 1 / ln b" moreover have "DERIV (\y. lb * ln y) x :> lb / x" using \x > 0\ by (auto intro!: derivative_eq_intros) ultimately show ?thesis by (simp add: log_def) qed lemmas DERIV_log[THEN DERIV_chain2, derivative_intros] and DERIV_log[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemma powr_log_cancel [simp]: "0 < a \ a \ 1 \ 0 < x \ a powr (log a x) = x" by (simp add: powr_def log_def) lemma log_powr_cancel [simp]: "0 < a \ a \ 1 \ log a (a powr y) = y" by (simp add: log_def powr_def) lemma log_mult: "0 < a \ a \ 1 \ 0 < x \ 0 < y \ log a (x * y) = log a x + log a y" by (simp add: log_def ln_mult divide_inverse distrib_right) lemma log_eq_div_ln_mult_log: "0 < a \ a \ 1 \ 0 < b \ b \ 1 \ 0 < x \ log a x = (ln b/ln a) * log b x" by (simp add: log_def divide_inverse) text\Base 10 logarithms\ lemma log_base_10_eq1: "0 < x \ log 10 x = (ln (exp 1) / ln 10) * ln x" by (simp add: log_def) lemma log_base_10_eq2: "0 < x \ log 10 x = (log 10 (exp 1)) * ln x" by (simp add: log_def) lemma log_one [simp]: "log a 1 = 0" by (simp add: log_def) lemma log_eq_one [simp]: "0 < a \ a \ 1 \ log a a = 1" by (simp add: log_def) lemma log_inverse: "0 < a \ a \ 1 \ 0 < x \ log a (inverse x) = - log a x" using ln_inverse log_def by auto lemma log_divide: "0 < a \ a \ 1 \ 0 < x \ 0 < y \ log a (x/y) = log a x - log a y" by (simp add: log_mult divide_inverse log_inverse) lemma powr_gt_zero [simp]: "0 < x powr a \ x \ 0" for a x :: real by (simp add: powr_def) lemma powr_nonneg_iff[simp]: "a powr x \ 0 \ a = 0" for a x::real by (meson not_less powr_gt_zero) lemma log_add_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x + y = log b (x * b powr y)" and add_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y + log b x = log b (b powr y * x)" and log_minus_eq_powr: "0 < b \ b \ 1 \ 0 < x \ log b x - y = log b (x * b powr -y)" and minus_log_eq_powr: "0 < b \ b \ 1 \ 0 < x \ y - log b x = log b (b powr y / x)" by (simp_all add: log_mult log_divide) lemma log_less_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x < log a y \ x < y" using powr_less_cancel_iff [of a] powr_log_cancel [of a x] powr_log_cancel [of a y] by (metis less_eq_real_def less_trans not_le zero_less_one) lemma log_inj: assumes "1 < b" shows "inj_on (log b) {0 <..}" proof (rule inj_onI, simp) fix x y assume pos: "0 < x" "0 < y" and *: "log b x = log b y" show "x = y" proof (cases rule: linorder_cases) assume "x = y" then show ?thesis by simp next assume "x < y" then have "log b x < log b y" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp next assume "y < x" then have "log b y < log b x" using log_less_cancel_iff[OF \1 < b\] pos by simp then show ?thesis using * by simp qed qed lemma log_le_cancel_iff [simp]: "1 < a \ 0 < x \ 0 < y \ log a x \ log a y \ x \ y" by (simp add: linorder_not_less [symmetric]) lemma zero_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 < log a x \ 1 < x" using log_less_cancel_iff[of a 1 x] by simp lemma zero_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 0 \ log a x \ 1 \ x" using log_le_cancel_iff[of a 1 x] by simp lemma log_less_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 0 \ x < 1" using log_less_cancel_iff[of a x 1] by simp lemma log_le_zero_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 0 \ x \ 1" using log_le_cancel_iff[of a x 1] by simp lemma one_less_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 < log a x \ a < x" using log_less_cancel_iff[of a a x] by simp lemma one_le_log_cancel_iff[simp]: "1 < a \ 0 < x \ 1 \ log a x \ a \ x" using log_le_cancel_iff[of a a x] by simp lemma log_less_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x < 1 \ x < a" using log_less_cancel_iff[of a x a] by simp lemma log_le_one_cancel_iff[simp]: "1 < a \ 0 < x \ log a x \ 1 \ x \ a" using log_le_cancel_iff[of a x a] by simp lemma le_log_iff: fixes b x y :: real assumes "1 < b" "x > 0" shows "y \ log b x \ b powr y \ x" using assms by (metis less_irrefl less_trans powr_le_cancel_iff powr_log_cancel zero_less_one) lemma less_log_iff: assumes "1 < b" "x > 0" shows "y < log b x \ b powr y < x" by (metis assms dual_order.strict_trans less_irrefl powr_less_cancel_iff powr_log_cancel zero_less_one) lemma assumes "1 < b" "x > 0" shows log_less_iff: "log b x < y \ x < b powr y" and log_le_iff: "log b x \ y \ x \ b powr y" using le_log_iff[OF assms, of y] less_log_iff[OF assms, of y] by auto lemmas powr_le_iff = le_log_iff[symmetric] and powr_less_iff = less_log_iff[symmetric] and less_powr_iff = log_less_iff[symmetric] and le_powr_iff = log_le_iff[symmetric] lemma le_log_of_power: assumes "b ^ n \ m" "1 < b" shows "n \ log b m" proof - from assms have "0 < m" by (metis less_trans zero_less_power less_le_trans zero_less_one) thus ?thesis using assms by (simp add: le_log_iff powr_realpow) qed lemma le_log2_of_power: "2 ^ n \ m \ n \ log 2 m" for m n :: nat using le_log_of_power[of 2] by simp lemma log_of_power_le: "\ m \ b ^ n; b > 1; m > 0 \ \ log b (real m) \ n" by (simp add: log_le_iff powr_realpow) lemma log2_of_power_le: "\ m \ 2 ^ n; m > 0 \ \ log 2 m \ n" for m n :: nat using log_of_power_le[of _ 2] by simp lemma log_of_power_less: "\ m < b ^ n; b > 1; m > 0 \ \ log b (real m) < n" by (simp add: log_less_iff powr_realpow) lemma log2_of_power_less: "\ m < 2 ^ n; m > 0 \ \ log 2 m < n" for m n :: nat using log_of_power_less[of _ 2] by simp lemma less_log_of_power: assumes "b ^ n < m" "1 < b" shows "n < log b m" proof - have "0 < m" by (metis assms less_trans zero_less_power zero_less_one) thus ?thesis using assms by (simp add: less_log_iff powr_realpow) qed lemma less_log2_of_power: "2 ^ n < m \ n < log 2 m" for m n :: nat using less_log_of_power[of 2] by simp lemma gr_one_powr[simp]: fixes x y :: real shows "\ x > 1; y > 0 \ \ 1 < x powr y" by(simp add: less_powr_iff) lemma log_pow_cancel [simp]: "a > 0 \ a \ 1 \ log a (a ^ b) = b" by (simp add: ln_realpow log_def) lemma floor_log_eq_powr_iff: "x > 0 \ b > 1 \ \log b x\ = k \ b powr k \ x \ x < b powr (k + 1)" by (auto simp: floor_eq_iff powr_le_iff less_powr_iff) lemma floor_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ floor (log b (real k)) = n \ b^n \ k \ k < b^(n+1)" by (auto simp: floor_log_eq_powr_iff powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma floor_log_nat_eq_if: fixes b n k :: nat assumes "b^n \ k" "k < b^(n+1)" "b \ 2" shows "floor (log b (real k)) = n" proof - have "k \ 1" using assms(1,3) one_le_power[of b n] by linarith with assms show ?thesis by(simp add: floor_log_nat_eq_powr_iff) qed lemma ceiling_log_eq_powr_iff: "\ x > 0; b > 1 \ \ \log b x\ = int k + 1 \ b powr k < x \ x \ b powr (k + 1)" by (auto simp: ceiling_eq_iff powr_less_iff le_powr_iff) lemma ceiling_log_nat_eq_powr_iff: fixes b n k :: nat shows "\ b \ 2; k > 0 \ \ ceiling (log b (real k)) = int n + 1 \ (b^n < k \ k \ b^(n+1))" using ceiling_log_eq_powr_iff by (auto simp: powr_add powr_realpow of_nat_power[symmetric] of_nat_mult[symmetric] ac_simps simp del: of_nat_power of_nat_mult) lemma ceiling_log_nat_eq_if: fixes b n k :: nat assumes "b^n < k" "k \ b^(n+1)" "b \ 2" shows "ceiling (log b (real k)) = int n + 1" proof - have "k \ 1" using assms(1,3) one_le_power[of b n] by linarith with assms show ?thesis by(simp add: ceiling_log_nat_eq_powr_iff) qed lemma floor_log2_div2: fixes n :: nat assumes "n \ 2" shows "floor(log 2 n) = floor(log 2 (n div 2)) + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "n div 2" assume "n\2" hence "1 \ ?m" using assms by arith then obtain i where i: "2 ^ i \ ?m" "?m < 2 ^ (i + 1)" using ex_power_ivl1[of 2 ?m] by auto have "2^(i+1) \ 2*?m" using i(1) by simp also have "2*?m \ n" by arith finally have *: "2^(i+1) \ \" . have "n < 2^(i+1+1)" using i(2) by simp from floor_log_nat_eq_if[OF * this] floor_log_nat_eq_if[OF i] show ?thesis by simp qed lemma ceiling_log2_div2: assumes "n \ 2" shows "ceiling(log 2 (real n)) = ceiling(log 2 ((n-1) div 2 + 1)) + 1" proof cases assume "n=2" thus ?thesis by simp next let ?m = "(n-1) div 2 + 1" assume "n\2" hence "2 \ ?m" using assms by arith then obtain i where i: "2 ^ i < ?m" "?m \ 2 ^ (i + 1)" using ex_power_ivl2[of 2 ?m] by auto have "n \ 2*?m" by arith also have "2*?m \ 2 ^ ((i+1)+1)" using i(2) by simp finally have *: "n \ \" . have "2^(i+1) < n" using i(1) by (auto simp: less_Suc_eq_0_disj) from ceiling_log_nat_eq_if[OF this *] ceiling_log_nat_eq_if[OF i] show ?thesis by simp qed lemma powr_real_of_int: "x > 0 \ x powr real_of_int n = (if n \ 0 then x ^ nat n else inverse (x ^ nat (- n)))" using powr_realpow[of x "nat n"] powr_realpow[of x "nat (-n)"] by (auto simp: field_simps powr_minus) lemma powr_numeral [simp]: "0 \ x \ x powr (numeral n :: real) = x ^ (numeral n)" by (metis less_le power_zero_numeral powr_0 of_nat_numeral powr_realpow) lemma powr_int: assumes "x > 0" shows "x powr i = (if i \ 0 then x ^ nat i else 1 / x ^ nat (-i))" proof (cases "i < 0") case True have r: "x powr i = 1 / x powr (- i)" by (simp add: powr_minus field_simps) show ?thesis using \i < 0\ \x > 0\ by (simp add: r field_simps powr_realpow[symmetric]) next case False then show ?thesis by (simp add: assms powr_realpow[symmetric]) qed definition powr_real :: "real \ real \ real" where [code_abbrev, simp]: "powr_real = Transcendental.powr" lemma compute_powr_real [code]: "powr_real b i = (if b \ 0 then Code.abort (STR ''powr_real with nonpositive base'') (\_. powr_real b i) else if \i\ = i then (if 0 \ i then b ^ nat \i\ else 1 / b ^ nat \- i\) else Code.abort (STR ''powr_real with non-integer exponent'') (\_. powr_real b i))" for b i :: real by (auto simp: powr_int) lemma powr_one: "0 \ x \ x powr 1 = x" for x :: real using powr_realpow [of x 1] by simp lemma powr_neg_one: "0 < x \ x powr - 1 = 1 / x" for x :: real using powr_int [of x "- 1"] by simp lemma powr_neg_numeral: "0 < x \ x powr - numeral n = 1 / x ^ numeral n" for x :: real using powr_int [of x "- numeral n"] by simp lemma root_powr_inverse: "0 < n \ 0 < x \ root n x = x powr (1/n)" by (rule real_root_pos_unique) (auto simp: powr_realpow[symmetric] powr_powr) lemma ln_powr: "x \ 0 \ ln (x powr y) = y * ln x" for x :: real by (simp add: powr_def) lemma ln_root: "n > 0 \ b > 0 \ ln (root n b) = ln b / n" by (simp add: root_powr_inverse ln_powr) lemma ln_sqrt: "0 < x \ ln (sqrt x) = ln x / 2" by (simp add: ln_powr ln_powr[symmetric] mult.commute) lemma log_root: "n > 0 \ a > 0 \ log b (root n a) = log b a / n" by (simp add: log_def ln_root) lemma log_powr: "x \ 0 \ log b (x powr y) = y * log b x" by (simp add: log_def ln_powr) (* [simp] is not worth it, interferes with some proofs *) lemma log_nat_power: "0 < x \ log b (x^n) = real n * log b x" by (simp add: log_powr powr_realpow [symmetric]) lemma log_of_power_eq: assumes "m = b ^ n" "b > 1" shows "n = log b (real m)" proof - have "n = log b (b ^ n)" using assms(2) by (simp add: log_nat_power) also have "\ = log b m" using assms by simp finally show ?thesis . qed lemma log2_of_power_eq: "m = 2 ^ n \ n = log 2 m" for m n :: nat using log_of_power_eq[of _ 2] by simp lemma log_base_change: "0 < a \ a \ 1 \ log b x = log a x / log a b" by (simp add: log_def) lemma log_base_pow: "0 < a \ log (a ^ n) x = log a x / n" by (simp add: log_def ln_realpow) lemma log_base_powr: "a \ 0 \ log (a powr b) x = log a x / b" by (simp add: log_def ln_powr) lemma log_base_root: "n > 0 \ b > 0 \ log (root n b) x = n * (log b x)" by (simp add: log_def ln_root) lemma ln_bound: "0 < x \ ln x \ x" for x :: real using ln_le_minus_one by force lemma powr_mono: fixes x :: real assumes "a \ b" and "1 \ x" shows "x powr a \ x powr b" using assms less_eq_real_def by auto lemma ge_one_powr_ge_zero: "1 \ x \ 0 \ a \ 1 \ x powr a" for x :: real using powr_mono by fastforce lemma powr_less_mono2: "0 < a \ 0 \ x \ x < y \ x powr a < y powr a" for x :: real by (simp add: powr_def) lemma powr_less_mono2_neg: "a < 0 \ 0 < x \ x < y \ y powr a < x powr a" for x :: real by (simp add: powr_def) lemma powr_mono2: "x powr a \ y powr a" if "0 \ a" "0 \ x" "x \ y" for x :: real using less_eq_real_def powr_less_mono2 that by auto lemma powr_le1: "0 \ a \ 0 \ x \ x \ 1 \ x powr a \ 1" for x :: real using powr_mono2 by fastforce lemma powr_mono2': fixes a x y :: real assumes "a \ 0" "x > 0" "x \ y" shows "x powr a \ y powr a" proof - from assms have "x powr - a \ y powr - a" by (intro powr_mono2) simp_all with assms show ?thesis by (auto simp: powr_minus field_simps) qed lemma powr_mono_both: fixes x :: real assumes "0 \ a" "a \ b" "1 \ x" "x \ y" shows "x powr a \ y powr b" by (meson assms order.trans powr_mono powr_mono2 zero_le_one) lemma powr_inj: "0 < a \ a \ 1 \ a powr x = a powr y \ x = y" for x :: real unfolding powr_def exp_inj_iff by simp lemma powr_half_sqrt: "0 \ x \ x powr (1/2) = sqrt x" by (simp add: powr_def root_powr_inverse sqrt_def) lemma square_powr_half [simp]: fixes x::real shows "x\<^sup>2 powr (1/2) = \x\" by (simp add: powr_half_sqrt) lemma ln_powr_bound: "1 \ x \ 0 < a \ ln x \ (x powr a) / a" for x :: real by (metis exp_gt_zero linear ln_eq_zero_iff ln_exp ln_less_self ln_powr mult.commute mult_imp_le_div_pos not_less powr_gt_zero) lemma ln_powr_bound2: fixes x :: real assumes "1 < x" and "0 < a" shows "(ln x) powr a \ (a powr a) * x" proof - from assms have "ln x \ (x powr (1 / a)) / (1 / a)" by (metis less_eq_real_def ln_powr_bound zero_less_divide_1_iff) also have "\ = a * (x powr (1 / a))" by simp finally have "(ln x) powr a \ (a * (x powr (1 / a))) powr a" by (metis assms less_imp_le ln_gt_zero powr_mono2) also have "\ = (a powr a) * ((x powr (1 / a)) powr a)" using assms powr_mult by auto also have "(x powr (1 / a)) powr a = x powr ((1 / a) * a)" by (rule powr_powr) also have "\ = x" using assms by auto finally show ?thesis . qed lemma tendsto_powr: fixes a b :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0" shows "((\x. f x powr g x) \ a powr b) F" unfolding powr_def proof (rule filterlim_If) from f show "((\x. 0) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x = 0}))" by simp (auto simp: filterlim_iff eventually_inf_principal elim: eventually_mono dest: t1_space_nhds) from f g a show "((\x. exp (g x * ln (f x))) \ (if a = 0 then 0 else exp (b * ln a))) (inf F (principal {x. f x \ 0}))" by (auto intro!: tendsto_intros intro: tendsto_mono inf_le1) qed lemma tendsto_powr'[tendsto_intros]: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and a: "a \ 0 \ (b > 0 \ eventually (\x. f x \ 0) F)" shows "((\x. f x powr g x) \ a powr b) F" proof - from a consider "a \ 0" | "a = 0" "b > 0" "eventually (\x. f x \ 0) F" by auto then show ?thesis proof cases case 1 with f g show ?thesis by (rule tendsto_powr) next case 2 have "((\x. if f x = 0 then 0 else exp (g x * ln (f x))) \ 0) F" proof (intro filterlim_If) have "filterlim f (principal {0<..}) (inf F (principal {z. f z \ 0}))" using \eventually (\x. f x \ 0) F\ by (auto simp: filterlim_iff eventually_inf_principal eventually_principal elim: eventually_mono) moreover have "filterlim f (nhds a) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ f]) simp_all ultimately have f: "filterlim f (at_right 0) (inf F (principal {x. f x \ 0}))" by (simp add: at_within_def filterlim_inf \a = 0\) have g: "(g \ b) (inf F (principal {z. f z \ 0}))" by (rule tendsto_mono[OF _ g]) simp_all show "((\x. exp (g x * ln (f x))) \ 0) (inf F (principal {x. f x \ 0}))" by (rule filterlim_compose[OF exp_at_bot] filterlim_tendsto_pos_mult_at_bot filterlim_compose[OF ln_at_0] f g \b > 0\)+ qed simp_all with \a = 0\ show ?thesis by (simp add: powr_def) qed qed lemma continuous_powr: assumes "continuous F f" and "continuous F g" and "f (Lim F (\x. x)) \ 0" shows "continuous F (\x. (f x) powr (g x :: real))" using assms unfolding continuous_def by (rule tendsto_powr) lemma continuous_at_within_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous (at a within s) f" and "continuous (at a within s) g" and "f a \ 0" shows "continuous (at a within s) (\x. (f x) powr (g x))" using assms unfolding continuous_within by (rule tendsto_powr) lemma isCont_powr[continuous_intros, simp]: fixes f g :: "_ \ real" assumes "isCont f a" "isCont g a" "f a \ 0" shows "isCont (\x. (f x) powr g x) a" using assms unfolding continuous_at by (rule tendsto_powr) lemma continuous_on_powr[continuous_intros]: fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0" shows "continuous_on s (\x. (f x) powr (g x))" using assms unfolding continuous_on_def by (fast intro: tendsto_powr) lemma tendsto_powr2: fixes a :: real assumes f: "(f \ a) F" and g: "(g \ b) F" and "\\<^sub>F x in F. 0 \ f x" and b: "0 < b" shows "((\x. f x powr g x) \ a powr b) F" using tendsto_powr'[of f a F g b] assms by auto lemma has_derivative_powr[derivative_intros]: assumes g[derivative_intros]: "(g has_derivative g') (at x within X)" and f[derivative_intros]:"(f has_derivative f') (at x within X)" assumes pos: "0 < g x" and "x \ X" shows "((\x. g x powr f x::real) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" proof - have "\\<^sub>F x in at x within X. g x > 0" by (rule order_tendstoD[OF _ pos]) (rule has_derivative_continuous[OF g, unfolded continuous_within]) then obtain d where "d > 0" and pos': "\x'. x' \ X \ dist x' x < d \ 0 < g x'" using pos unfolding eventually_at by force have "((\x. exp (f x * ln (g x))) has_derivative (\h. (g x powr f x) * (f' h * ln (g x) + g' h * f x / g x))) (at x within X)" using pos by (auto intro!: derivative_eq_intros simp: field_split_simps powr_def) then show ?thesis by (rule has_derivative_transform_within[OF _ \d > 0\ \x \ X\]) (auto simp: powr_def dest: pos') qed lemma DERIV_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" and f: "DERIV f x :> r" shows "DERIV (\x. g x powr f x) x :> (g x powr f x) * (r * ln (g x) + m * f x / g x)" using assms by (auto intro!: derivative_eq_intros ext simp: has_field_derivative_def algebra_simps) lemma DERIV_fun_powr: fixes r :: real assumes g: "DERIV g x :> m" and pos: "g x > 0" shows "DERIV (\x. (g x) powr r) x :> r * (g x) powr (r - of_nat 1) * m" using DERIV_powr[OF g pos DERIV_const, of r] pos by (simp add: powr_diff field_simps) lemma has_real_derivative_powr: assumes "z > 0" shows "((\z. z powr r) has_real_derivative r * z powr (r - 1)) (at z)" proof (subst DERIV_cong_ev[OF refl _ refl]) from assms have "eventually (\z. z \ 0) (nhds z)" by (intro t1_space_nhds) auto then show "eventually (\z. z powr r = exp (r * ln z)) (nhds z)" unfolding powr_def by eventually_elim simp from assms show "((\z. exp (r * ln z)) has_real_derivative r * z powr (r - 1)) (at z)" by (auto intro!: derivative_eq_intros simp: powr_def field_simps exp_diff) qed declare has_real_derivative_powr[THEN DERIV_chain2, derivative_intros] lemma tendsto_zero_powrI: assumes "(f \ (0::real)) F" "(g \ b) F" "\\<^sub>F x in F. 0 \ f x" "0 < b" shows "((\x. f x powr g x) \ 0) F" using tendsto_powr2[OF assms] by simp lemma continuous_on_powr': fixes f g :: "_ \ real" assumes "continuous_on s f" "continuous_on s g" and "\x\s. f x \ 0 \ (f x = 0 \ g x > 0)" shows "continuous_on s (\x. (f x) powr (g x))" unfolding continuous_on_def proof fix x assume x: "x \ s" from assms x show "((\x. f x powr g x) \ f x powr g x) (at x within s)" proof (cases "f x = 0") case True from assms(3) have "eventually (\x. f x \ 0) (at x within s)" by (auto simp: at_within_def eventually_inf_principal) with True x assms show ?thesis by (auto intro!: tendsto_zero_powrI[of f _ g "g x"] simp: continuous_on_def) next case False with assms x show ?thesis by (auto intro!: tendsto_powr' simp: continuous_on_def) qed qed lemma tendsto_neg_powr: assumes "s < 0" and f: "LIM x F. f x :> at_top" shows "((\x. f x powr s) \ (0::real)) F" proof - have "((\x. exp (s * ln (f x))) \ (0::real)) F" (is "?X") by (auto intro!: filterlim_compose[OF exp_at_bot] filterlim_compose[OF ln_at_top] filterlim_tendsto_neg_mult_at_bot assms) also have "?X \ ((\x. f x powr s) \ (0::real)) F" using f filterlim_at_top_dense[of f F] by (intro filterlim_cong[OF refl refl]) (auto simp: neq_iff powr_def elim: eventually_mono) finally show ?thesis . qed lemma tendsto_exp_limit_at_right: "((\y. (1 + x * y) powr (1 / y)) \ exp x) (at_right 0)" for x :: real proof (cases "x = 0") case True then show ?thesis by simp next case False have "((\y. ln (1 + x * y)::real) has_real_derivative 1 * x) (at 0)" by (auto intro!: derivative_eq_intros) then have "((\y. ln (1 + x * y) / y) \ x) (at 0)" by (auto simp: has_field_derivative_def field_has_derivative_at) then have *: "((\y. exp (ln (1 + x * y) / y)) \ exp x) (at 0)" by (rule tendsto_intros) then show ?thesis proof (rule filterlim_mono_eventually) show "eventually (\xa. exp (ln (1 + x * xa) / xa) = (1 + x * xa) powr (1 / xa)) (at_right 0)" unfolding eventually_at_right[OF zero_less_one] using False by (intro exI[of _ "1 / \x\"]) (auto simp: field_simps powr_def abs_if add_nonneg_eq_0_iff) qed (simp_all add: at_eq_sup_left_right) qed lemma tendsto_exp_limit_at_top: "((\y. (1 + x / y) powr y) \ exp x) at_top" for x :: real by (simp add: filterlim_at_top_to_right inverse_eq_divide tendsto_exp_limit_at_right) lemma tendsto_exp_limit_sequentially: "(\n. (1 + x / n) ^ n) \ exp x" for x :: real proof (rule filterlim_mono_eventually) from reals_Archimedean2 [of "\x\"] obtain n :: nat where *: "real n > \x\" .. then have "eventually (\n :: nat. 0 < 1 + x / real n) at_top" by (intro eventually_sequentiallyI [of n]) (auto simp: field_split_simps) then show "eventually (\n. (1 + x / n) powr n = (1 + x / n) ^ n) at_top" by (rule eventually_mono) (erule powr_realpow) show "(\n. (1 + x / real n) powr real n) \ exp x" by (rule filterlim_compose [OF tendsto_exp_limit_at_top filterlim_real_sequentially]) qed auto subsection \Sine and Cosine\ definition sin_coeff :: "nat \ real" where "sin_coeff = (\n. if even n then 0 else (- 1) ^ ((n - Suc 0) div 2) / (fact n))" definition cos_coeff :: "nat \ real" where "cos_coeff = (\n. if even n then ((- 1) ^ (n div 2)) / (fact n) else 0)" definition sin :: "'a \ 'a::{real_normed_algebra_1,banach}" where "sin = (\x. \n. sin_coeff n *\<^sub>R x^n)" definition cos :: "'a \ 'a::{real_normed_algebra_1,banach}" where "cos = (\x. \n. cos_coeff n *\<^sub>R x^n)" lemma sin_coeff_0 [simp]: "sin_coeff 0 = 0" unfolding sin_coeff_def by simp lemma cos_coeff_0 [simp]: "cos_coeff 0 = 1" unfolding cos_coeff_def by simp lemma sin_coeff_Suc: "sin_coeff (Suc n) = cos_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) lemma cos_coeff_Suc: "cos_coeff (Suc n) = - sin_coeff n / real (Suc n)" unfolding cos_coeff_def sin_coeff_def by (simp del: mult_Suc) (auto elim: oddE) lemma summable_norm_sin: "summable (\n. norm (sin_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_comparison_test [OF _ summable_norm_exp]) show "\N. \n\N. norm (norm (sin_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" unfolding sin_coeff_def by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) qed lemma summable_norm_cos: "summable (\n. norm (cos_coeff n *\<^sub>R x^n))" for x :: "'a::{real_normed_algebra_1,banach}" proof (rule summable_comparison_test [OF _ summable_norm_exp]) show "\N. \n\N. norm (norm (cos_coeff n *\<^sub>R x ^ n)) \ norm (x ^ n /\<^sub>R fact n)" unfolding cos_coeff_def by (auto simp: divide_inverse abs_mult power_abs [symmetric] zero_le_mult_iff) qed lemma sin_converges: "(\n. sin_coeff n *\<^sub>R x^n) sums sin x" unfolding sin_def by (metis (full_types) summable_norm_cancel summable_norm_sin summable_sums) lemma cos_converges: "(\n. cos_coeff n *\<^sub>R x^n) sums cos x" unfolding cos_def by (metis (full_types) summable_norm_cancel summable_norm_cos summable_sums) lemma sin_of_real: "sin (of_real x) = of_real (sin x)" for x :: real proof - have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) = (\n. sin_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (sin_coeff n *\<^sub>R x^n) = sin_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (sin (of_real x))" by (rule sin_converges) finally have "(\n. of_real (sin_coeff n *\<^sub>R x^n)) sums (sin (of_real x))" . then show ?thesis using sums_unique2 sums_of_real [OF sin_converges] by blast qed corollary sin_in_Reals [simp]: "z \ \ \ sin z \ \" by (metis Reals_cases Reals_of_real sin_of_real) lemma cos_of_real: "cos (of_real x) = of_real (cos x)" for x :: real proof - have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) = (\n. cos_coeff n *\<^sub>R (of_real x)^n)" proof show "of_real (cos_coeff n *\<^sub>R x^n) = cos_coeff n *\<^sub>R of_real x^n" for n by (simp add: scaleR_conv_of_real) qed also have "\ sums (cos (of_real x))" by (rule cos_converges) finally have "(\n. of_real (cos_coeff n *\<^sub>R x^n)) sums (cos (of_real x))" . then show ?thesis using sums_unique2 sums_of_real [OF cos_converges] by blast qed corollary cos_in_Reals [simp]: "z \ \ \ cos z \ \" by (metis Reals_cases Reals_of_real cos_of_real) lemma diffs_sin_coeff: "diffs sin_coeff = cos_coeff" by (simp add: diffs_def sin_coeff_Suc del: of_nat_Suc) lemma diffs_cos_coeff: "diffs cos_coeff = (\n. - sin_coeff n)" by (simp add: diffs_def cos_coeff_Suc del: of_nat_Suc) lemma sin_int_times_real: "sin (of_int m * of_real x) = of_real (sin (of_int m * x))" by (metis sin_of_real of_real_mult of_real_of_int_eq) lemma cos_int_times_real: "cos (of_int m * of_real x) = of_real (cos (of_int m * x))" by (metis cos_of_real of_real_mult of_real_of_int_eq) text \Now at last we can get the derivatives of exp, sin and cos.\ lemma DERIV_sin [simp]: "DERIV sin x :> cos x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_sin[THEN DERIV_chain2, derivative_intros] and DERIV_sin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_sin[derivative_intros] = DERIV_sin[THEN DERIV_compose_FDERIV] lemma DERIV_cos [simp]: "DERIV cos x :> - sin x" for x :: "'a::{real_normed_field,banach}" unfolding sin_def cos_def scaleR_conv_of_real apply (rule DERIV_cong) apply (rule termdiffs [where K="of_real (norm x) + 1 :: 'a"]) apply (simp_all add: norm_less_p1 diffs_of_real diffs_minus suminf_minus diffs_sin_coeff diffs_cos_coeff summable_minus_iff scaleR_conv_of_real [symmetric] summable_norm_sin [THEN summable_norm_cancel] summable_norm_cos [THEN summable_norm_cancel]) done declare DERIV_cos[THEN DERIV_chain2, derivative_intros] and DERIV_cos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_cos[derivative_intros] = DERIV_cos[THEN DERIV_compose_FDERIV] lemma isCont_sin: "isCont sin x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_sin [THEN DERIV_isCont]) lemma continuous_on_sin_real: "continuous_on {a..b} sin" for a::real using continuous_at_imp_continuous_on isCont_sin by blast lemma isCont_cos: "isCont cos x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cos [THEN DERIV_isCont]) lemma continuous_on_cos_real: "continuous_on {a..b} cos" for a::real using continuous_at_imp_continuous_on isCont_cos by blast context fixes f :: "'a::t2_space \ 'b::{real_normed_field,banach}" begin lemma isCont_sin' [simp]: "isCont f a \ isCont (\x. sin (f x)) a" by (rule isCont_o2 [OF _ isCont_sin]) lemma isCont_cos' [simp]: "isCont f a \ isCont (\x. cos (f x)) a" by (rule isCont_o2 [OF _ isCont_cos]) lemma tendsto_sin [tendsto_intros]: "(f \ a) F \ ((\x. sin (f x)) \ sin a) F" by (rule isCont_tendsto_compose [OF isCont_sin]) lemma tendsto_cos [tendsto_intros]: "(f \ a) F \ ((\x. cos (f x)) \ cos a) F" by (rule isCont_tendsto_compose [OF isCont_cos]) lemma continuous_sin [continuous_intros]: "continuous F f \ continuous F (\x. sin (f x))" unfolding continuous_def by (rule tendsto_sin) lemma continuous_on_sin [continuous_intros]: "continuous_on s f \ continuous_on s (\x. sin (f x))" unfolding continuous_on_def by (auto intro: tendsto_sin) lemma continuous_cos [continuous_intros]: "continuous F f \ continuous F (\x. cos (f x))" unfolding continuous_def by (rule tendsto_cos) lemma continuous_on_cos [continuous_intros]: "continuous_on s f \ continuous_on s (\x. cos (f x))" unfolding continuous_on_def by (auto intro: tendsto_cos) end lemma continuous_within_sin: "continuous (at z within s) sin" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_sin) lemma continuous_within_cos: "continuous (at z within s) cos" for z :: "'a::{real_normed_field,banach}" by (simp add: continuous_within tendsto_cos) subsection \Properties of Sine and Cosine\ lemma sin_zero [simp]: "sin 0 = 0" by (simp add: sin_def sin_coeff_def scaleR_conv_of_real) lemma cos_zero [simp]: "cos 0 = 1" by (simp add: cos_def cos_coeff_def scaleR_conv_of_real) lemma DERIV_fun_sin: "DERIV g x :> m \ DERIV (\x. sin (g x)) x :> cos (g x) * m" by (fact derivative_intros) lemma DERIV_fun_cos: "DERIV g x :> m \ DERIV (\x. cos(g x)) x :> - sin (g x) * m" by (fact derivative_intros) subsection \Deriving the Addition Formulas\ text \The product of two cosine series.\ lemma cos_x_cos_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (cos x * cos y)" proof - have "(cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p - n)) = (if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p - n) else 0)" if "n \ p" for n p :: nat proof - from that have *: "even n \ even p \ (-1) ^ (n div 2) * (-1) ^ ((p - n) div 2) = (-1 :: real) ^ (p div 2)" by (metis div_add power_add le_add_diff_inverse odd_add) with that show ?thesis by (auto simp: algebra_simps cos_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ even n then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (cos_coeff n * cos_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (cos_coeff n *\<^sub>R x^n) * (cos_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (cos x * cos y)" using summable_norm_cos by (auto simp: cos_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed text \The product of two sine series.\ lemma sin_x_sin_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums (sin x * sin y)" proof - have "(sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)) = (if even p \ odd n then -((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat proof - have "(-1) ^ ((n - Suc 0) div 2) * (-1) ^ ((p - Suc n) div 2) = - ((-1 :: real) ^ (p div 2))" if np: "odd n" "even p" proof - have "p > 0" using \n \ p\ neq0_conv that(1) by blast then have \
: "(- 1::real) ^ (p div 2 - Suc 0) = - ((- 1) ^ (p div 2))" using \even p\ by (auto simp add: dvd_def power_eq_if) from \n \ p\ np have *: "n - Suc 0 + (p - Suc n) = p - Suc (Suc 0)" "Suc (Suc 0) \ p" by arith+ have "(p - Suc (Suc 0)) div 2 = p div 2 - Suc 0" by simp with \n \ p\ np \
* show ?thesis by (simp add: flip: div_add power_add) qed then show ?thesis using \n\p\ by (auto simp: algebra_simps sin_coeff_def binomial_fact) qed then have "(\p. \n\p. if even p \ odd n then - ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. \n\p. (sin_coeff n * sin_coeff (p - n)) *\<^sub>R (x^n * y^(p-n)))" by simp also have "\ = (\p. \n\p. (sin_coeff n *\<^sub>R x^n) * (sin_coeff (p - n) *\<^sub>R y^(p-n)))" by (simp add: algebra_simps) also have "\ sums (sin x * sin y)" using summable_norm_sin by (auto simp: sin_def scaleR_conv_of_real intro!: Cauchy_product_sums) finally show ?thesis . qed lemma sums_cos_x_plus_y: fixes x :: "'a::{real_normed_field,banach}" shows "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) sums cos (x + y)" proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = cos_coeff p *\<^sub>R ((x + y) ^ p)" for p :: nat proof - have "(\n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then \n\p. ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" by simp also have "\ = (if even p then of_real ((-1) ^ (p div 2) / (fact p)) * (\n\p. (p choose n) *\<^sub>R (x^n) * y^(p-n)) else 0)" by (auto simp: sum_distrib_left field_simps scaleR_conv_of_real nonzero_of_real_divide) also have "\ = cos_coeff p *\<^sub>R ((x + y) ^ p)" by (simp add: cos_coeff_def binomial_ring [of x y] scaleR_conv_of_real atLeast0AtMost) finally show ?thesis . qed then have "(\p. \n\p. if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (\p. cos_coeff p *\<^sub>R ((x+y)^p))" by simp also have "\ sums cos (x + y)" by (rule cos_converges) finally show ?thesis . qed theorem cos_add: fixes x :: "'a::{real_normed_field,banach}" shows "cos (x + y) = cos x * cos y - sin x * sin y" proof - have "(if even p \ even n then ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) - (if even p \ odd n then - ((- 1) ^ (p div 2) * int (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0) = (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)" if "n \ p" for n p :: nat by simp then have "(\p. \n\p. (if even p then ((-1) ^ (p div 2) * (p choose n) / (fact p)) *\<^sub>R (x^n) * y^(p-n) else 0)) sums (cos x * cos y - sin x * sin y)" using sums_diff [OF cos_x_cos_y [of x y] sin_x_sin_y [of x y]] by (simp add: sum_subtractf [symmetric]) then show ?thesis by (blast intro: sums_cos_x_plus_y sums_unique2) qed lemma sin_minus_converges: "(\n. - (sin_coeff n *\<^sub>R (-x)^n)) sums sin x" proof - have [simp]: "\n. - (sin_coeff n *\<^sub>R (-x)^n) = (sin_coeff n *\<^sub>R x^n)" by (auto simp: sin_coeff_def elim!: oddE) show ?thesis by (simp add: sin_def summable_norm_sin [THEN summable_norm_cancel, THEN summable_sums]) qed lemma sin_minus [simp]: "sin (- x) = - sin x" for x :: "'a::{real_normed_algebra_1,banach}" using sin_minus_converges [of x] by (auto simp: sin_def summable_norm_sin [THEN summable_norm_cancel] suminf_minus sums_iff equation_minus_iff) lemma cos_minus_converges: "(\n. (cos_coeff n *\<^sub>R (-x)^n)) sums cos x" proof - have [simp]: "\n. (cos_coeff n *\<^sub>R (-x)^n) = (cos_coeff n *\<^sub>R x^n)" by (auto simp: Transcendental.cos_coeff_def elim!: evenE) show ?thesis by (simp add: cos_def summable_norm_cos [THEN summable_norm_cancel, THEN summable_sums]) qed lemma cos_minus [simp]: "cos (-x) = cos x" for x :: "'a::{real_normed_algebra_1,banach}" using cos_minus_converges [of x] by (simp add: cos_def summable_norm_cos [THEN summable_norm_cancel] suminf_minus sums_iff equation_minus_iff) lemma sin_cos_squared_add [simp]: "(sin x)\<^sup>2 + (cos x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "-x"] by (simp add: power2_eq_square algebra_simps) lemma sin_cos_squared_add2 [simp]: "(cos x)\<^sup>2 + (sin x)\<^sup>2 = 1" for x :: "'a::{real_normed_field,banach}" by (subst add.commute, rule sin_cos_squared_add) lemma sin_cos_squared_add3 [simp]: "cos x * cos x + sin x * sin x = 1" for x :: "'a::{real_normed_field,banach}" using sin_cos_squared_add2 [unfolded power2_eq_square] . lemma sin_squared_eq: "(sin x)\<^sup>2 = 1 - (cos x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add) lemma cos_squared_eq: "(cos x)\<^sup>2 = 1 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" unfolding eq_diff_eq by (rule sin_cos_squared_add2) lemma abs_sin_le_one [simp]: "\sin x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: sin_squared_eq) lemma sin_ge_minus_one [simp]: "- 1 \ sin x" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma sin_le_one [simp]: "sin x \ 1" for x :: real using abs_sin_le_one [of x] by (simp add: abs_le_iff) lemma abs_cos_le_one [simp]: "\cos x\ \ 1" for x :: real by (rule power2_le_imp_le) (simp_all add: cos_squared_eq) lemma cos_ge_minus_one [simp]: "- 1 \ cos x" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_le_one [simp]: "cos x \ 1" for x :: real using abs_cos_le_one [of x] by (simp add: abs_le_iff) lemma cos_diff: "cos (x - y) = cos x * cos y + sin x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of x "- y"] by simp lemma cos_double: "cos(2*x) = (cos x)\<^sup>2 - (sin x)\<^sup>2" for x :: "'a::{real_normed_field,banach}" using cos_add [where x=x and y=x] by (simp add: power2_eq_square) lemma sin_cos_le1: "\sin x * sin y + cos x * cos y\ \ 1" for x :: real using cos_diff [of x y] by (metis abs_cos_le_one add.commute) lemma DERIV_fun_pow: "DERIV g x :> m \ DERIV (\x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m" by (auto intro!: derivative_eq_intros simp:) lemma DERIV_fun_exp: "DERIV g x :> m \ DERIV (\x. exp (g x)) x :> exp (g x) * m" by (auto intro!: derivative_intros) subsection \The Constant Pi\ definition pi :: real where "pi = 2 * (THE x. 0 \ x \ x \ 2 \ cos x = 0)" text \Show that there's a least positive \<^term>\x\ with \<^term>\cos x = 0\; hence define pi.\ lemma sin_paired: "(\n. (- 1) ^ n / (fact (2 * n + 1)) * x ^ (2 * n + 1)) sums sin x" for x :: real proof - have "(\n. \k = n*2..n. 0 < ?f n" proof fix n :: nat let ?k2 = "real (Suc (Suc (4 * n)))" let ?k3 = "real (Suc (Suc (Suc (4 * n))))" have "x * x < ?k2 * ?k3" using assms by (intro mult_strict_mono', simp_all) then have "x * x * x * x ^ (n * 4) < ?k2 * ?k3 * x * x ^ (n * 4)" by (intro mult_strict_right_mono zero_less_power \0 < x\) then show "0 < ?f n" by (simp add: ac_simps divide_less_eq) qed have sums: "?f sums sin x" by (rule sin_paired [THEN sums_group]) simp show "0 < sin x" unfolding sums_unique [OF sums] using sums_summable [OF sums] pos by (simp add: suminf_pos) qed lemma cos_double_less_one: "0 < x \ x < 2 \ cos (2 * x) < 1" for x :: real using sin_gt_zero_02 [where x = x] by (auto simp: cos_squared_eq cos_double) lemma cos_paired: "(\n. (- 1) ^ n / (fact (2 * n)) * x ^ (2 * n)) sums cos x" for x :: real proof - have "(\n. \k = n * 2.. real" assumes f: "summable f" and fplus: "\d. 0 < f (k + (Suc(Suc 0) * d)) + f (k + ((Suc (Suc 0) * d) + 1))" shows "sum f {..n. \n = n * Suc (Suc 0)..n. f (n + k))" proof (rule sums_group) show "(\n. f (n + k)) sums (\n. f (n + k))" by (simp add: f summable_iff_shift summable_sums) qed auto with fplus have "0 < (\n. f (n + k))" apply (simp add: add.commute) apply (metis (no_types, lifting) suminf_pos summable_def sums_unique) done then show ?thesis by (simp add: f suminf_minus_initial_segment) qed lemma cos_two_less_zero [simp]: "cos 2 < (0::real)" proof - note fact_Suc [simp del] from sums_minus [OF cos_paired] have *: "(\n. - ((- 1) ^ n * 2 ^ (2 * n) / fact (2 * n))) sums - cos (2::real)" by simp then have sm: "summable (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_summable) have "0 < (\nnn. - ((- 1) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" proof - { fix d let ?six4d = "Suc (Suc (Suc (Suc (Suc (Suc (4 * d))))))" have "(4::real) * (fact (?six4d)) < (Suc (Suc (?six4d)) * fact (Suc (?six4d)))" unfolding of_nat_mult by (rule mult_strict_mono) (simp_all add: fact_less_mono) then have "(4::real) * (fact (?six4d)) < (fact (Suc (Suc (?six4d))))" by (simp only: fact_Suc [of "Suc (?six4d)"] of_nat_mult of_nat_fact) then have "(4::real) * inverse (fact (Suc (Suc (?six4d)))) < inverse (fact (?six4d))" by (simp add: inverse_eq_divide less_divide_eq) } then show ?thesis by (force intro!: sum_pos_lt_pair [OF sm] simp add: divide_inverse algebra_simps) qed ultimately have "0 < (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule order_less_trans) moreover from * have "- cos 2 = (\n. - ((- 1::real) ^ n * 2 ^ (2 * n) / (fact (2 * n))))" by (rule sums_unique) ultimately have "(0::real) < - cos 2" by simp then show ?thesis by simp qed lemmas cos_two_neq_zero [simp] = cos_two_less_zero [THEN less_imp_neq] lemmas cos_two_le_zero [simp] = cos_two_less_zero [THEN order_less_imp_le] lemma cos_is_zero: "\!x::real. 0 \ x \ x \ 2 \ cos x = 0" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ 2 \ cos x = 0" by (rule IVT2) simp_all next fix a b :: real assume ab: "0 \ a \ a \ 2 \ cos a = 0" "0 \ b \ b \ 2 \ cos b = 0" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero_02) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero_02) qed auto qed lemma pi_half: "pi/2 = (THE x. 0 \ x \ x \ 2 \ cos x = 0)" by (simp add: pi_def) lemma cos_pi_half [simp]: "cos (pi/2) = 0" by (simp add: pi_half cos_is_zero [THEN theI']) lemma cos_of_real_pi_half [simp]: "cos ((of_real pi/2) :: 'a) = 0" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" by (metis cos_pi_half cos_of_real eq_numeral_simps(4) nonzero_of_real_divide of_real_0 of_real_numeral) lemma pi_half_gt_zero [simp]: "0 < pi/2" proof - have "0 \ pi/2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_zero less_eq_real_def one_neq_zero) qed lemmas pi_half_neq_zero [simp] = pi_half_gt_zero [THEN less_imp_neq, symmetric] lemmas pi_half_ge_zero [simp] = pi_half_gt_zero [THEN order_less_imp_le] lemma pi_half_less_two [simp]: "pi/2 < 2" proof - have "pi/2 \ 2" by (simp add: pi_half cos_is_zero [THEN theI']) then show ?thesis by (metis cos_pi_half cos_two_neq_zero le_less) qed lemmas pi_half_neq_two [simp] = pi_half_less_two [THEN less_imp_neq] lemmas pi_half_le_two [simp] = pi_half_less_two [THEN order_less_imp_le] lemma pi_gt_zero [simp]: "0 < pi" using pi_half_gt_zero by simp lemma pi_ge_zero [simp]: "0 \ pi" by (rule pi_gt_zero [THEN order_less_imp_le]) lemma pi_neq_zero [simp]: "pi \ 0" by (rule pi_gt_zero [THEN less_imp_neq, symmetric]) lemma pi_not_less_zero [simp]: "\ pi < 0" by (simp add: linorder_not_less) lemma minus_pi_half_less_zero: "-(pi/2) < 0" by simp lemma m2pi_less_pi: "- (2*pi) < pi" by simp lemma sin_pi_half [simp]: "sin(pi/2) = 1" using sin_cos_squared_add2 [where x = "pi/2"] using sin_gt_zero_02 [OF pi_half_gt_zero pi_half_less_two] by (simp add: power2_eq_1_iff) lemma sin_of_real_pi_half [simp]: "sin ((of_real pi/2) :: 'a) = 1" if "SORT_CONSTRAINT('a::{real_field,banach,real_normed_algebra_1})" using sin_pi_half by (metis sin_pi_half eq_numeral_simps(4) nonzero_of_real_divide of_real_1 of_real_numeral sin_of_real) lemma sin_cos_eq: "sin x = cos (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_diff) lemma minus_sin_cos_eq: "- sin x = cos (x + of_real pi/2)" for x :: "'a::{real_normed_field,banach}" by (simp add: cos_add nonzero_of_real_divide) lemma cos_sin_eq: "cos x = sin (of_real pi/2 - x)" for x :: "'a::{real_normed_field,banach}" using sin_cos_eq [of "of_real pi/2 - x"] by simp lemma sin_add: "sin (x + y) = sin x * cos y + cos x * sin y" for x :: "'a::{real_normed_field,banach}" using cos_add [of "of_real pi/2 - x" "-y"] by (simp add: cos_sin_eq) (simp add: sin_cos_eq) lemma sin_diff: "sin (x - y) = sin x * cos y - cos x * sin y" for x :: "'a::{real_normed_field,banach}" using sin_add [of x "- y"] by simp lemma sin_double: "sin(2 * x) = 2 * sin x * cos x" for x :: "'a::{real_normed_field,banach}" using sin_add [where x=x and y=x] by simp lemma cos_of_real_pi [simp]: "cos (of_real pi) = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by (simp add: cos_of_real) lemma sin_of_real_pi [simp]: "sin (of_real pi) = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by (simp add: sin_of_real) lemma cos_pi [simp]: "cos pi = -1" using cos_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_pi [simp]: "sin pi = 0" using sin_add [where x = "pi/2" and y = "pi/2"] by simp lemma sin_periodic_pi [simp]: "sin (x + pi) = - sin x" by (simp add: sin_add) lemma sin_periodic_pi2 [simp]: "sin (pi + x) = - sin x" by (simp add: sin_add) lemma cos_periodic_pi [simp]: "cos (x + pi) = - cos x" by (simp add: cos_add) lemma cos_periodic_pi2 [simp]: "cos (pi + x) = - cos x" by (simp add: cos_add) lemma sin_periodic [simp]: "sin (x + 2 * pi) = sin x" by (simp add: sin_add sin_double cos_double) lemma cos_periodic [simp]: "cos (x + 2 * pi) = cos x" by (simp add: cos_add sin_double cos_double) lemma cos_npi [simp]: "cos (real n * pi) = (- 1) ^ n" by (induct n) (auto simp: distrib_right) lemma cos_npi2 [simp]: "cos (pi * real n) = (- 1) ^ n" by (metis cos_npi mult.commute) lemma sin_npi [simp]: "sin (real n * pi) = 0" for n :: nat by (induct n) (auto simp: distrib_right) lemma sin_npi2 [simp]: "sin (pi * real n) = 0" for n :: nat by (simp add: mult.commute [of pi]) lemma cos_two_pi [simp]: "cos (2 * pi) = 1" by (simp add: cos_double) lemma sin_two_pi [simp]: "sin (2 * pi) = 0" by (simp add: sin_double) context fixes w :: "'a::{real_normed_field,banach}" begin lemma sin_times_sin: "sin w * sin z = (cos (w - z) - cos (w + z)) / 2" by (simp add: cos_diff cos_add) lemma sin_times_cos: "sin w * cos z = (sin (w + z) + sin (w - z)) / 2" by (simp add: sin_diff sin_add) lemma cos_times_sin: "cos w * sin z = (sin (w + z) - sin (w - z)) / 2" by (simp add: sin_diff sin_add) lemma cos_times_cos: "cos w * cos z = (cos (w - z) + cos (w + z)) / 2" by (simp add: cos_diff cos_add) lemma cos_double_cos: "cos (2 * w) = 2 * cos w ^ 2 - 1" by (simp add: cos_double sin_squared_eq) lemma cos_double_sin: "cos (2 * w) = 1 - 2 * sin w ^ 2" by (simp add: cos_double sin_squared_eq) end lemma sin_plus_sin: "sin w + sin z = 2 * sin ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma sin_diff_sin: "sin w - sin z = 2 * sin ((w - z) / 2) * cos ((w + z) / 2)" for w :: "'a::{real_normed_field,banach}" apply (simp add: mult.assoc sin_times_cos) apply (simp add: field_simps) done lemma cos_plus_cos: "cos w + cos z = 2 * cos ((w + z) / 2) * cos ((w - z) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc cos_times_cos) apply (simp add: field_simps) done lemma cos_diff_cos: "cos w - cos z = 2 * sin ((w + z) / 2) * sin ((z - w) / 2)" for w :: "'a::{real_normed_field,banach,field}" apply (simp add: mult.assoc sin_times_sin) apply (simp add: field_simps) done lemma sin_pi_minus [simp]: "sin (pi - x) = sin x" by (metis sin_minus sin_periodic_pi minus_minus uminus_add_conv_diff) lemma cos_pi_minus [simp]: "cos (pi - x) = - (cos x)" by (metis cos_minus cos_periodic_pi uminus_add_conv_diff) lemma sin_minus_pi [simp]: "sin (x - pi) = - (sin x)" by (simp add: sin_diff) lemma cos_minus_pi [simp]: "cos (x - pi) = - (cos x)" by (simp add: cos_diff) lemma sin_2pi_minus [simp]: "sin (2 * pi - x) = - (sin x)" by (metis sin_periodic_pi2 add_diff_eq mult_2 sin_pi_minus) lemma cos_2pi_minus [simp]: "cos (2 * pi - x) = cos x" by (metis (no_types, opaque_lifting) cos_add cos_minus cos_two_pi sin_minus sin_two_pi diff_0_right minus_diff_eq mult_1 mult_zero_left uminus_add_conv_diff) lemma sin_gt_zero2: "0 < x \ x < pi/2 \ 0 < sin x" by (metis sin_gt_zero_02 order_less_trans pi_half_less_two) lemma sin_less_zero: assumes "- pi/2 < x" and "x < 0" shows "sin x < 0" proof - have "0 < sin (- x)" using assms by (simp only: sin_gt_zero2) then show ?thesis by simp qed lemma pi_less_4: "pi < 4" using pi_half_less_two by auto lemma cos_gt_zero: "0 < x \ x < pi/2 \ 0 < cos x" by (simp add: cos_sin_eq sin_gt_zero2) lemma cos_gt_zero_pi: "-(pi/2) < x \ x < pi/2 \ 0 < cos x" using cos_gt_zero [of x] cos_gt_zero [of "-x"] by (cases rule: linorder_cases [of x 0]) auto lemma cos_ge_zero: "-(pi/2) \ x \ x \ pi/2 \ 0 \ cos x" by (auto simp: order_le_less cos_gt_zero_pi) (metis cos_pi_half eq_divide_eq eq_numeral_simps(4)) lemma sin_gt_zero: "0 < x \ x < pi \ 0 < sin x" by (simp add: sin_cos_eq cos_gt_zero_pi) lemma sin_lt_zero: "pi < x \ x < 2 * pi \ sin x < 0" using sin_gt_zero [of "x - pi"] by (simp add: sin_diff) lemma pi_ge_two: "2 \ pi" proof (rule ccontr) assume "\ ?thesis" then have "pi < 2" by auto have "\y > pi. y < 2 \ y < 2 * pi" proof (cases "2 < 2 * pi") case True with dense[OF \pi < 2\] show ?thesis by auto next case False have "pi < 2 * pi" by auto from dense[OF this] and False show ?thesis by auto qed then obtain y where "pi < y" and "y < 2" and "y < 2 * pi" by blast then have "0 < sin y" using sin_gt_zero_02 by auto moreover have "sin y < 0" using sin_gt_zero[of "y - pi"] \pi < y\ and \y < 2 * pi\ sin_periodic_pi[of "y - pi"] by auto ultimately show False by auto qed lemma sin_ge_zero: "0 \ x \ x \ pi \ 0 \ sin x" by (auto simp: order_le_less sin_gt_zero) lemma sin_le_zero: "pi \ x \ x < 2 * pi \ sin x \ 0" using sin_ge_zero [of "x - pi"] by (simp add: sin_diff) lemma sin_pi_divide_n_ge_0 [simp]: assumes "n \ 0" shows "0 \ sin (pi / real n)" by (rule sin_ge_zero) (use assms in \simp_all add: field_split_simps\) lemma sin_pi_divide_n_gt_0: assumes "2 \ n" shows "0 < sin (pi / real n)" by (rule sin_gt_zero) (use assms in \simp_all add: field_split_simps\) text\Proof resembles that of \cos_is_zero\ but with \<^term>\pi\ for the upper bound\ lemma cos_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. 0 \ x \ x \ pi \ cos x = y" proof (rule ex_ex1I) show "\x::real. 0 \ x \ x \ pi \ cos x = y" by (rule IVT2) (simp_all add: y) next fix a b :: real assume ab: "0 \ a \ a \ pi \ cos a = y" "0 \ b \ b \ pi \ cos b = y" have cosd: "\x::real. cos differentiable (at x)" unfolding real_differentiable_def by (auto intro: DERIV_cos) show "a = b" proof (cases a b rule: linorder_cases) case less then obtain z where "a < z" "z < b" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \a < z\ \z < b\ ab order_less_le_trans less_le sin_gt_zero) next case greater then obtain z where "b < z" "z < a" "(cos has_real_derivative 0) (at z)" using Rolle by (metis cosd continuous_on_cos_real ab) then have "sin z = 0" using DERIV_cos DERIV_unique neg_equal_0_iff_equal by blast then show ?thesis by (metis \b < z\ \z < a\ ab order_less_le_trans less_le sin_gt_zero) qed auto qed lemma sin_total: assumes y: "-1 \ y" "y \ 1" shows "\!x. - (pi/2) \ x \ x \ pi/2 \ sin x = y" proof - from cos_total [OF y] obtain x where x: "0 \ x" "x \ pi" "cos x = y" and uniq: "\x'. 0 \ x' \ x' \ pi \ cos x' = y \ x' = x " by blast show ?thesis unfolding sin_cos_eq proof (rule ex1I [where a="pi/2 - x"]) show "- (pi/2) \ z \ z \ pi/2 \ cos (of_real pi/2 - z) = y \ z = pi/2 - x" for z using uniq [of "pi/2 -z"] by auto qed (use x in auto) qed lemma cos_zero_lemma: assumes "0 \ x" "cos x = 0" shows "\n. odd n \ x = of_nat n * (pi/2)" proof - have xle: "x < (1 + real_of_int \x/pi\) * pi" using floor_correct [of "x/pi"] by (simp add: add.commute divide_less_eq) obtain n where "real n * pi \ x" "x < real (Suc n) * pi" proof show "real (nat \x / pi\) * pi \ x" using assms floor_divide_lower [of pi x] by auto show "x < real (Suc (nat \x / pi\)) * pi" using assms floor_divide_upper [of pi x] by (simp add: xle) qed then have x: "0 \ x - n * pi" "(x - n * pi) \ pi" "cos (x - n * pi) = 0" by (auto simp: algebra_simps cos_diff assms) then have "\!x. 0 \ x \ x \ pi \ cos x = 0" by (auto simp: intro!: cos_total) then obtain \ where \: "0 \ \" "\ \ pi" "cos \ = 0" and uniq: "\\. 0 \ \ \ \ \ pi \ cos \ = 0 \ \ = \" by blast then have "x - real n * pi = \" using x by blast moreover have "pi/2 = \" using pi_half_ge_zero uniq by fastforce ultimately show ?thesis by (rule_tac x = "Suc (2 * n)" in exI) (simp add: algebra_simps) qed lemma sin_zero_lemma: assumes "0 \ x" "sin x = 0" shows "\n::nat. even n \ x = real n * (pi/2)" proof - obtain n where "odd n" and n: "x + pi/2 = of_nat n * (pi/2)" "n > 0" using cos_zero_lemma [of "x + pi/2"] assms by (auto simp add: cos_add) then have "x = real (n - 1) * (pi / 2)" by (simp add: algebra_simps of_nat_diff) then show ?thesis by (simp add: \odd n\) qed lemma cos_zero_iff: "cos x = 0 \ ((\n. odd n \ x = real n * (pi/2)) \ (\n. odd n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof - have *: "cos (real n * pi/2) = 0" if "odd n" for n :: nat proof - from that obtain m where "n = 2 * m + 1" .. then show ?thesis by (simp add: field_simps) (simp add: cos_add add_divide_distrib) qed show ?thesis proof show ?rhs if ?lhs using that cos_zero_lemma [of x] cos_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto dest: * simp del: eq_divide_eq_numeral1) qed qed lemma sin_zero_iff: "sin x = 0 \ ((\n. even n \ x = real n * (pi/2)) \ (\n. even n \ x = - (real n * (pi/2))))" (is "?lhs = ?rhs") proof show ?rhs if ?lhs using that sin_zero_lemma [of x] sin_zero_lemma [of "-x"] by force show ?lhs if ?rhs using that by (auto elim: evenE) qed lemma sin_zero_pi_iff: fixes x::real assumes "\x\ < pi" shows "sin x = 0 \ x = 0" proof show "x = 0" if "sin x = 0" using that assms by (auto simp: sin_zero_iff) qed auto lemma cos_zero_iff_int: "cos x = 0 \ (\i. odd i \ x = of_int i * (pi/2))" proof - have 1: "\n. odd n \ \i. odd i \ real n = real_of_int i" by (metis even_of_nat_iff of_int_of_nat_eq) have 2: "\n. odd n \ \i. odd i \ - (real n * pi) = real_of_int i * pi" by (metis even_minus even_of_nat_iff mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) have 3: "\odd i; \n. even n \ real_of_int i \ - (real n)\ \ \n. odd n \ real_of_int i = real n" for i by (cases i rule: int_cases2) auto show ?thesis by (force simp: cos_zero_iff intro!: 1 2 3) qed lemma sin_zero_iff_int: "sin x = 0 \ (\i. even i \ x = of_int i * (pi/2))" (is "?lhs = ?rhs") proof safe assume ?lhs then consider (plus) n where "even n" "x = real n * (pi/2)" | (minus) n where "even n" "x = - (real n * (pi/2))" using sin_zero_iff by auto then show "\n. even n \ x = of_int n * (pi/2)" proof cases case plus then show ?rhs by (metis even_of_nat_iff of_int_of_nat_eq) next case minus then show ?thesis by (rule_tac x="- (int n)" in exI) simp qed next fix i :: int assume "even i" then show "sin (of_int i * (pi/2)) = 0" by (cases i rule: int_cases2, simp_all add: sin_zero_iff) qed lemma sin_zero_iff_int2: "sin x = 0 \ (\i::int. x = of_int i * pi)" proof - have "sin x = 0 \ (\i. even i \ x = real_of_int i * (pi / 2))" by (auto simp: sin_zero_iff_int) also have "... = (\j. x = real_of_int (2*j) * (pi / 2))" using dvd_triv_left by blast also have "... = (\i::int. x = of_int i * pi)" by auto finally show ?thesis . qed lemma sin_npi_int [simp]: "sin (pi * of_int n) = 0" by (simp add: sin_zero_iff_int2) lemma cos_monotone_0_pi: assumes "0 \ y" and "y < x" and "x \ pi" shows "cos x < cos y" proof - have "- (x - y) < 0" using assms by auto from MVT2[OF \y < x\ DERIV_cos] obtain z where "y < z" and "z < x" and cos_diff: "cos x - cos y = (x - y) * - sin z" by auto then have "0 < z" and "z < pi" using assms by auto then have "0 < sin z" using sin_gt_zero by auto then have "cos x - cos y < 0" unfolding cos_diff minus_mult_commute[symmetric] using \- (x - y) < 0\ by (rule mult_pos_neg2) then show ?thesis by auto qed lemma cos_monotone_0_pi_le: assumes "0 \ y" and "y \ x" and "x \ pi" shows "cos x \ cos y" proof (cases "y < x") case True show ?thesis using cos_monotone_0_pi[OF \0 \ y\ True \x \ pi\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma cos_monotone_minus_pi_0: assumes "- pi \ y" and "y < x" and "x \ 0" shows "cos y < cos x" proof - have "0 \ - x" and "- x < - y" and "- y \ pi" using assms by auto from cos_monotone_0_pi[OF this] show ?thesis unfolding cos_minus . qed lemma cos_monotone_minus_pi_0': assumes "- pi \ y" and "y \ x" and "x \ 0" shows "cos y \ cos x" proof (cases "y < x") case True show ?thesis using cos_monotone_minus_pi_0[OF \-pi \ y\ True \x \ 0\] by auto next case False then have "y = x" using \y \ x\ by auto then show ?thesis by auto qed lemma sin_monotone_2pi: assumes "- (pi/2) \ y" and "y < x" and "x \ pi/2" shows "sin y < sin x" unfolding sin_cos_eq using assms by (auto intro: cos_monotone_0_pi) lemma sin_monotone_2pi_le: assumes "- (pi/2) \ y" and "y \ x" and "x \ pi/2" shows "sin y \ sin x" by (metis assms le_less sin_monotone_2pi) lemma sin_x_le_x: fixes x :: real assumes "x \ 0" shows "sin x \ x" proof - let ?f = "\x. x - sin x" have "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 - cos u) (at u)" by (auto intro!: derivative_eq_intros simp: field_simps) then have "?f x \ ?f 0" by (metis cos_le_one diff_ge_0_iff_ge DERIV_nonneg_imp_nondecreasing [OF assms]) then show "sin x \ x" by simp qed lemma sin_x_ge_neg_x: fixes x :: real assumes x: "x \ 0" shows "sin x \ - x" proof - let ?f = "\x. x + sin x" have \
: "\u. \0 \ u; u \ x\ \ \y. (?f has_real_derivative 1 + cos u) (at u)" by (auto intro!: derivative_eq_intros simp: field_simps) have "?f x \ ?f 0" by (rule DERIV_nonneg_imp_nondecreasing [OF assms]) (use \
real_0_le_add_iff in force) then show "sin x \ -x" by simp qed lemma abs_sin_x_le_abs_x: "\sin x\ \ \x\" for x :: real using sin_x_ge_neg_x [of x] sin_x_le_x [of x] sin_x_ge_neg_x [of "-x"] sin_x_le_x [of "-x"] by (auto simp: abs_real_def) subsection \More Corollaries about Sine and Cosine\ lemma sin_cos_npi [simp]: "sin (real (Suc (2 * n)) * pi/2) = (-1) ^ n" proof - have "sin ((real n + 1/2) * pi) = cos (real n * pi)" by (auto simp: algebra_simps sin_add) then show ?thesis by (simp add: distrib_right add_divide_distrib add.commute mult.commute [of pi]) qed lemma cos_2npi [simp]: "cos (2 * real n * pi) = 1" for n :: nat by (cases "even n") (simp_all add: cos_double mult.assoc) lemma cos_3over2_pi [simp]: "cos (3/2*pi) = 0" proof - have "cos (3/2*pi) = cos (pi + pi/2)" by simp also have "... = 0" by (subst cos_add, simp) finally show ?thesis . qed lemma sin_2npi [simp]: "sin (2 * real n * pi) = 0" for n :: nat by (auto simp: mult.assoc sin_double) lemma sin_3over2_pi [simp]: "sin (3/2*pi) = - 1" proof - have "sin (3/2*pi) = sin (pi + pi/2)" by simp also have "... = -1" by (subst sin_add, simp) finally show ?thesis . qed lemma cos_pi_eq_zero [simp]: "cos (pi * real (Suc (2 * m)) / 2) = 0" by (simp only: cos_add sin_add of_nat_Suc distrib_right distrib_left add_divide_distrib, auto) lemma DERIV_cos_add [simp]: "DERIV (\x. cos (x + k)) xa :> - sin (xa + k)" by (auto intro!: derivative_eq_intros) lemma sin_zero_norm_cos_one: fixes x :: "'a::{real_normed_field,banach}" assumes "sin x = 0" shows "norm (cos x) = 1" using sin_cos_squared_add [of x, unfolded assms] by (simp add: square_norm_one) lemma sin_zero_abs_cos_one: "sin x = 0 \ \cos x\ = (1::real)" using sin_zero_norm_cos_one by fastforce lemma cos_one_sin_zero: fixes x :: "'a::{real_normed_field,banach}" assumes "cos x = 1" shows "sin x = 0" using sin_cos_squared_add [of x, unfolded assms] by simp lemma sin_times_pi_eq_0: "sin (x * pi) = 0 \ x \ \" by (simp add: sin_zero_iff_int2) (metis Ints_cases Ints_of_int) lemma cos_one_2pi: "cos x = 1 \ (\n::nat. x = n * 2 * pi) \ (\n::nat. x = - (n * 2 * pi))" (is "?lhs = ?rhs") proof assume ?lhs then have "sin x = 0" by (simp add: cos_one_sin_zero) then show ?rhs proof (simp only: sin_zero_iff, elim exE disjE conjE) fix n :: nat assume n: "even n" "x = real n * (pi/2)" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) next fix n :: nat assume n: "even n" "x = - (real n * (pi/2))" then obtain m where m: "n = 2 * m" using dvdE by blast then have me: "even m" using \?lhs\ n by (auto simp: field_simps) (metis one_neq_neg_one power_minus_odd power_one) show ?rhs using m me n by (auto simp: field_simps elim!: evenE) qed next assume ?rhs then show "cos x = 1" by (metis cos_2npi cos_minus mult.assoc mult.left_commute) qed lemma cos_one_2pi_int: "cos x = 1 \ (\n::int. x = n * 2 * pi)" (is "?lhs = ?rhs") proof assume "cos x = 1" then show ?rhs by (metis cos_one_2pi mult.commute mult_minus_right of_int_minus of_int_of_nat_eq) next assume ?rhs then show "cos x = 1" by (clarsimp simp add: cos_one_2pi) (metis mult_minus_right of_int_of_nat) qed lemma cos_npi_int [simp]: fixes n::int shows "cos (pi * of_int n) = (if even n then 1 else -1)" by (auto simp: algebra_simps cos_one_2pi_int elim!: oddE evenE) lemma sin_cos_sqrt: "0 \ sin x \ sin x = sqrt (1 - (cos(x) ^ 2))" using sin_squared_eq real_sqrt_unique by fastforce lemma sin_eq_0_pi: "- pi < x \ x < pi \ sin x = 0 \ x = 0" by (metis sin_gt_zero sin_minus minus_less_iff neg_0_less_iff_less not_less_iff_gr_or_eq) lemma cos_treble_cos: "cos (3 * x) = 4 * cos x ^ 3 - 3 * cos x" for x :: "'a::{real_normed_field,banach}" proof - have *: "(sin x * (sin x * 3)) = 3 - (cos x * (cos x * 3))" by (simp add: mult.assoc [symmetric] sin_squared_eq [unfolded power2_eq_square]) have "cos(3 * x) = cos(2*x + x)" by simp also have "\ = 4 * cos x ^ 3 - 3 * cos x" unfolding cos_add cos_double sin_double by (simp add: * field_simps power2_eq_square power3_eq_cube) finally show ?thesis . qed lemma cos_45: "cos (pi / 4) = sqrt 2 / 2" proof - let ?c = "cos (pi / 4)" let ?s = "sin (pi / 4)" have nonneg: "0 \ ?c" by (simp add: cos_ge_zero) have "0 = cos (pi / 4 + pi / 4)" by simp also have "cos (pi / 4 + pi / 4) = ?c\<^sup>2 - ?s\<^sup>2" by (simp only: cos_add power2_eq_square) also have "\ = 2 * ?c\<^sup>2 - 1" by (simp add: sin_squared_eq) finally have "?c\<^sup>2 = (sqrt 2 / 2)\<^sup>2" by (simp add: power_divide) then show ?thesis using nonneg by (rule power2_eq_imp_eq) simp qed lemma cos_30: "cos (pi / 6) = sqrt 3/2" proof - let ?c = "cos (pi / 6)" let ?s = "sin (pi / 6)" have pos_c: "0 < ?c" by (rule cos_gt_zero) simp_all have "0 = cos (pi / 6 + pi / 6 + pi / 6)" by simp also have "\ = (?c * ?c - ?s * ?s) * ?c - (?s * ?c + ?c * ?s) * ?s" by (simp only: cos_add sin_add) also have "\ = ?c * (?c\<^sup>2 - 3 * ?s\<^sup>2)" by (simp add: algebra_simps power2_eq_square) finally have "?c\<^sup>2 = (sqrt 3/2)\<^sup>2" using pos_c by (simp add: sin_squared_eq power_divide) then show ?thesis using pos_c [THEN order_less_imp_le] by (rule power2_eq_imp_eq) simp qed lemma sin_45: "sin (pi / 4) = sqrt 2 / 2" by (simp add: sin_cos_eq cos_45) lemma sin_60: "sin (pi / 3) = sqrt 3/2" by (simp add: sin_cos_eq cos_30) lemma cos_60: "cos (pi / 3) = 1 / 2" proof - have "0 \ cos (pi / 3)" by (rule cos_ge_zero) (use pi_half_ge_zero in \linarith+\) then show ?thesis by (simp add: cos_squared_eq sin_60 power_divide power2_eq_imp_eq) qed lemma sin_30: "sin (pi / 6) = 1 / 2" by (simp add: sin_cos_eq cos_60) lemma cos_integer_2pi: "n \ \ \ cos(2 * pi * n) = 1" by (metis Ints_cases cos_one_2pi_int mult.assoc mult.commute) lemma sin_integer_2pi: "n \ \ \ sin(2 * pi * n) = 0" by (metis sin_two_pi Ints_mult mult.assoc mult.commute sin_times_pi_eq_0) lemma cos_int_2pin [simp]: "cos ((2 * pi) * of_int n) = 1" by (simp add: cos_one_2pi_int) lemma sin_int_2pin [simp]: "sin ((2 * pi) * of_int n) = 0" by (metis Ints_of_int sin_integer_2pi) lemma sincos_principal_value: "\y. (- pi < y \ y \ pi) \ (sin y = sin x \ cos y = cos x)" proof - define y where "y \ pi - (2 * pi) * frac ((pi - x) / (2 * pi))" have "-pi < y"" y \ pi" by (auto simp: field_simps frac_lt_1 y_def) moreover have "sin y = sin x" "cos y = cos x" unfolding y_def apply (simp_all add: frac_def divide_simps sin_add cos_add) by (metis sin_int_2pin cos_int_2pin diff_zero add.right_neutral mult.commute mult.left_neutral mult_zero_left)+ ultimately show ?thesis by metis qed subsection \Tangent\ definition tan :: "'a \ 'a::{real_normed_field,banach}" where "tan = (\x. sin x / cos x)" lemma tan_of_real: "of_real (tan x) = (tan (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: tan_def sin_of_real cos_of_real) lemma tan_in_Reals [simp]: "z \ \ \ tan z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: tan_def) lemma tan_zero [simp]: "tan 0 = 0" by (simp add: tan_def) lemma tan_pi [simp]: "tan pi = 0" by (simp add: tan_def) lemma tan_npi [simp]: "tan (real n * pi) = 0" for n :: nat by (simp add: tan_def) lemma tan_minus [simp]: "tan (- x) = - tan x" by (simp add: tan_def) lemma tan_periodic [simp]: "tan (x + 2 * pi) = tan x" by (simp add: tan_def) lemma lemma_tan_add1: "cos x \ 0 \ cos y \ 0 \ 1 - tan x * tan y = cos (x + y)/(cos x * cos y)" by (simp add: tan_def cos_add field_simps) lemma add_tan_eq: "cos x \ 0 \ cos y \ 0 \ tan x + tan y = sin(x + y)/(cos x * cos y)" for x :: "'a::{real_normed_field,banach}" by (simp add: tan_def sin_add field_simps) lemma tan_add: "cos x \ 0 \ cos y \ 0 \ cos (x + y) \ 0 \ tan (x + y) = (tan x + tan y)/(1 - tan x * tan y)" for x :: "'a::{real_normed_field,banach}" by (simp add: add_tan_eq lemma_tan_add1 field_simps) (simp add: tan_def) lemma tan_double: "cos x \ 0 \ cos (2 * x) \ 0 \ tan (2 * x) = (2 * tan x) / (1 - (tan x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x x] by (simp add: power2_eq_square) lemma tan_gt_zero: "0 < x \ x < pi/2 \ 0 < tan x" by (simp add: tan_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma tan_less_zero: assumes "- pi/2 < x" and "x < 0" shows "tan x < 0" proof - have "0 < tan (- x)" using assms by (simp only: tan_gt_zero) then show ?thesis by simp qed lemma tan_half: "tan x = sin (2 * x) / (cos (2 * x) + 1)" for x :: "'a::{real_normed_field,banach,field}" unfolding tan_def sin_double cos_double sin_squared_eq by (simp add: power2_eq_square) lemma tan_30: "tan (pi / 6) = 1 / sqrt 3" unfolding tan_def by (simp add: sin_30 cos_30) lemma tan_45: "tan (pi / 4) = 1" unfolding tan_def by (simp add: sin_45 cos_45) lemma tan_60: "tan (pi / 3) = sqrt 3" unfolding tan_def by (simp add: sin_60 cos_60) lemma DERIV_tan [simp]: "cos x \ 0 \ DERIV tan x :> inverse ((cos x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding tan_def by (auto intro!: derivative_eq_intros, simp add: divide_inverse power2_eq_square) declare DERIV_tan[THEN DERIV_chain2, derivative_intros] and DERIV_tan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_tan[derivative_intros] = DERIV_tan[THEN DERIV_compose_FDERIV] lemma isCont_tan: "cos x \ 0 \ isCont tan x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_tan [THEN DERIV_isCont]) lemma isCont_tan' [simp,continuous_intros]: fixes a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" shows "isCont f a \ cos (f a) \ 0 \ isCont (\x. tan (f x)) a" by (rule isCont_o2 [OF _ isCont_tan]) lemma tendsto_tan [tendsto_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "(f \ a) F \ cos a \ 0 \ ((\x. tan (f x)) \ tan a) F" by (rule isCont_tendsto_compose [OF isCont_tan]) lemma continuous_tan: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous F f \ cos (f (Lim F (\x. x))) \ 0 \ continuous F (\x. tan (f x))" unfolding continuous_def by (rule tendsto_tan) lemma continuous_on_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. cos (f x) \ 0) \ continuous_on s (\x. tan (f x))" unfolding continuous_on_def by (auto intro: tendsto_tan) lemma continuous_within_tan [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ cos (f x) \ 0 \ continuous (at x within s) (\x. tan (f x))" unfolding continuous_within by (rule tendsto_tan) lemma LIM_cos_div_sin: "(\x. cos(x)/sin(x)) \pi/2\ 0" by (rule tendsto_cong_limit, (rule tendsto_intros)+, simp_all) lemma lemma_tan_total: assumes "0 < y" shows "\x. 0 < x \ x < pi/2 \ y < tan x" proof - obtain s where "0 < s" and s: "\x. \x \ pi/2; norm (x - pi/2) < s\ \ norm (cos x / sin x - 0) < inverse y" using LIM_D [OF LIM_cos_div_sin, of "inverse y"] that assms by force obtain e where e: "0 < e" "e < s" "e < pi/2" using \0 < s\ field_lbound_gt_zero pi_half_gt_zero by blast show ?thesis proof (intro exI conjI) have "0 < sin e" "0 < cos e" using e by (auto intro: cos_gt_zero sin_gt_zero2 simp: mult.commute) then show "y < tan (pi/2 - e)" using s [of "pi/2 - e"] e assms by (simp add: tan_def sin_diff cos_diff) (simp add: field_simps split: if_split_asm) qed (use e in auto) qed lemma tan_total_pos: assumes "0 \ y" shows "\x. 0 \ x \ x < pi/2 \ tan x = y" proof (cases "y = 0") case True then show ?thesis using pi_half_gt_zero tan_zero by blast next case False with assms have "y > 0" by linarith obtain x where x: "0 < x" "x < pi/2" "y < tan x" using lemma_tan_total \0 < y\ by blast have "\u\0. u \ x \ tan u = y" proof (intro IVT allI impI) show "isCont tan u" if "0 \ u \ u \ x" for u proof - have "cos u \ 0" using antisym_conv2 cos_gt_zero that x(2) by fastforce with assms show ?thesis by (auto intro!: DERIV_tan [THEN DERIV_isCont]) qed qed (use assms x in auto) then show ?thesis using x(2) by auto qed lemma lemma_tan_total1: "\x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof (cases "0::real" y rule: le_cases) case le then show ?thesis by (meson less_le_trans minus_pi_half_less_zero tan_total_pos) next case ge with tan_total_pos [of "-y"] obtain x where "0 \ x" "x < pi / 2" "tan x = - y" by force then show ?thesis by (rule_tac x="-x" in exI) auto qed proposition tan_total: "\! x. -(pi/2) < x \ x < (pi/2) \ tan x = y" proof - have "u = v" if u: "- (pi / 2) < u" "u < pi / 2" and v: "- (pi / 2) < v" "v < pi / 2" and eq: "tan u = tan v" for u v proof (cases u v rule: linorder_cases) case less have "\x. u \ x \ x \ v \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(1) v(2)) then have "continuous_on {u..v} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. u < x \ x < v \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(1) v(2)) ultimately obtain z where "u < z" "z < v" "DERIV tan z :> 0" by (metis less Rolle eq) moreover have "cos z \ 0" by (metis (no_types) \u < z\ \z < v\ cos_gt_zero_pi less_le_trans linorder_not_less not_less_iff_gr_or_eq u(1) v(2)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce next case greater have "\x. v \ x \ x \ u \ isCont tan x" by (metis cos_gt_zero_pi isCont_tan le_less_trans less_irrefl less_le_trans u(2) v(1)) then have "continuous_on {v..u} tan" by (simp add: continuous_at_imp_continuous_on) moreover have "\x. v < x \ x < u \ tan differentiable (at x)" by (metis DERIV_tan cos_gt_zero_pi real_differentiable_def less_numeral_extra(3) order.strict_trans u(2) v(1)) ultimately obtain z where "v < z" "z < u" "DERIV tan z :> 0" by (metis greater Rolle eq) moreover have "cos z \ 0" by (metis \v < z\ \z < u\ cos_gt_zero_pi less_eq_real_def less_le_trans order_less_irrefl u(2) v(1)) ultimately show ?thesis using DERIV_unique [OF _ DERIV_tan] by fastforce qed auto then have "\!x. - (pi / 2) < x \ x < pi / 2 \ tan x = y" if x: "- (pi / 2) < x" "x < pi / 2" "tan x = y" for x using that by auto then show ?thesis using lemma_tan_total1 [where y = y] by auto qed lemma tan_monotone: assumes "- (pi/2) < y" and "y < x" and "x < pi/2" shows "tan y < tan x" proof - have "DERIV tan x' :> inverse ((cos x')\<^sup>2)" if "y \ x'" "x' \ x" for x' proof - have "-(pi/2) < x'" and "x' < pi/2" using that assms by auto with cos_gt_zero_pi have "cos x' \ 0" by force then show "DERIV tan x' :> inverse ((cos x')\<^sup>2)" by (rule DERIV_tan) qed from MVT2[OF \y < x\ this] obtain z where "y < z" and "z < x" and tan_diff: "tan x - tan y = (x - y) * inverse ((cos z)\<^sup>2)" by auto then have "- (pi/2) < z" and "z < pi/2" using assms by auto then have "0 < cos z" using cos_gt_zero_pi by auto then have inv_pos: "0 < inverse ((cos z)\<^sup>2)" by auto have "0 < x - y" using \y < x\ by auto with inv_pos have "0 < tan x - tan y" unfolding tan_diff by auto then show ?thesis by auto qed lemma tan_monotone': assumes "- (pi/2) < y" and "y < pi/2" and "- (pi/2) < x" and "x < pi/2" shows "y < x \ tan y < tan x" proof assume "y < x" then show "tan y < tan x" using tan_monotone and \- (pi/2) < y\ and \x < pi/2\ by auto next assume "tan y < tan x" show "y < x" proof (rule ccontr) assume "\ ?thesis" then have "x \ y" by auto then have "tan x \ tan y" proof (cases "x = y") case True then show ?thesis by auto next case False then have "x < y" using \x \ y\ by auto from tan_monotone[OF \- (pi/2) < x\ this \y < pi/2\] show ?thesis by auto qed then show False using \tan y < tan x\ by auto qed qed lemma tan_inverse: "1 / (tan y) = tan (pi/2 - y)" unfolding tan_def sin_cos_eq[of y] cos_sin_eq[of y] by auto lemma tan_periodic_pi[simp]: "tan (x + pi) = tan x" by (simp add: tan_def) lemma tan_periodic_nat[simp]: "tan (x + real n * pi) = tan x" for n :: nat proof (induct n arbitrary: x) case 0 then show ?case by simp next case (Suc n) have split_pi_off: "x + real (Suc n) * pi = (x + real n * pi) + pi" unfolding Suc_eq_plus1 of_nat_add distrib_right by auto show ?case unfolding split_pi_off using Suc by auto qed lemma tan_periodic_int[simp]: "tan (x + of_int i * pi) = tan x" proof (cases "0 \ i") case True then have i_nat: "of_int i = of_int (nat i)" by auto show ?thesis unfolding i_nat by (metis of_int_of_nat_eq tan_periodic_nat) next case False then have i_nat: "of_int i = - of_int (nat (- i))" by auto have "tan x = tan (x + of_int i * pi - of_int i * pi)" by auto also have "\ = tan (x + of_int i * pi)" unfolding i_nat mult_minus_left diff_minus_eq_add by (metis of_int_of_nat_eq tan_periodic_nat) finally show ?thesis by auto qed lemma tan_periodic_n[simp]: "tan (x + numeral n * pi) = tan x" using tan_periodic_int[of _ "numeral n" ] by simp lemma tan_minus_45: "tan (-(pi/4)) = -1" unfolding tan_def by (simp add: sin_45 cos_45) lemma tan_diff: "cos x \ 0 \ cos y \ 0 \ cos (x - y) \ 0 \ tan (x - y) = (tan x - tan y)/(1 + tan x * tan y)" for x :: "'a::{real_normed_field,banach}" using tan_add [of x "-y"] by simp lemma tan_pos_pi2_le: "0 \ x \ x < pi/2 \ 0 \ tan x" using less_eq_real_def tan_gt_zero by auto lemma cos_tan: "\x\ < pi/2 \ cos x = 1 / sqrt (1 + tan x ^ 2)" using cos_gt_zero_pi [of x] by (simp add: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) +lemma cos_tan_half: "cos x \0 \ cos (2*x) = (1 - (tan x)^2) / (1 + (tan x)^2)" + unfolding cos_double tan_def by (auto simp add:field_simps ) + lemma sin_tan: "\x\ < pi/2 \ sin x = tan x / sqrt (1 + tan x ^ 2)" using cos_gt_zero [of "x"] cos_gt_zero [of "-x"] by (force simp: field_split_simps tan_def real_sqrt_divide abs_if split: if_split_asm) +lemma sin_tan_half: "sin (2*x) = 2 * tan x / (1 + (tan x)^2)" + unfolding sin_double tan_def + by (cases "cos x=0") (auto simp add:field_simps power2_eq_square) + lemma tan_mono_le: "-(pi/2) < x \ x \ y \ y < pi/2 \ tan x \ tan y" using less_eq_real_def tan_monotone by auto lemma tan_mono_lt_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x < tan y \ x < y" using tan_monotone' by blast lemma tan_mono_le_eq: "-(pi/2) < x \ x < pi/2 \ -(pi/2) < y \ y < pi/2 \ tan x \ tan y \ x \ y" by (meson tan_mono_le not_le tan_monotone) lemma tan_bound_pi2: "\x\ < pi/4 \ \tan x\ < 1" using tan_45 tan_monotone [of x "pi/4"] tan_monotone [of "-x" "pi/4"] by (auto simp: abs_if split: if_split_asm) lemma tan_cot: "tan(pi/2 - x) = inverse(tan x)" by (simp add: tan_def sin_diff cos_diff) subsection \Cotangent\ definition cot :: "'a \ 'a::{real_normed_field,banach}" where "cot = (\x. cos x / sin x)" lemma cot_of_real: "of_real (cot x) = (cot (of_real x) :: 'a::{real_normed_field,banach})" by (simp add: cot_def sin_of_real cos_of_real) lemma cot_in_Reals [simp]: "z \ \ \ cot z \ \" for z :: "'a::{real_normed_field,banach}" by (simp add: cot_def) lemma cot_zero [simp]: "cot 0 = 0" by (simp add: cot_def) lemma cot_pi [simp]: "cot pi = 0" by (simp add: cot_def) lemma cot_npi [simp]: "cot (real n * pi) = 0" for n :: nat by (simp add: cot_def) lemma cot_minus [simp]: "cot (- x) = - cot x" by (simp add: cot_def) lemma cot_periodic [simp]: "cot (x + 2 * pi) = cot x" by (simp add: cot_def) lemma cot_altdef: "cot x = inverse (tan x)" by (simp add: cot_def tan_def) lemma tan_altdef: "tan x = inverse (cot x)" by (simp add: cot_def tan_def) lemma tan_cot': "tan (pi/2 - x) = cot x" by (simp add: tan_cot cot_altdef) lemma cot_gt_zero: "0 < x \ x < pi/2 \ 0 < cot x" by (simp add: cot_def zero_less_divide_iff sin_gt_zero2 cos_gt_zero_pi) lemma cot_less_zero: assumes lb: "- pi/2 < x" and "x < 0" shows "cot x < 0" proof - have "0 < cot (- x)" using assms by (simp only: cot_gt_zero) then show ?thesis by simp qed lemma DERIV_cot [simp]: "sin x \ 0 \ DERIV cot x :> -inverse ((sin x)\<^sup>2)" for x :: "'a::{real_normed_field,banach}" unfolding cot_def using cos_squared_eq[of x] by (auto intro!: derivative_eq_intros) (simp add: divide_inverse power2_eq_square) lemma isCont_cot: "sin x \ 0 \ isCont cot x" for x :: "'a::{real_normed_field,banach}" by (rule DERIV_cot [THEN DERIV_isCont]) lemma isCont_cot' [simp,continuous_intros]: "isCont f a \ sin (f a) \ 0 \ isCont (\x. cot (f x)) a" for a :: "'a::{real_normed_field,banach}" and f :: "'a \ 'a" by (rule isCont_o2 [OF _ isCont_cot]) lemma tendsto_cot [tendsto_intros]: "(f \ a) F \ sin a \ 0 \ ((\x. cot (f x)) \ cot a) F" for f :: "'a \ 'a::{real_normed_field,banach}" by (rule isCont_tendsto_compose [OF isCont_cot]) lemma continuous_cot: "continuous F f \ sin (f (Lim F (\x. x))) \ 0 \ continuous F (\x. cot (f x))" for f :: "'a \ 'a::{real_normed_field,banach}" unfolding continuous_def by (rule tendsto_cot) lemma continuous_on_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous_on s f \ (\x\s. sin (f x) \ 0) \ continuous_on s (\x. cot (f x))" unfolding continuous_on_def by (auto intro: tendsto_cot) lemma continuous_within_cot [continuous_intros]: fixes f :: "'a \ 'a::{real_normed_field,banach}" shows "continuous (at x within s) f \ sin (f x) \ 0 \ continuous (at x within s) (\x. cot (f x))" unfolding continuous_within by (rule tendsto_cot) subsection \Inverse Trigonometric Functions\ definition arcsin :: "real \ real" where "arcsin y = (THE x. -(pi/2) \ x \ x \ pi/2 \ sin x = y)" definition arccos :: "real \ real" where "arccos y = (THE x. 0 \ x \ x \ pi \ cos x = y)" definition arctan :: "real \ real" where "arctan y = (THE x. -(pi/2) < x \ x < pi/2 \ tan x = y)" lemma arcsin: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2 \ sin (arcsin y) = y" unfolding arcsin_def by (rule theI' [OF sin_total]) lemma arcsin_pi: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi \ sin (arcsin y) = y" by (drule (1) arcsin) (force intro: order_trans) lemma sin_arcsin [simp]: "- 1 \ y \ y \ 1 \ sin (arcsin y) = y" by (blast dest: arcsin) lemma arcsin_bounded: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lbound: "- 1 \ y \ y \ 1 \ - (pi/2) \ arcsin y" by (blast dest: arcsin) lemma arcsin_ubound: "- 1 \ y \ y \ 1 \ arcsin y \ pi/2" by (blast dest: arcsin) lemma arcsin_lt_bounded: assumes "- 1 < y" "y < 1" shows "- (pi/2) < arcsin y \ arcsin y < pi/2" proof - have "arcsin y \ pi/2" by (metis arcsin assms not_less not_less_iff_gr_or_eq sin_pi_half) moreover have "arcsin y \ - pi/2" by (metis arcsin assms minus_divide_left not_less not_less_iff_gr_or_eq sin_minus sin_pi_half) ultimately show ?thesis using arcsin_bounded [of y] assms by auto qed lemma arcsin_sin: "- (pi/2) \ x \ x \ pi/2 \ arcsin (sin x) = x" unfolding arcsin_def using the1_equality [OF sin_total] by simp lemma arcsin_0 [simp]: "arcsin 0 = 0" using arcsin_sin [of 0] by simp lemma arcsin_1 [simp]: "arcsin 1 = pi/2" using arcsin_sin [of "pi/2"] by simp lemma arcsin_minus_1 [simp]: "arcsin (- 1) = - (pi/2)" using arcsin_sin [of "- pi/2"] by simp lemma arcsin_minus: "- 1 \ x \ x \ 1 \ arcsin (- x) = - arcsin x" by (metis (no_types, opaque_lifting) arcsin arcsin_sin minus_minus neg_le_iff_le sin_minus) lemma arcsin_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arcsin x = arcsin y \ x = y" by (metis abs_le_iff arcsin minus_le_iff) lemma cos_arcsin_nonzero: "- 1 < x \ x < 1 \ cos (arcsin x) \ 0" using arcsin_lt_bounded cos_gt_zero_pi by force lemma arccos: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi \ cos (arccos y) = y" unfolding arccos_def by (rule theI' [OF cos_total]) lemma cos_arccos [simp]: "- 1 \ y \ y \ 1 \ cos (arccos y) = y" by (blast dest: arccos) lemma arccos_bounded: "- 1 \ y \ y \ 1 \ 0 \ arccos y \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lbound: "- 1 \ y \ y \ 1 \ 0 \ arccos y" by (blast dest: arccos) lemma arccos_ubound: "- 1 \ y \ y \ 1 \ arccos y \ pi" by (blast dest: arccos) lemma arccos_lt_bounded: assumes "- 1 < y" "y < 1" shows "0 < arccos y \ arccos y < pi" proof - have "arccos y \ 0" by (metis (no_types) arccos assms(1) assms(2) cos_zero less_eq_real_def less_irrefl) moreover have "arccos y \ -pi" by (metis arccos assms(1) assms(2) cos_minus cos_pi not_less not_less_iff_gr_or_eq) ultimately show ?thesis using arccos_bounded [of y] assms by (metis arccos cos_pi not_less not_less_iff_gr_or_eq) qed lemma arccos_cos: "0 \ x \ x \ pi \ arccos (cos x) = x" by (auto simp: arccos_def intro!: the1_equality cos_total) lemma arccos_cos2: "x \ 0 \ - pi \ x \ arccos (cos x) = -x" by (auto simp: arccos_def intro!: the1_equality cos_total) lemma cos_arcsin: assumes "- 1 \ x" "x \ 1" shows "cos (arcsin x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(cos (arcsin x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms cos_squared_eq) show "0 \ cos (arcsin x)" using arcsin assms cos_ge_zero by blast show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma sin_arccos: assumes "- 1 \ x" "x \ 1" shows "sin (arccos x) = sqrt (1 - x\<^sup>2)" proof (rule power2_eq_imp_eq) show "(sin (arccos x))\<^sup>2 = (sqrt (1 - x\<^sup>2))\<^sup>2" by (simp add: square_le_1 assms sin_squared_eq) show "0 \ sin (arccos x)" by (simp add: arccos_bounded assms sin_ge_zero) show "0 \ sqrt (1 - x\<^sup>2)" by (simp add: square_le_1 assms) qed lemma arccos_0 [simp]: "arccos 0 = pi/2" by (metis arccos_cos cos_gt_zero cos_pi cos_pi_half pi_gt_zero pi_half_ge_zero not_le not_zero_less_neg_numeral numeral_One) lemma arccos_1 [simp]: "arccos 1 = 0" using arccos_cos by force lemma arccos_minus_1 [simp]: "arccos (- 1) = pi" by (metis arccos_cos cos_pi order_refl pi_ge_zero) lemma arccos_minus: "-1 \ x \ x \ 1 \ arccos (- x) = pi - arccos x" by (metis arccos_cos arccos_cos2 cos_minus_pi cos_total diff_le_0_iff_le le_add_same_cancel1 minus_diff_eq uminus_add_conv_diff) corollary arccos_minus_abs: assumes "\x\ \ 1" shows "arccos (- x) = pi - arccos x" using assms by (simp add: arccos_minus) lemma sin_arccos_nonzero: "- 1 < x \ x < 1 \ sin (arccos x) \ 0" using arccos_lt_bounded sin_gt_zero by force lemma arctan: "- (pi/2) < arctan y \ arctan y < pi/2 \ tan (arctan y) = y" unfolding arctan_def by (rule theI' [OF tan_total]) lemma tan_arctan: "tan (arctan y) = y" by (simp add: arctan) lemma arctan_bounded: "- (pi/2) < arctan y \ arctan y < pi/2" by (auto simp only: arctan) lemma arctan_lbound: "- (pi/2) < arctan y" by (simp add: arctan) lemma arctan_ubound: "arctan y < pi/2" by (auto simp only: arctan) lemma arctan_unique: assumes "-(pi/2) < x" and "x < pi/2" and "tan x = y" shows "arctan y = x" using assms arctan [of y] tan_total [of y] by (fast elim: ex1E) lemma arctan_tan: "-(pi/2) < x \ x < pi/2 \ arctan (tan x) = x" by (rule arctan_unique) simp_all lemma arctan_zero_zero [simp]: "arctan 0 = 0" by (rule arctan_unique) simp_all lemma arctan_minus: "arctan (- x) = - arctan x" using arctan [of "x"] by (auto simp: arctan_unique) lemma cos_arctan_not_zero [simp]: "cos (arctan x) \ 0" by (intro less_imp_neq [symmetric] cos_gt_zero_pi arctan_lbound arctan_ubound) lemma cos_arctan: "cos (arctan x) = 1 / sqrt (1 + x\<^sup>2)" proof (rule power2_eq_imp_eq) have "0 < 1 + x\<^sup>2" by (simp add: add_pos_nonneg) show "0 \ 1 / sqrt (1 + x\<^sup>2)" by simp show "0 \ cos (arctan x)" by (intro less_imp_le cos_gt_zero_pi arctan_lbound arctan_ubound) have "(cos (arctan x))\<^sup>2 * (1 + (tan (arctan x))\<^sup>2) = 1" unfolding tan_def by (simp add: distrib_left power_divide) then show "(cos (arctan x))\<^sup>2 = (1 / sqrt (1 + x\<^sup>2))\<^sup>2" using \0 < 1 + x\<^sup>2\ by (simp add: arctan power_divide eq_divide_eq) qed lemma sin_arctan: "sin (arctan x) = x / sqrt (1 + x\<^sup>2)" using add_pos_nonneg [OF zero_less_one zero_le_power2 [of x]] using tan_arctan [of x] unfolding tan_def cos_arctan by (simp add: eq_divide_eq) lemma tan_sec: "cos x \ 0 \ 1 + (tan x)\<^sup>2 = (inverse (cos x))\<^sup>2" for x :: "'a::{real_normed_field,banach,field}" by (simp add: add_divide_eq_iff inverse_eq_divide power2_eq_square tan_def) lemma arctan_less_iff: "arctan x < arctan y \ x < y" by (metis tan_monotone' arctan_lbound arctan_ubound tan_arctan) lemma arctan_le_iff: "arctan x \ arctan y \ x \ y" by (simp only: not_less [symmetric] arctan_less_iff) lemma arctan_eq_iff: "arctan x = arctan y \ x = y" by (simp only: eq_iff [where 'a=real] arctan_le_iff) lemma zero_less_arctan_iff [simp]: "0 < arctan x \ 0 < x" using arctan_less_iff [of 0 x] by simp lemma arctan_less_zero_iff [simp]: "arctan x < 0 \ x < 0" using arctan_less_iff [of x 0] by simp lemma zero_le_arctan_iff [simp]: "0 \ arctan x \ 0 \ x" using arctan_le_iff [of 0 x] by simp lemma arctan_le_zero_iff [simp]: "arctan x \ 0 \ x \ 0" using arctan_le_iff [of x 0] by simp lemma arctan_eq_zero_iff [simp]: "arctan x = 0 \ x = 0" using arctan_eq_iff [of x 0] by simp lemma continuous_on_arcsin': "continuous_on {-1 .. 1} arcsin" proof - have "continuous_on (sin ` {- pi/2 .. pi/2}) arcsin" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arcsin_sin) also have "sin ` {- pi/2 .. pi/2} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ sin ` {- pi/2..pi/2}" using arcsin_lbound arcsin_ubound by (intro image_eqI[where x="arcsin x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arcsin [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arcsin (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arcsin']] by (auto simp: comp_def subset_eq) lemma isCont_arcsin: "-1 < x \ x < 1 \ isCont arcsin x" using continuous_on_arcsin'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma continuous_on_arccos': "continuous_on {-1 .. 1} arccos" proof - have "continuous_on (cos ` {0 .. pi}) arccos" by (rule continuous_on_inv) (auto intro: continuous_intros simp: arccos_cos) also have "cos ` {0 .. pi} = {-1 .. 1}" proof safe fix x :: real assume "x \ {-1..1}" then show "x \ cos ` {0..pi}" using arccos_lbound arccos_ubound by (intro image_eqI[where x="arccos x"]) auto qed simp finally show ?thesis . qed lemma continuous_on_arccos [continuous_intros]: "continuous_on s f \ (\x\s. -1 \ f x \ f x \ 1) \ continuous_on s (\x. arccos (f x))" using continuous_on_compose[of s f, OF _ continuous_on_subset[OF continuous_on_arccos']] by (auto simp: comp_def subset_eq) lemma isCont_arccos: "-1 < x \ x < 1 \ isCont arccos x" using continuous_on_arccos'[THEN continuous_on_subset, of "{ -1 <..< 1 }"] by (auto simp: continuous_on_eq_continuous_at subset_eq) lemma isCont_arctan: "isCont arctan x" proof - obtain u where u: "- (pi / 2) < u" "u < arctan x" by (meson arctan arctan_less_iff linordered_field_no_lb) obtain v where v: "arctan x < v" "v < pi / 2" by (meson arctan_less_iff arctan_ubound linordered_field_no_ub) have "isCont arctan (tan (arctan x))" proof (rule isCont_inverse_function2 [of u "arctan x" v]) show "\z. \u \ z; z \ v\ \ arctan (tan z) = z" using arctan_unique u(1) v(2) by auto then show "\z. \u \ z; z \ v\ \ isCont tan z" by (metis arctan cos_gt_zero_pi isCont_tan less_irrefl) qed (use u v in auto) then show ?thesis by (simp add: arctan) qed lemma tendsto_arctan [tendsto_intros]: "(f \ x) F \ ((\x. arctan (f x)) \ arctan x) F" by (rule isCont_tendsto_compose [OF isCont_arctan]) lemma continuous_arctan [continuous_intros]: "continuous F f \ continuous F (\x. arctan (f x))" unfolding continuous_def by (rule tendsto_arctan) lemma continuous_on_arctan [continuous_intros]: "continuous_on s f \ continuous_on s (\x. arctan (f x))" unfolding continuous_on_def by (auto intro: tendsto_arctan) lemma DERIV_arcsin: assumes "- 1 < x" "x < 1" shows "DERIV arcsin x :> inverse (sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(sin has_real_derivative sqrt (1 - x\<^sup>2)) (at (arcsin x))" by (rule derivative_eq_intros | use assms cos_arcsin in force)+ show "sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arcsin in auto) lemma DERIV_arccos: assumes "- 1 < x" "x < 1" shows "DERIV arccos x :> inverse (- sqrt (1 - x\<^sup>2))" proof (rule DERIV_inverse_function) show "(cos has_real_derivative - sqrt (1 - x\<^sup>2)) (at (arccos x))" by (rule derivative_eq_intros | use assms sin_arccos in force)+ show "- sqrt (1 - x\<^sup>2) \ 0" using abs_square_eq_1 assms by force qed (use assms isCont_arccos in auto) lemma DERIV_arctan: "DERIV arctan x :> inverse (1 + x\<^sup>2)" proof (rule DERIV_inverse_function) have "inverse ((cos (arctan x))\<^sup>2) = 1 + x\<^sup>2" by (metis arctan cos_arctan_not_zero power_inverse tan_sec) then show "(tan has_real_derivative 1 + x\<^sup>2) (at (arctan x))" by (auto intro!: derivative_eq_intros) show "\y. \x - 1 < y; y < x + 1\ \ tan (arctan y) = y" using tan_arctan by blast show "1 + x\<^sup>2 \ 0" by (metis power_one sum_power2_eq_zero_iff zero_neq_one) qed (use isCont_arctan in auto) declare DERIV_arcsin[THEN DERIV_chain2, derivative_intros] DERIV_arcsin[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arccos[THEN DERIV_chain2, derivative_intros] DERIV_arccos[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] DERIV_arctan[THEN DERIV_chain2, derivative_intros] DERIV_arctan[THEN DERIV_chain2, unfolded has_field_derivative_def, derivative_intros] lemmas has_derivative_arctan[derivative_intros] = DERIV_arctan[THEN DERIV_compose_FDERIV] and has_derivative_arccos[derivative_intros] = DERIV_arccos[THEN DERIV_compose_FDERIV] and has_derivative_arcsin[derivative_intros] = DERIV_arcsin[THEN DERIV_compose_FDERIV] lemma filterlim_tan_at_right: "filterlim tan at_bot (at_right (- (pi/2)))" by (rule filterlim_at_bot_at_right[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma filterlim_tan_at_left: "filterlim tan at_top (at_left (pi/2))" by (rule filterlim_at_top_at_left[where Q="\x. - pi/2 < x \ x < pi/2" and P="\x. True" and g=arctan]) (auto simp: arctan le_less eventually_at dist_real_def simp del: less_divide_eq_numeral1 intro!: tan_monotone exI[of _ "pi/2"]) lemma tendsto_arctan_at_top: "(arctan \ (pi/2)) at_top" proof (rule tendstoI) fix e :: real assume "0 < e" define y where "y = pi/2 - min (pi/2) e" then have y: "0 \ y" "y < pi/2" "pi/2 \ e + y" using \0 < e\ by auto show "eventually (\x. dist (arctan x) (pi/2) < e) at_top" proof (intro eventually_at_top_dense[THEN iffD2] exI allI impI) fix x assume "tan y < x" then have "arctan (tan y) < arctan x" by (simp add: arctan_less_iff) with y have "y < arctan x" by (subst (asm) arctan_tan) simp_all with arctan_ubound[of x, arith] y \0 < e\ show "dist (arctan x) (pi/2) < e" by (simp add: dist_real_def) qed qed lemma tendsto_arctan_at_bot: "(arctan \ - (pi/2)) at_bot" unfolding filterlim_at_bot_mirror arctan_minus by (intro tendsto_minus tendsto_arctan_at_top) subsection \Prove Totality of the Trigonometric Functions\ lemma cos_arccos_abs: "\y\ \ 1 \ cos (arccos y) = y" by (simp add: abs_le_iff) lemma sin_arccos_abs: "\y\ \ 1 \ sin (arccos y) = sqrt (1 - y\<^sup>2)" by (simp add: sin_arccos abs_le_iff) lemma sin_mono_less_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x < sin y \ x < y" by (metis not_less_iff_gr_or_eq sin_monotone_2pi) lemma sin_mono_le_eq: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x \ sin y \ x \ y" by (meson leD le_less_linear sin_monotone_2pi sin_monotone_2pi_le) lemma sin_inj_pi: "- (pi/2) \ x \ x \ pi/2 \ - (pi/2) \ y \ y \ pi/2 \ sin x = sin y \ x = y" by (metis arcsin_sin) lemma arcsin_le_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma le_arcsin_iff: assumes "x \ -1" "x \ 1" "y \ -pi/2" "y \ pi/2" shows "arcsin x \ y \ x \ sin y" proof - have "arcsin x \ y \ sin (arcsin x) \ sin y" using arcsin_bounded[of x] assms by (subst sin_mono_le_eq) auto also from assms have "sin (arcsin x) = x" by simp finally show ?thesis . qed lemma cos_mono_less_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x < cos y \ y < x" by (meson cos_monotone_0_pi cos_monotone_0_pi_le leD le_less_linear) lemma cos_mono_le_eq: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x \ cos y \ y \ x" by (metis arccos_cos cos_monotone_0_pi_le eq_iff linear) lemma cos_inj_pi: "0 \ x \ x \ pi \ 0 \ y \ y \ pi \ cos x = cos y \ x = y" by (metis arccos_cos) lemma arccos_le_pi2: "\0 \ y; y \ 1\ \ arccos y \ pi/2" by (metis (mono_tags) arccos_0 arccos cos_le_one cos_monotone_0_pi_le cos_pi cos_pi_half pi_half_ge_zero antisym_conv less_eq_neg_nonpos linear minus_minus order.trans order_refl) lemma sincos_total_pi_half: assumes "0 \ x" "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi/2 \ x = cos t \ y = sin t" proof - have x1: "x \ 1" using assms by (metis le_add_same_cancel1 power2_le_imp_le power_one zero_le_power2) with assms have *: "0 \ arccos x" "cos (arccos x) = x" by (auto simp: arccos) from assms have "y = sqrt (1 - x\<^sup>2)" by (metis abs_of_nonneg add.commute add_diff_cancel real_sqrt_abs) with x1 * assms arccos_le_pi2 [of x] show ?thesis by (rule_tac x="arccos x" in exI) (auto simp: sin_arccos) qed lemma sincos_total_pi: assumes "0 \ y" "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 x]) case le from sincos_total_pi_half [OF le] show ?thesis by (metis pi_ge_two pi_half_le_two add.commute add_le_cancel_left add_mono assms) next case ge then have "0 \ -x" by simp then obtain t where t: "t\0" "t \ pi/2" "-x = cos t" "y = sin t" using sincos_total_pi_half assms by auto (metis \0 \ - x\ power2_minus) show ?thesis by (rule exI [where x = "pi -t"]) (use t in auto) qed lemma sincos_total_2pi_le: assumes "x\<^sup>2 + y\<^sup>2 = 1" shows "\t. 0 \ t \ t \ 2 * pi \ x = cos t \ y = sin t" proof (cases rule: le_cases [of 0 y]) case le from sincos_total_pi [OF le] show ?thesis by (metis assms le_add_same_cancel1 mult.commute mult_2_right order.trans) next case ge then have "0 \ -y" by simp then obtain t where t: "t\0" "t \ pi" "x = cos t" "-y = sin t" using sincos_total_pi assms by auto (metis \0 \ - y\ power2_minus) show ?thesis by (rule exI [where x = "2 * pi - t"]) (use t in auto) qed lemma sincos_total_2pi: assumes "x\<^sup>2 + y\<^sup>2 = 1" obtains t where "0 \ t" "t < 2*pi" "x = cos t" "y = sin t" proof - from sincos_total_2pi_le [OF assms] obtain t where t: "0 \ t" "t \ 2*pi" "x = cos t" "y = sin t" by blast show ?thesis by (cases "t = 2 * pi") (use t that in \force+\) qed lemma arcsin_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x < arcsin y \ x < y" by (rule trans [OF sin_mono_less_eq [symmetric]]) (use arcsin_ubound arcsin_lbound in auto) lemma arcsin_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arcsin x \ arcsin y \ x \ y" using arcsin_less_mono not_le by blast lemma arcsin_less_arcsin: "- 1 \ x \ x < y \ y \ 1 \ arcsin x < arcsin y" using arcsin_less_mono by auto lemma arcsin_le_arcsin: "- 1 \ x \ x \ y \ y \ 1 \ arcsin x \ arcsin y" using arcsin_le_mono by auto lemma arccos_less_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x < arccos y \ y < x" by (rule trans [OF cos_mono_less_eq [symmetric]]) (use arccos_ubound arccos_lbound in auto) lemma arccos_le_mono: "\x\ \ 1 \ \y\ \ 1 \ arccos x \ arccos y \ y \ x" using arccos_less_mono [of y x] by (simp add: not_le [symmetric]) lemma arccos_less_arccos: "- 1 \ x \ x < y \ y \ 1 \ arccos y < arccos x" using arccos_less_mono by auto lemma arccos_le_arccos: "- 1 \ x \ x \ y \ y \ 1 \ arccos y \ arccos x" using arccos_le_mono by auto lemma arccos_eq_iff: "\x\ \ 1 \ \y\ \ 1 \ arccos x = arccos y \ x = y" using cos_arccos_abs by fastforce lemma arccos_cos_eq_abs: assumes "\\\ \ pi" shows "arccos (cos \) = \\\" unfolding arccos_def proof (intro the_equality conjI; clarify?) show "cos \\\ = cos \" by (simp add: abs_real_def) show "x = \\\" if "cos x = cos \" "0 \ x" "x \ pi" for x by (simp add: \cos \\\ = cos \\ assms cos_inj_pi that) qed (use assms in auto) lemma arccos_cos_eq_abs_2pi: obtains k where "arccos (cos \) = \\ - of_int k * (2 * pi)\" proof - define k where "k \ \(\ + pi) / (2 * pi)\" have lepi: "\\ - of_int k * (2 * pi)\ \ pi" using floor_divide_lower [of "2*pi" "\ + pi"] floor_divide_upper [of "2*pi" "\ + pi"] by (auto simp: k_def abs_if algebra_simps) have "arccos (cos \) = arccos (cos (\ - of_int k * (2 * pi)))" using cos_int_2pin sin_int_2pin by (simp add: cos_diff mult.commute) also have "\ = \\ - of_int k * (2 * pi)\" using arccos_cos_eq_abs lepi by blast finally show ?thesis using that by metis qed 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 arcsin_arccos_sqrt_pos: "0 \ x \ x \ 1 \ arcsin x = arccos(sqrt(1 - x\<^sup>2))" by (smt (verit, del_insts) arccos_cos arcsin_0 arcsin_le_arcsin arcsin_pi cos_arcsin) 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))" by (smt (verit, del_insts) arccos_lbound arccos_le_pi2 arcsin_sin sin_arccos) 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) lemma cos_limit_1: assumes "(\j. cos (\ j)) \ 1" shows "\k. (\j. \ j - of_int (k j) * (2 * pi)) \ 0" proof - have "\\<^sub>F j in sequentially. cos (\ j) \ {- 1..1}" by auto then have "(\j. arccos (cos (\ j))) \ arccos 1" using continuous_on_tendsto_compose [OF continuous_on_arccos' assms] by auto moreover have "\j. \k. arccos (cos (\ j)) = \\ j - of_int k * (2 * pi)\" using arccos_cos_eq_abs_2pi by metis then have "\k. \j. arccos (cos (\ j)) = \\ j - of_int (k j) * (2 * pi)\" by metis ultimately have "\k. (\j. \\ j - of_int (k j) * (2 * pi)\) \ 0" by auto then show ?thesis by (simp add: tendsto_rabs_zero_iff) qed lemma cos_diff_limit_1: assumes "(\j. cos (\ j - \)) \ 1" obtains k where "(\j. \ j - of_int (k j) * (2 * pi)) \ \" proof - obtain k where "(\j. (\ j - \) - of_int (k j) * (2 * pi)) \ 0" using cos_limit_1 [OF assms] by auto then have "(\j. \ + ((\ j - \) - of_int (k j) * (2 * pi))) \ \ + 0" by (rule tendsto_add [OF tendsto_const]) with that show ?thesis by auto qed subsection \Machin's formula\ lemma arctan_one: "arctan 1 = pi / 4" by (rule arctan_unique) (simp_all add: tan_45 m2pi_less_pi) lemma tan_total_pi4: assumes "\x\ < 1" shows "\z. - (pi / 4) < z \ z < pi / 4 \ tan z = x" proof show "- (pi / 4) < arctan x \ arctan x < pi / 4 \ tan (arctan x) = x" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_less_iff using assms by (auto simp: arctan) qed lemma arctan_add: assumes "\x\ \ 1" "\y\ < 1" shows "arctan x + arctan y = arctan ((x + y) / (1 - x * y))" proof (rule arctan_unique [symmetric]) have "- (pi / 4) \ arctan x" "- (pi / 4) < arctan y" unfolding arctan_one [symmetric] arctan_minus [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 1: "- (pi/2) < arctan x + arctan y" by simp have "arctan x \ pi / 4" "arctan y < pi / 4" unfolding arctan_one [symmetric] unfolding arctan_le_iff arctan_less_iff using assms by auto from add_le_less_mono [OF this] show 2: "arctan x + arctan y < pi/2" by simp show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)" using cos_gt_zero_pi [OF 1 2] by (simp add: arctan tan_add) qed lemma arctan_double: "\x\ < 1 \ 2 * arctan x = arctan ((2 * x) / (1 - x\<^sup>2))" by (metis arctan_add linear mult_2 not_less power2_eq_square) theorem machin: "pi / 4 = 4 * arctan (1 / 5) - arctan (1 / 239)" proof - have "\1 / 5\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (1 / 5) = arctan (5 / 12)" by auto moreover have "\5 / 12\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF this] this] have "2 * arctan (5 / 12) = arctan (120 / 119)" by auto moreover have "\1\ \ (1::real)" and "\1 / 239\ < (1::real)" by auto from arctan_add[OF this] have "arctan 1 + arctan (1 / 239) = arctan (120 / 119)" by auto ultimately have "arctan 1 + arctan (1 / 239) = 4 * arctan (1 / 5)" by auto then show ?thesis unfolding arctan_one by algebra qed lemma machin_Euler: "5 * arctan (1 / 7) + 2 * arctan (3 / 79) = pi / 4" proof - have 17: "\1 / 7\ < (1 :: real)" by auto with arctan_double have "2 * arctan (1 / 7) = arctan (7 / 24)" by simp (simp add: field_simps) moreover have "\7 / 24\ < (1 :: real)" by auto with arctan_double have "2 * arctan (7 / 24) = arctan (336 / 527)" by simp (simp add: field_simps) moreover have "\336 / 527\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF 17] this] have "arctan(1/7) + arctan (336 / 527) = arctan (2879 / 3353)" by auto ultimately have I: "5 * arctan (1 / 7) = arctan (2879 / 3353)" by auto have 379: "\3 / 79\ < (1 :: real)" by auto with arctan_double have II: "2 * arctan (3 / 79) = arctan (237 / 3116)" by simp (simp add: field_simps) have *: "\2879 / 3353\ < (1 :: real)" by auto have "\237 / 3116\ < (1 :: real)" by auto from arctan_add[OF less_imp_le[OF *] this] have "arctan (2879/3353) + arctan (237/3116) = pi/4" by (simp add: arctan_one) with I II show ?thesis by auto qed (*But could also prove MACHIN_GAUSS: 12 * arctan(1/18) + 8 * arctan(1/57) - 5 * arctan(1/239) = pi/4*) subsection \Introducing the inverse tangent power series\ lemma monoseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "monoseq (\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1))" (is "monoseq ?a") proof (cases "x = 0") case True then show ?thesis by (auto simp: monoseq_def) next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "monoseq ?a" proof - have mono: "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" if "0 \ x" and "x \ 1" for n and x :: real proof (rule mult_mono) show "1 / real (Suc (Suc n * 2)) \ 1 / real (Suc (n * 2))" by (rule frac_le) simp_all show "0 \ 1 / real (Suc (n * 2))" by auto show "x ^ Suc (Suc n * 2) \ x ^ Suc (n * 2)" by (rule power_decreasing) (simp_all add: \0 \ x\ \x \ 1\) show "0 \ x ^ Suc (Suc n * 2)" by (rule zero_le_power) (simp add: \0 \ x\) qed show ?thesis proof (cases "0 \ x") case True from mono[OF this \x \ 1\, THEN allI] show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI2) next case False then have "0 \ - x" and "- x \ 1" using \-1 \ x\ by auto from mono[OF this] have "1 / real (Suc (Suc n * 2)) * x ^ Suc (Suc n * 2) \ 1 / real (Suc (n * 2)) * x ^ Suc (n * 2)" for n using \0 \ -x\ by auto then show ?thesis unfolding Suc_eq_plus1[symmetric] by (rule mono_SucI1[OF allI]) qed qed qed lemma zeroseq_arctan_series: fixes x :: real assumes "\x\ \ 1" shows "(\n. 1 / real (n * 2 + 1) * x^(n * 2 + 1)) \ 0" (is "?a \ 0") proof (cases "x = 0") case True then show ?thesis by simp next case False have "norm x \ 1" and "x \ 1" and "-1 \ x" using assms by auto show "?a \ 0" proof (cases "\x\ < 1") case True then have "norm x < 1" by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat LIMSEQ_power_zero[OF \norm x < 1\, THEN LIMSEQ_Suc]] have "(\n. 1 / real (n + 1) * x ^ (n + 1)) \ 0" unfolding inverse_eq_divide Suc_eq_plus1 by simp then show ?thesis using pos2 by (rule LIMSEQ_linear) next case False then have "x = -1 \ x = 1" using \\x\ \ 1\ by auto then have n_eq: "\ n. x ^ (n * 2 + 1) = x" unfolding One_nat_def by auto from tendsto_mult[OF LIMSEQ_inverse_real_of_nat[THEN LIMSEQ_linear, OF pos2, unfolded inverse_eq_divide] tendsto_const[of x]] show ?thesis unfolding n_eq Suc_eq_plus1 by auto qed qed lemma summable_arctan_series: fixes n :: nat assumes "\x\ \ 1" shows "summable (\ k. (-1)^k * (1 / real (k*2+1) * x ^ (k*2+1)))" (is "summable (?c x)") by (rule summable_Leibniz(1), rule zeroseq_arctan_series[OF assms], rule monoseq_arctan_series[OF assms]) lemma DERIV_arctan_series: assumes "\x\ < 1" shows "DERIV (\x'. \k. (-1)^k * (1 / real (k * 2 + 1) * x' ^ (k * 2 + 1))) x :> (\k. (-1)^k * x^(k * 2))" (is "DERIV ?arctan _ :> ?Int") proof - let ?f = "\n. if even n then (-1)^(n div 2) * 1 / real (Suc n) else 0" have n_even: "even n \ 2 * (n div 2) = n" for n :: nat by presburger then have if_eq: "?f n * real (Suc n) * x'^n = (if even n then (-1)^(n div 2) * x'^(2 * (n div 2)) else 0)" for n x' by auto have summable_Integral: "summable (\ n. (- 1) ^ n * x^(2 * n))" if "\x\ < 1" for x :: real proof - from that have "x\<^sup>2 < 1" by (simp add: abs_square_less_1) have "summable (\ n. (- 1) ^ n * (x\<^sup>2) ^n)" by (rule summable_Leibniz(1)) (auto intro!: LIMSEQ_realpow_zero monoseq_realpow \x\<^sup>2 < 1\ order_less_imp_le[OF \x\<^sup>2 < 1\]) then show ?thesis by (simp only: power_mult) qed have sums_even: "(sums) f = (sums) (\ n. if even n then f (n div 2) else 0)" for f :: "nat \ real" proof - have "f sums x = (\ n. if even n then f (n div 2) else 0) sums x" for x :: real proof assume "f sums x" from sums_if[OF sums_zero this] show "(\n. if even n then f (n div 2) else 0) sums x" by auto next assume "(\ n. if even n then f (n div 2) else 0) sums x" from LIMSEQ_linear[OF this[simplified sums_def] pos2, simplified sum_split_even_odd[simplified mult.commute]] show "f sums x" unfolding sums_def by auto qed then show ?thesis .. qed have Int_eq: "(\n. ?f n * real (Suc n) * x^n) = ?Int" unfolding if_eq mult.commute[of _ 2] suminf_def sums_even[of "\ n. (- 1) ^ n * x ^ (2 * n)", symmetric] by auto have arctan_eq: "(\n. ?f n * x^(Suc n)) = ?arctan x" for x proof - have if_eq': "\n. (if even n then (- 1) ^ (n div 2) * 1 / real (Suc n) else 0) * x ^ Suc n = (if even n then (- 1) ^ (n div 2) * (1 / real (Suc (2 * (n div 2))) * x ^ Suc (2 * (n div 2))) else 0)" using n_even by auto have idx_eq: "\n. n * 2 + 1 = Suc (2 * n)" by auto then show ?thesis unfolding if_eq' idx_eq suminf_def sums_even[of "\ n. (- 1) ^ n * (1 / real (Suc (2 * n)) * x ^ Suc (2 * n))", symmetric] by auto qed have "DERIV (\ x. \ n. ?f n * x^(Suc n)) x :> (\n. ?f n * real (Suc n) * x^n)" proof (rule DERIV_power_series') show "x \ {- 1 <..< 1}" using \\ x \ < 1\ by auto show "summable (\ n. ?f n * real (Suc n) * x'^n)" if x'_bounds: "x' \ {- 1 <..< 1}" for x' :: real proof - from that have "\x'\ < 1" by auto then show ?thesis using that sums_summable sums_if [OF sums_0 [of "\x. 0"] summable_sums [OF summable_Integral]] by (auto simp add: if_distrib [of "\x. x * y" for y] cong: if_cong) qed qed auto then show ?thesis by (simp only: Int_eq arctan_eq) qed lemma arctan_series: assumes "\x\ \ 1" shows "arctan x = (\k. (-1)^k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))" (is "_ = suminf (\ n. ?c x n)") proof - let ?c' = "\x n. (-1)^n * x^(n*2)" have DERIV_arctan_suminf: "DERIV (\ x. suminf (?c x)) x :> (suminf (?c' x))" if "0 < r" and "r < 1" and "\x\ < r" for r x :: real proof (rule DERIV_arctan_series) from that show "\x\ < 1" using \r < 1\ and \\x\ < r\ by auto qed { fix x :: real assume "\x\ \ 1" note summable_Leibniz[OF zeroseq_arctan_series[OF this] monoseq_arctan_series[OF this]] } note arctan_series_borders = this have when_less_one: "arctan x = (\k. ?c x k)" if "\x\ < 1" for x :: real proof - obtain r where "\x\ < r" and "r < 1" using dense[OF \\x\ < 1\] by blast then have "0 < r" and "- r < x" and "x < r" by auto have suminf_eq_arctan_bounded: "suminf (?c x) - arctan x = suminf (?c a) - arctan a" if "-r < a" and "b < r" and "a < b" and "a \ x" and "x \ b" for x a b proof - from that have "\x\ < r" by auto show "suminf (?c x) - arctan x = suminf (?c a) - arctan a" proof (rule DERIV_isconst2[of "a" "b"]) show "a < b" and "a \ x" and "x \ b" using \a < b\ \a \ x\ \x \ b\ by auto have "\x. - r < x \ x < r \ DERIV (\ x. suminf (?c x) - arctan x) x :> 0" proof (rule allI, rule impI) fix x assume "-r < x \ x < r" then have "\x\ < r" by auto with \r < 1\ have "\x\ < 1" by auto have "\- (x\<^sup>2)\ < 1" using abs_square_less_1 \\x\ < 1\ by auto then have "(\n. (- (x\<^sup>2)) ^ n) sums (1 / (1 - (- (x\<^sup>2))))" unfolding real_norm_def[symmetric] by (rule geometric_sums) then have "(?c' x) sums (1 / (1 - (- (x\<^sup>2))))" unfolding power_mult_distrib[symmetric] power_mult mult.commute[of _ 2] by auto then have suminf_c'_eq_geom: "inverse (1 + x\<^sup>2) = suminf (?c' x)" using sums_unique unfolding inverse_eq_divide by auto have "DERIV (\ x. suminf (?c x)) x :> (inverse (1 + x\<^sup>2))" unfolding suminf_c'_eq_geom by (rule DERIV_arctan_suminf[OF \0 < r\ \r < 1\ \\x\ < r\]) from DERIV_diff [OF this DERIV_arctan] show "DERIV (\x. suminf (?c x) - arctan x) x :> 0" by auto qed then have DERIV_in_rball: "\y. a \ y \ y \ b \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \-r < a\ \b < r\ by auto then show "\y. \a < y; y < b\ \ DERIV (\x. suminf (?c x) - arctan x) y :> 0" using \\x\ < r\ by auto show "continuous_on {a..b} (\x. suminf (?c x) - arctan x)" using DERIV_in_rball DERIV_atLeastAtMost_imp_continuous_on by blast qed qed have suminf_arctan_zero: "suminf (?c 0) - arctan 0 = 0" unfolding Suc_eq_plus1[symmetric] power_Suc2 mult_zero_right arctan_zero_zero suminf_zero by auto have "suminf (?c x) - arctan x = 0" proof (cases "x = 0") case True then show ?thesis using suminf_arctan_zero by auto next case False then have "0 < \x\" and "- \x\ < \x\" by auto have "suminf (?c (- \x\)) - arctan (- \x\) = suminf (?c 0) - arctan 0" by (rule suminf_eq_arctan_bounded[where x1=0 and a1="-\x\" and b1="\x\", symmetric]) (simp_all only: \\x\ < r\ \-\x\ < \x\\ neg_less_iff_less) moreover have "suminf (?c x) - arctan x = suminf (?c (- \x\)) - arctan (- \x\)" by (rule suminf_eq_arctan_bounded[where x1=x and a1="- \x\" and b1="\x\"]) (simp_all only: \\x\ < r\ \- \x\ < \x\\ neg_less_iff_less) ultimately show ?thesis using suminf_arctan_zero by auto qed then show ?thesis by auto qed show "arctan x = suminf (\n. ?c x n)" proof (cases "\x\ < 1") case True then show ?thesis by (rule when_less_one) next case False then have "\x\ = 1" using \\x\ \ 1\ by auto let ?a = "\x n. \1 / real (n * 2 + 1) * x^(n * 2 + 1)\" let ?diff = "\x n. \arctan x - (\i" have "?diff 1 n \ ?a 1 n" for n :: nat proof - have "0 < (1 :: real)" by auto moreover have "?diff x n \ ?a x n" if "0 < x" and "x < 1" for x :: real proof - from that have "\x\ \ 1" and "\x\ < 1" by auto from \0 < x\ have "0 < 1 / real (0 * 2 + (1::nat)) * x ^ (0 * 2 + 1)" by auto note bounds = mp[OF arctan_series_borders(2)[OF \\x\ \ 1\] this, unfolded when_less_one[OF \\x\ < 1\, symmetric], THEN spec] have "0 < 1 / real (n*2+1) * x^(n*2+1)" by (rule mult_pos_pos) (simp_all only: zero_less_power[OF \0 < x\], auto) then have a_pos: "?a x n = 1 / real (n*2+1) * x^(n*2+1)" by (rule abs_of_pos) show ?thesis proof (cases "even n") case True then have sgn_pos: "(-1)^n = (1::real)" by auto from \even n\ obtain m where "n = 2 * m" .. then have "2 * m = n" .. from bounds[of m, unfolded this atLeastAtMost_iff] have "\arctan x - (\i \ (\ii = ?c x n" by auto also have "\ = ?a x n" unfolding sgn_pos a_pos by auto finally show ?thesis . next case False then have sgn_neg: "(-1)^n = (-1::real)" by auto from \odd n\ obtain m where "n = 2 * m + 1" .. then have m_def: "2 * m + 1 = n" .. then have m_plus: "2 * (m + 1) = n + 1" by auto from bounds[of "m + 1", unfolded this atLeastAtMost_iff, THEN conjunct1] bounds[of m, unfolded m_def atLeastAtMost_iff, THEN conjunct2] have "\arctan x - (\i \ (\ii = - ?c x n" by auto also have "\ = ?a x n" unfolding sgn_neg a_pos by auto finally show ?thesis . qed qed hence "\x \ { 0 <..< 1 }. 0 \ ?a x n - ?diff x n" by auto moreover have "isCont (\ x. ?a x n - ?diff x n) x" for x unfolding diff_conv_add_uminus divide_inverse by (auto intro!: isCont_add isCont_rabs continuous_ident isCont_minus isCont_arctan continuous_at_within_inverse isCont_mult isCont_power continuous_const isCont_sum simp del: add_uminus_conv_diff) ultimately have "0 \ ?a 1 n - ?diff 1 n" by (rule LIM_less_bound) then show ?thesis by auto qed have "?a 1 \ 0" unfolding tendsto_rabs_zero_iff power_one divide_inverse One_nat_def by (auto intro!: tendsto_mult LIMSEQ_linear LIMSEQ_inverse_real_of_nat simp del: of_nat_Suc) have "?diff 1 \ 0" proof (rule LIMSEQ_I) fix r :: real assume "0 < r" obtain N :: nat where N_I: "N \ n \ ?a 1 n < r" for n using LIMSEQ_D[OF \?a 1 \ 0\ \0 < r\] by auto have "norm (?diff 1 n - 0) < r" if "N \ n" for n using \?diff 1 n \ ?a 1 n\ N_I[OF that] by auto then show "\N. \ n \ N. norm (?diff 1 n - 0) < r" by blast qed from this [unfolded tendsto_rabs_zero_iff, THEN tendsto_add [OF _ tendsto_const], of "- arctan 1", THEN tendsto_minus] have "(?c 1) sums (arctan 1)" unfolding sums_def by auto then have "arctan 1 = (\i. ?c 1 i)" by (rule sums_unique) show ?thesis proof (cases "x = 1") case True then show ?thesis by (simp add: \arctan 1 = (\ i. ?c 1 i)\) next case False then have "x = -1" using \\x\ = 1\ by auto have "- (pi/2) < 0" using pi_gt_zero by auto have "- (2 * pi) < 0" using pi_gt_zero by auto have c_minus_minus: "?c (- 1) i = - ?c 1 i" for i by auto have "arctan (- 1) = arctan (tan (-(pi / 4)))" unfolding tan_45 tan_minus .. also have "\ = - (pi / 4)" by (rule arctan_tan) (auto simp: order_less_trans[OF \- (pi/2) < 0\ pi_gt_zero]) also have "\ = - (arctan (tan (pi / 4)))" unfolding neg_equal_iff_equal by (rule arctan_tan[symmetric]) (auto simp: order_less_trans[OF \- (2 * pi) < 0\ pi_gt_zero]) also have "\ = - (arctan 1)" unfolding tan_45 .. also have "\ = - (\ i. ?c 1 i)" using \arctan 1 = (\ i. ?c 1 i)\ by auto also have "\ = (\ i. ?c (- 1) i)" using suminf_minus[OF sums_summable[OF \(?c 1) sums (arctan 1)\]] unfolding c_minus_minus by auto finally show ?thesis using \x = -1\ by auto qed qed qed lemma arctan_half: "arctan x = 2 * arctan (x / (1 + sqrt(1 + x\<^sup>2)))" for x :: real proof - obtain y where low: "- (pi/2) < y" and high: "y < pi/2" and y_eq: "tan y = x" using tan_total by blast then have low2: "- (pi/2) < y / 2" and high2: "y / 2 < pi/2" by auto have "0 < cos y" by (rule cos_gt_zero_pi[OF low high]) then have "cos y \ 0" and cos_sqrt: "sqrt ((cos y)\<^sup>2) = cos y" by auto have "1 + (tan y)\<^sup>2 = 1 + (sin y)\<^sup>2 / (cos y)\<^sup>2" unfolding tan_def power_divide .. also have "\ = (cos y)\<^sup>2 / (cos y)\<^sup>2 + (sin y)\<^sup>2 / (cos y)\<^sup>2" using \cos y \ 0\ by auto also have "\ = 1 / (cos y)\<^sup>2" unfolding add_divide_distrib[symmetric] sin_cos_squared_add2 .. finally have "1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2" . have "sin y / (cos y + 1) = tan y / ((cos y + 1) / cos y)" unfolding tan_def using \cos y \ 0\ by (simp add: field_simps) also have "\ = tan y / (1 + 1 / cos y)" using \cos y \ 0\ unfolding add_divide_distrib by auto also have "\ = tan y / (1 + 1 / sqrt ((cos y)\<^sup>2))" unfolding cos_sqrt .. also have "\ = tan y / (1 + sqrt (1 / (cos y)\<^sup>2))" unfolding real_sqrt_divide by auto finally have eq: "sin y / (cos y + 1) = tan y / (1 + sqrt(1 + (tan y)\<^sup>2))" unfolding \1 + (tan y)\<^sup>2 = 1 / (cos y)\<^sup>2\ . have "arctan x = y" using arctan_tan low high y_eq by auto also have "\ = 2 * (arctan (tan (y/2)))" using arctan_tan[OF low2 high2] by auto also have "\ = 2 * (arctan (sin y / (cos y + 1)))" unfolding tan_half by auto finally show ?thesis unfolding eq \tan y = x\ . qed lemma arctan_monotone: "x < y \ arctan x < arctan y" by (simp only: arctan_less_iff) lemma arctan_monotone': "x \ y \ arctan x \ arctan y" by (simp only: arctan_le_iff) lemma arctan_inverse: assumes "x \ 0" shows "arctan (1 / x) = sgn x * pi/2 - arctan x" proof (rule arctan_unique) have \
: "x > 0 \ arctan x < pi" using arctan_bounded [of x] by linarith show "- (pi/2) < sgn x * pi/2 - arctan x" using assms by (auto simp: sgn_real_def arctan algebra_simps \
) show "sgn x * pi/2 - arctan x < pi/2" using arctan_bounded [of "- x"] assms by (auto simp: algebra_simps sgn_real_def arctan_minus) show "tan (sgn x * pi/2 - arctan x) = 1 / x" unfolding tan_inverse [of "arctan x", unfolded tan_arctan] sgn_real_def by (simp add: tan_def cos_arctan sin_arctan sin_diff cos_diff) qed theorem pi_series: "pi / 4 = (\k. (-1)^k * 1 / real (k * 2 + 1))" (is "_ = ?SUM") proof - have "pi / 4 = arctan 1" using arctan_one by auto also have "\ = ?SUM" using arctan_series[of 1] by auto finally show ?thesis by auto qed subsection \Existence of Polar Coordinates\ lemma cos_x_y_le_one: "\x / sqrt (x\<^sup>2 + y\<^sup>2)\ \ 1" by (rule power2_le_imp_le [OF _ zero_le_one]) (simp add: power_divide divide_le_eq not_sum_power2_lt_zero) lemma polar_Ex: "\r::real. \a. x = r * cos a \ y = r * sin a" proof - have polar_ex1: "\r a. x = r * cos a \ y = r * sin a" if "0 < y" for y proof - have "x = sqrt (x\<^sup>2 + y\<^sup>2) * cos (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" by (simp add: cos_arccos_abs [OF cos_x_y_le_one]) moreover have "y = sqrt (x\<^sup>2 + y\<^sup>2) * sin (arccos (x / sqrt (x\<^sup>2 + y\<^sup>2)))" using that by (simp add: sin_arccos_abs [OF cos_x_y_le_one] power_divide right_diff_distrib flip: real_sqrt_mult) ultimately show ?thesis by blast qed show ?thesis proof (cases "0::real" y rule: linorder_cases) case less then show ?thesis by (rule polar_ex1) next case equal then show ?thesis by (force simp: intro!: cos_zero sin_zero) next case greater with polar_ex1 [where y="-y"] show ?thesis by auto (metis cos_minus minus_minus minus_mult_right sin_minus) qed qed subsection \Basics about polynomial functions: products, extremal behaviour and root counts\ lemma pairs_le_eq_Sigma: "{(i, j). i + j \ m} = Sigma (atMost m) (\r. atMost (m - r))" for m :: nat by auto lemma sum_up_index_split: "(\k\m + n. f k) = (\k\m. f k) + (\k = Suc m..m + n. f k)" by (metis atLeast0AtMost Suc_eq_plus1 le0 sum.ub_add_nat) lemma Sigma_interval_disjoint: "(SIGMA i:A. {..v i}) \ (SIGMA i:A.{v i<..w}) = {}" for w :: "'a::order" by auto lemma product_atMost_eq_Un: "A \ {..m} = (SIGMA i:A.{..m - i}) \ (SIGMA i:A.{m - i<..m})" for m :: nat by auto lemma polynomial_product: (*with thanks to Chaitanya Mangla*) fixes x :: "'a::idom" assumes m: "\i. i > m \ a i = 0" and n: "\j. j > n \ b j = 0" shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" proof - have "\i j. \m + n - i < j; a i \ 0\ \ b j = 0" by (meson le_add_diff leI le_less_trans m n) then have \
: "(\(i,j)\(SIGMA i:{..m+n}. {m+n - i<..m+n}). a i * x ^ i * (b j * x ^ j)) = 0" by (clarsimp simp add: sum_Un Sigma_interval_disjoint intro!: sum.neutral) have "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\i\m. \j\n. (a i * x ^ i) * (b j * x ^ j))" by (rule sum_product) also have "\ = (\i\m + n. \j\n + m. a i * x ^ i * (b j * x ^ j))" using assms by (auto simp: sum_up_index_split) also have "\ = (\r\m + n. \j\m + n - r. a r * x ^ r * (b j * x ^ j))" by (simp add: add_ac sum.Sigma product_atMost_eq_Un sum_Un Sigma_interval_disjoint \
) also have "\ = (\(i,j)\{(i,j). i+j \ m+n}. (a i * x ^ i) * (b j * x ^ j))" by (auto simp: pairs_le_eq_Sigma sum.Sigma) also have "... = (\k\m + n. \i\k. a i * x ^ i * (b (k - i) * x ^ (k - i)))" by (rule sum.triangle_reindex_eq) also have "\ = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" by (auto simp: algebra_simps sum_distrib_left simp flip: power_add intro!: sum.cong) finally show ?thesis . qed lemma polynomial_product_nat: fixes x :: nat assumes m: "\i. i > m \ a i = 0" and n: "\j. j > n \ b j = 0" shows "(\i\m. (a i) * x ^ i) * (\j\n. (b j) * x ^ j) = (\r\m + n. (\k\r. (a k) * (b (r - k))) * x ^ r)" using polynomial_product [of m a n b x] assms by (simp only: of_nat_mult [symmetric] of_nat_power [symmetric] of_nat_eq_iff Int.int_sum [symmetric]) lemma polyfun_diff: (*COMPLEX_SUB_POLYFUN in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" proof - have h: "bij_betw (\(i,j). (j,i)) ((SIGMA i : atMost n. lessThan i)) (SIGMA j : lessThan n. {Suc j..n})" by (auto simp: bij_betw_def inj_on_def) have "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (\i\n. a i * (x^i - y^i))" by (simp add: right_diff_distrib sum_subtractf) also have "\ = (\i\n. a i * (x - y) * (\j = (\i\n. \j = (\(i,j) \ (SIGMA i : atMost n. lessThan i). a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (\(j,i) \ (SIGMA j : lessThan n. {Suc j..n}). a i * (x - y) * (y^(i - Suc j) * x^j))" by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) also have "\ = (\ji=Suc j..n. a i * (x - y) * (y^(i - Suc j) * x^j))" by (simp add: sum.Sigma) also have "\ = (x - y) * (\ji=Suc j..n. a i * y^(i - j - 1)) * x^j)" by (simp add: sum_distrib_left mult_ac) finally show ?thesis . qed lemma polyfun_diff_alt: (*COMPLEX_SUB_POLYFUN_ALT in HOL Light*) fixes x :: "'a::idom" assumes "1 \ n" shows "(\i\n. a i * x^i) - (\i\n. a i * y^i) = (x - y) * ((\jki=Suc j..n. a i * y^(i - j - 1)) = (\kk. k < n - j \ k \ (\i. i - Suc j) ` {Suc j..n}" by (rule_tac x="k + Suc j" in image_eqI, auto) then have h: "bij_betw (\i. i - (j + 1)) {Suc j..n} (lessThan (n-j))" by (auto simp: bij_betw_def inj_on_def) then show ?thesis by (auto simp: sum.reindex_bij_betw [OF h, symmetric] intro: sum.cong_simp) qed then show ?thesis by (simp add: polyfun_diff [OF assms] sum_distrib_right) qed lemma polyfun_linear_factor: (*COMPLEX_POLYFUN_LINEAR_FACTOR in HOL Light*) fixes a :: "'a::idom" shows "\b. \z. (\i\n. c(i) * z^i) = (z - a) * (\ii\n. c(i) * a^i)" proof (cases "n = 0") case True then show ?thesis by simp next case False have "(\b. \z. (\i\n. c i * z^i) = (z - a) * (\ii\n. c i * a^i)) \ (\b. \z. (\i\n. c i * z^i) - (\i\n. c i * a^i) = (z - a) * (\i \ (\b. \z. (z - a) * (\ji = Suc j..n. c i * a^(i - Suc j)) * z^j) = (z - a) * (\i = True" by auto finally show ?thesis by simp qed lemma polyfun_linear_factor_root: (*COMPLEX_POLYFUN_LINEAR_FACTOR_ROOT in HOL Light*) fixes a :: "'a::idom" assumes "(\i\n. c(i) * a^i) = 0" obtains b where "\z. (\i\n. c i * z^i) = (z - a) * (\iw. \i\n. c i * w^i) a" for c :: "nat \ 'a::real_normed_div_algebra" by simp lemma zero_polynom_imp_zero_coeffs: fixes c :: "nat \ 'a::{ab_semigroup_mult,real_normed_div_algebra}" assumes "\w. (\i\n. c i * w^i) = 0" "k \ n" shows "c k = 0" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc n c k) have [simp]: "c 0 = 0" using Suc.prems(1) [of 0] by simp have "(\i\Suc n. c i * w^i) = w * (\i\n. c (Suc i) * w^i)" for w proof - have "(\i\Suc n. c i * w^i) = (\i\n. c (Suc i) * w ^ Suc i)" unfolding Set_Interval.sum.atMost_Suc_shift by simp also have "\ = w * (\i\n. c (Suc i) * w^i)" by (simp add: sum_distrib_left ac_simps) finally show ?thesis . qed then have w: "\w. w \ 0 \ (\i\n. c (Suc i) * w^i) = 0" using Suc by auto then have "(\h. \i\n. c (Suc i) * h^i) \0\ 0" by (simp cong: LIM_cong) \ \the case \w = 0\ by continuity\ then have "(\i\n. c (Suc i) * 0^i) = 0" using isCont_polynom [of 0 "\i. c (Suc i)" n] LIM_unique by (force simp: Limits.isCont_iff) then have "\w. (\i\n. c (Suc i) * w^i) = 0" using w by metis then have "\i. i \ n \ c (Suc i) = 0" using Suc.IH [of "\i. c (Suc i)"] by blast then show ?case using \k \ Suc n\ by (cases k) auto qed lemma polyfun_rootbound: (*COMPLEX_POLYFUN_ROOTBOUND in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows "finite {z. (\i\n. c(i) * z^i) = 0} \ card {z. (\i\n. c(i) * z^i) = 0} \ n" using assms proof (induction n arbitrary: c k) case 0 then show ?case by simp next case (Suc m c k) let ?succase = ?case show ?case proof (cases "{z. (\i\Suc m. c(i) * z^i) = 0} = {}") case True then show ?succase by simp next case False then obtain z0 where z0: "(\i\Suc m. c(i) * z0^i) = 0" by blast then obtain b where b: "\w. (\i\Suc m. c i * w^i) = (w - z0) * (\i\m. b i * w^i)" using polyfun_linear_factor_root [OF z0, unfolded lessThan_Suc_atMost] by blast then have eq: "{z. (\i\Suc m. c i * z^i) = 0} = insert z0 {z. (\i\m. b i * z^i) = 0}" by auto have "\ (\k\m. b k = 0)" proof assume [simp]: "\k\m. b k = 0" then have "\w. (\i\m. b i * w^i) = 0" by simp then have "\w. (\i\Suc m. c i * w^i) = 0" using b by simp then have "\k. k \ Suc m \ c k = 0" using zero_polynom_imp_zero_coeffs by blast then show False using Suc.prems by blast qed then obtain k' where bk': "b k' \ 0" "k' \ m" by blast show ?succase using Suc.IH [of b k'] bk' by (simp add: eq card_insert_if del: sum.atMost_Suc) qed qed lemma fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" assumes "c k \ 0" "k\n" shows polyfun_roots_finite: "finite {z. (\i\n. c(i) * z^i) = 0}" and polyfun_roots_card: "card {z. (\i\n. c(i) * z^i) = 0} \ n" using polyfun_rootbound assms by auto lemma polyfun_finite_roots: (*COMPLEX_POLYFUN_FINITE_ROOTS in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "finite {x. (\i\n. c i * x^i) = 0} \ (\i\n. c i \ 0)" (is "?lhs = ?rhs") proof assume ?lhs moreover have "\ finite {x. (\i\n. c i * x^i) = 0}" if "\i\n. c i = 0" proof - from that have "\x. (\i\n. c i * x^i) = 0" by simp then show ?thesis using ex_new_if_finite [OF infinite_UNIV_char_0 [where 'a='a]] by auto qed ultimately show ?rhs by metis next assume ?rhs with polyfun_rootbound show ?lhs by blast qed lemma polyfun_eq_0: "(\x. (\i\n. c i * x^i) = 0) \ (\i\n. c i = 0)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" (*COMPLEX_POLYFUN_EQ_0 in HOL Light*) using zero_polynom_imp_zero_coeffs by auto lemma polyfun_eq_coeffs: "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\i\n. c i = d i)" for c :: "nat \ 'a::{idom,real_normed_div_algebra}" proof - have "(\x. (\i\n. c i * x^i) = (\i\n. d i * x^i)) \ (\x. (\i\n. (c i - d i) * x^i) = 0)" by (simp add: left_diff_distrib Groups_Big.sum_subtractf) also have "\ \ (\i\n. c i - d i = 0)" by (rule polyfun_eq_0) finally show ?thesis by simp qed lemma polyfun_eq_const: (*COMPLEX_POLYFUN_EQ_CONST in HOL Light*) fixes c :: "nat \ 'a::{idom,real_normed_div_algebra}" shows "(\x. (\i\n. c i * x^i) = k) \ c 0 = k \ (\i \ {1..n}. c i = 0)" (is "?lhs = ?rhs") proof - have *: "\x. (\i\n. (if i=0 then k else 0) * x^i) = k" by (induct n) auto show ?thesis proof assume ?lhs with * have "(\i\n. c i = (if i=0 then k else 0))" by (simp add: polyfun_eq_coeffs [symmetric]) then show ?rhs by simp next assume ?rhs then show ?lhs by (induct n) auto qed qed lemma root_polyfun: fixes z :: "'a::idom" assumes "1 \ n" shows "z^n = a \ (\i\n. (if i = 0 then -a else if i=n then 1 else 0) * z^i) = 0" using assms by (cases n) (simp_all add: sum.atLeast_Suc_atMost atLeast0AtMost [symmetric]) lemma assumes "SORT_CONSTRAINT('a::{idom,real_normed_div_algebra})" and "1 \ n" shows finite_roots_unity: "finite {z::'a. z^n = 1}" and card_roots_unity: "card {z::'a. z^n = 1} \ n" using polyfun_rootbound [of "\i. if i = 0 then -1 else if i=n then 1 else 0" n n] assms(2) by (auto simp: root_polyfun [OF assms(2)]) subsection \Hyperbolic functions\ definition sinh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "sinh x = (exp x - exp (-x)) /\<^sub>R 2" definition cosh :: "'a :: {banach, real_normed_algebra_1} \ 'a" where "cosh x = (exp x + exp (-x)) /\<^sub>R 2" definition tanh :: "'a :: {banach, real_normed_field} \ 'a" where "tanh x = sinh x / cosh x" definition arsinh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arsinh x = ln (x + (x^2 + 1) powr of_real (1/2))" definition arcosh :: "'a :: {banach, real_normed_algebra_1, ln} \ 'a" where "arcosh x = ln (x + (x^2 - 1) powr of_real (1/2))" definition artanh :: "'a :: {banach, real_normed_field, ln} \ 'a" where "artanh x = ln ((1 + x) / (1 - x)) / 2" lemma arsinh_0 [simp]: "arsinh 0 = 0" by (simp add: arsinh_def) lemma arcosh_1 [simp]: "arcosh 1 = 0" by (simp add: arcosh_def) lemma artanh_0 [simp]: "artanh 0 = 0" by (simp add: artanh_def) lemma tanh_altdef: "tanh x = (exp x - exp (-x)) / (exp x + exp (-x))" proof - have "tanh x = (2 *\<^sub>R sinh x) / (2 *\<^sub>R cosh x)" by (simp add: tanh_def scaleR_conv_of_real) also have "2 *\<^sub>R sinh x = exp x - exp (-x)" by (simp add: sinh_def) also have "2 *\<^sub>R cosh x = exp x + exp (-x)" by (simp add: cosh_def) finally show ?thesis . qed lemma tanh_real_altdef: "tanh (x::real) = (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))" proof - have [simp]: "exp (2 * x) = exp x * exp x" "exp (x * 2) = exp x * exp x" by (subst exp_add [symmetric]; simp)+ have "tanh x = (2 * exp (-x) * sinh x) / (2 * exp (-x) * cosh x)" by (simp add: tanh_def) also have "2 * exp (-x) * sinh x = 1 - exp (-2*x)" by (simp add: exp_minus field_simps sinh_def) also have "2 * exp (-x) * cosh x = 1 + exp (-2*x)" by (simp add: exp_minus field_simps cosh_def) finally show ?thesis . qed lemma sinh_converges: "(\n. if even n then 0 else x ^ n /\<^sub>R fact n) sums sinh x" proof - have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums sinh x" unfolding sinh_def by (intro sums_scaleR_right sums_diff exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n - (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then 0 else x ^ n /\<^sub>R fact n)" by auto finally show ?thesis . qed lemma cosh_converges: "(\n. if even n then x ^ n /\<^sub>R fact n else 0) sums cosh x" proof - have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) sums cosh x" unfolding cosh_def by (intro sums_scaleR_right sums_add exp_converges) also have "(\n. (x ^ n /\<^sub>R fact n + (-x) ^ n /\<^sub>R fact n) /\<^sub>R 2) = (\n. if even n then x ^ n /\<^sub>R fact n else 0)" by auto finally show ?thesis . qed lemma sinh_0 [simp]: "sinh 0 = 0" by (simp add: sinh_def) lemma cosh_0 [simp]: "cosh 0 = 1" proof - have "cosh 0 = (1/2) *\<^sub>R (1 + 1)" by (simp add: cosh_def) also have "\ = 1" by (rule scaleR_half_double) finally show ?thesis . qed lemma tanh_0 [simp]: "tanh 0 = 0" by (simp add: tanh_def) lemma sinh_minus [simp]: "sinh (- x) = -sinh x" by (simp add: sinh_def algebra_simps) lemma cosh_minus [simp]: "cosh (- x) = cosh x" by (simp add: cosh_def algebra_simps) lemma tanh_minus [simp]: "tanh (-x) = -tanh x" by (simp add: tanh_def) lemma sinh_ln_real: "x > 0 \ sinh (ln x :: real) = (x - inverse x) / 2" by (simp add: sinh_def exp_minus) lemma cosh_ln_real: "x > 0 \ cosh (ln x :: real) = (x + inverse x) / 2" by (simp add: cosh_def exp_minus) lemma tanh_ln_real: "tanh (ln x :: real) = (x ^ 2 - 1) / (x ^ 2 + 1)" if "x > 0" proof - from that have "(x * 2 - inverse x * 2) * (x\<^sup>2 + 1) = (x\<^sup>2 - 1) * (2 * x + 2 * inverse x)" by (simp add: field_simps power2_eq_square) moreover have "x\<^sup>2 + 1 > 0" using that by (simp add: ac_simps add_pos_nonneg) moreover have "2 * x + 2 * inverse x > 0" using that by (simp add: add_pos_pos) ultimately have "(x * 2 - inverse x * 2) / (2 * x + 2 * inverse x) = (x\<^sup>2 - 1) / (x\<^sup>2 + 1)" by (simp add: frac_eq_eq) with that show ?thesis by (simp add: tanh_def sinh_ln_real cosh_ln_real) qed lemma has_field_derivative_scaleR_right [derivative_intros]: "(f has_field_derivative D) F \ ((\x. c *\<^sub>R f x) has_field_derivative (c *\<^sub>R D)) F" unfolding has_field_derivative_def using has_derivative_scaleR_right[of f "\x. D * x" F c] by (simp add: mult_scaleR_left [symmetric] del: mult_scaleR_left) lemma has_field_derivative_sinh [THEN DERIV_chain2, derivative_intros]: "(sinh has_field_derivative cosh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_cosh [THEN DERIV_chain2, derivative_intros]: "(cosh has_field_derivative sinh x) (at (x :: 'a :: {banach, real_normed_field}))" unfolding sinh_def cosh_def by (auto intro!: derivative_eq_intros) lemma has_field_derivative_tanh [THEN DERIV_chain2, derivative_intros]: "cosh x \ 0 \ (tanh has_field_derivative 1 - tanh x ^ 2) (at (x :: 'a :: {banach, real_normed_field}))" unfolding tanh_def by (auto intro!: derivative_eq_intros simp: power2_eq_square field_split_simps) lemma has_derivative_sinh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\x. Db * x)) (at x within s)" shows "((\x. sinh (g x)) has_derivative (\y. (cosh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\x. (-Db) * x)" by (simp add: fun_eq_iff) finally have "((\x. sinh (g x)) has_derivative (\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding sinh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y - exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (cosh (g x) * Db) * y)" by (simp add: fun_eq_iff cosh_def algebra_simps) finally show ?thesis . qed lemma has_derivative_cosh [derivative_intros]: fixes g :: "'a \ ('a :: {banach, real_normed_field})" assumes "(g has_derivative (\y. Db * y)) (at x within s)" shows "((\x. cosh (g x)) has_derivative (\y. (sinh (g x) * Db) * y)) (at x within s)" proof - have "((\x. - g x) has_derivative (\y. -(Db * y))) (at x within s)" using assms by (intro derivative_intros) also have "(\y. -(Db * y)) = (\y. (-Db) * y)" by (simp add: fun_eq_iff) finally have "((\x. cosh (g x)) has_derivative (\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2)) (at x within s)" unfolding cosh_def by (intro derivative_intros assms) also have "(\y. (exp (g x) * Db * y + exp (-g x) * (-Db) * y) /\<^sub>R 2) = (\y. (sinh (g x) * Db) * y)" by (simp add: fun_eq_iff sinh_def algebra_simps) finally show ?thesis . qed lemma sinh_plus_cosh: "sinh x + cosh x = exp x" proof - have "sinh x + cosh x = (1 / 2) *\<^sub>R (exp x + exp x)" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp x" by (rule scaleR_half_double) finally show ?thesis . qed lemma cosh_plus_sinh: "cosh x + sinh x = exp x" by (subst add.commute) (rule sinh_plus_cosh) lemma cosh_minus_sinh: "cosh x - sinh x = exp (-x)" proof - have "cosh x - sinh x = (1 / 2) *\<^sub>R (exp (-x) + exp (-x))" by (simp add: sinh_def cosh_def algebra_simps) also have "\ = exp (-x)" by (rule scaleR_half_double) finally show ?thesis . qed lemma sinh_minus_cosh: "sinh x - cosh x = -exp (-x)" using cosh_minus_sinh[of x] by (simp add: algebra_simps) context fixes x :: "'a :: {real_normed_field, banach}" begin lemma sinh_zero_iff: "sinh x = 0 \ exp x \ {1, -1}" by (auto simp: sinh_def field_simps exp_minus power2_eq_square square_eq_1_iff) lemma cosh_zero_iff: "cosh x = 0 \ exp x ^ 2 = -1" by (auto simp: cosh_def exp_minus field_simps power2_eq_square eq_neg_iff_add_eq_0) lemma cosh_square_eq: "cosh x ^ 2 = sinh x ^ 2 + 1" by (simp add: cosh_def sinh_def algebra_simps power2_eq_square exp_add [symmetric] scaleR_conv_of_real) lemma sinh_square_eq: "sinh x ^ 2 = cosh x ^ 2 - 1" by (simp add: cosh_square_eq) lemma hyperbolic_pythagoras: "cosh x ^ 2 - sinh x ^ 2 = 1" by (simp add: cosh_square_eq) lemma sinh_add: "sinh (x + y) = sinh x * cosh y + cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma sinh_diff: "sinh (x - y) = sinh x * cosh y - cosh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_add: "cosh (x + y) = cosh x * cosh y + sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma cosh_diff: "cosh (x - y) = cosh x * cosh y - sinh x * sinh y" by (simp add: sinh_def cosh_def algebra_simps scaleR_conv_of_real exp_add [symmetric]) lemma tanh_add: "tanh (x + y) = (tanh x + tanh y) / (1 + tanh x * tanh y)" if "cosh x \ 0" "cosh y \ 0" proof - have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (cosh x * cosh y + sinh x * sinh y) * ((sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x))" using that by (simp add: field_split_simps) also have "(sinh x * cosh y + sinh y * cosh x) / (cosh y * cosh x) = sinh x / cosh x + sinh y / cosh y" using that by (simp add: field_split_simps) finally have "(sinh x * cosh y + cosh x * sinh y) * (1 + sinh x * sinh y / (cosh x * cosh y)) = (sinh x / cosh x + sinh y / cosh y) * (cosh x * cosh y + sinh x * sinh y)" by simp then show ?thesis using that by (auto simp add: tanh_def sinh_add cosh_add eq_divide_eq) (simp_all add: field_split_simps) qed lemma sinh_double: "sinh (2 * x) = 2 * sinh x * cosh x" using sinh_add[of x] by simp lemma cosh_double: "cosh (2 * x) = cosh x ^ 2 + sinh x ^ 2" using cosh_add[of x] by (simp add: power2_eq_square) end lemma sinh_field_def: "sinh z = (exp z - exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: sinh_def scaleR_conv_of_real) lemma cosh_field_def: "cosh z = (exp z + exp (-z)) / (2 :: 'a :: {banach, real_normed_field})" by (simp add: cosh_def scaleR_conv_of_real) subsubsection \More specific properties of the real functions\ lemma sinh_real_zero_iff [simp]: "sinh (x::real) = 0 \ x = 0" proof - have "(-1 :: real) < 0" by simp also have "0 < exp x" by simp finally have "exp x \ -1" by (intro notI) simp thus ?thesis by (subst sinh_zero_iff) simp qed lemma plus_inverse_ge_2: fixes x :: real assumes "x > 0" shows "x + inverse x \ 2" proof - have "0 \ (x - 1) ^ 2" by simp also have "\ = x^2 - 2*x + 1" by (simp add: power2_eq_square algebra_simps) finally show ?thesis using assms by (simp add: field_simps power2_eq_square) qed lemma sinh_real_nonneg_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_pos_iff [simp]: "sinh (x :: real) > 0 \ x > 0" by (simp add: sinh_def) lemma sinh_real_nonpos_iff [simp]: "sinh (x :: real) \ 0 \ x \ 0" by (simp add: sinh_def) lemma sinh_real_neg_iff [simp]: "sinh (x :: real) < 0 \ x < 0" by (simp add: sinh_def) lemma cosh_real_ge_1: "cosh (x :: real) \ 1" using plus_inverse_ge_2[of "exp x"] by (simp add: cosh_def exp_minus) lemma cosh_real_pos [simp]: "cosh (x :: real) > 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonneg[simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma cosh_real_nonzero [simp]: "cosh (x :: real) \ 0" using cosh_real_ge_1[of x] by simp lemma tanh_real_nonneg_iff [simp]: "tanh (x :: real) \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_pos_iff [simp]: "tanh (x :: real) > 0 \ x > 0" by (simp add: tanh_def field_simps) lemma tanh_real_nonpos_iff [simp]: "tanh (x :: real) \ 0 \ x \ 0" by (simp add: tanh_def field_simps) lemma tanh_real_neg_iff [simp]: "tanh (x :: real) < 0 \ x < 0" by (simp add: tanh_def field_simps) lemma tanh_real_zero_iff [simp]: "tanh (x :: real) = 0 \ x = 0" by (simp add: tanh_def field_simps) lemma arsinh_real_def: "arsinh (x::real) = ln (x + sqrt (x^2 + 1))" by (simp add: arsinh_def powr_half_sqrt) lemma arcosh_real_def: "x \ 1 \ arcosh (x::real) = ln (x + sqrt (x^2 - 1))" by (simp add: arcosh_def powr_half_sqrt) lemma arsinh_real_aux: "0 < x + sqrt (x ^ 2 + 1 :: real)" proof (cases "x < 0") case True have "(-x) ^ 2 = x ^ 2" by simp also have "x ^ 2 < x ^ 2 + 1" by simp finally have "sqrt ((-x) ^ 2) < sqrt (x ^ 2 + 1)" by (rule real_sqrt_less_mono) thus ?thesis using True by simp qed (auto simp: add_nonneg_pos) lemma arsinh_minus_real [simp]: "arsinh (-x::real) = -arsinh x" proof - have "arsinh (-x) = ln (sqrt (x\<^sup>2 + 1) - x)" by (simp add: arsinh_real_def) also have "sqrt (x^2 + 1) - x = inverse (sqrt (x^2 + 1) + x)" using arsinh_real_aux[of x] by (simp add: field_split_simps algebra_simps power2_eq_square) also have "ln \ = -arsinh x" using arsinh_real_aux[of x] by (simp add: arsinh_real_def ln_inverse) finally show ?thesis . qed lemma artanh_minus_real [simp]: assumes "abs x < 1" shows "artanh (-x::real) = -artanh x" using assms by (simp add: artanh_def ln_div field_simps) lemma sinh_less_cosh_real: "sinh (x :: real) < cosh x" by (simp add: sinh_def cosh_def) lemma sinh_le_cosh_real: "sinh (x :: real) \ cosh x" by (simp add: sinh_def cosh_def) lemma tanh_real_lt_1: "tanh (x :: real) < 1" by (simp add: tanh_def sinh_less_cosh_real) lemma tanh_real_gt_neg1: "tanh (x :: real) > -1" proof - have "- cosh x < sinh x" by (simp add: sinh_def cosh_def field_split_simps) thus ?thesis by (simp add: tanh_def field_simps) qed lemma tanh_real_bounds: "tanh (x :: real) \ {-1<..<1}" using tanh_real_lt_1 tanh_real_gt_neg1 by simp context fixes x :: real begin lemma arsinh_sinh_real: "arsinh (sinh x) = x" by (simp add: arsinh_real_def powr_def sinh_square_eq sinh_plus_cosh) lemma arcosh_cosh_real: "x \ 0 \ arcosh (cosh x) = x" by (simp add: arcosh_real_def powr_def cosh_square_eq cosh_real_ge_1 cosh_plus_sinh) lemma artanh_tanh_real: "artanh (tanh x) = x" proof - have "artanh (tanh x) = ln (cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x))) / 2" by (simp add: artanh_def tanh_def field_split_simps) also have "cosh x * (cosh x + sinh x) / (cosh x * (cosh x - sinh x)) = (cosh x + sinh x) / (cosh x - sinh x)" by simp also have "\ = (exp x)^2" by (simp add: cosh_plus_sinh cosh_minus_sinh exp_minus field_simps power2_eq_square) also have "ln ((exp x)^2) / 2 = x" by (simp add: ln_realpow) finally show ?thesis . qed end lemma sinh_real_strict_mono: "strict_mono (sinh :: real \ real)" by (rule pos_deriv_imp_strict_mono derivative_intros)+ auto lemma cosh_real_strict_mono: assumes "0 \ x" and "x < (y::real)" shows "cosh x < cosh y" proof - from assms have "\z>x. z < y \ cosh y - cosh x = (y - x) * sinh z" by (intro MVT2) (auto dest: connectedD_interval intro!: derivative_eq_intros) then obtain z where z: "z > x" "z < y" "cosh y - cosh x = (y - x) * sinh z" by blast note \cosh y - cosh x = (y - x) * sinh z\ also from \z > x\ and assms have "(y - x) * sinh z > 0" by (intro mult_pos_pos) auto finally show "cosh x < cosh y" by simp qed lemma tanh_real_strict_mono: "strict_mono (tanh :: real \ real)" proof - have *: "tanh x ^ 2 < 1" for x :: real using tanh_real_bounds[of x] by (simp add: abs_square_less_1 abs_if) show ?thesis by (rule pos_deriv_imp_strict_mono) (insert *, auto intro!: derivative_intros) qed lemma sinh_real_abs [simp]: "sinh (abs x :: real) = abs (sinh x)" by (simp add: abs_if) lemma cosh_real_abs [simp]: "cosh (abs x :: real) = cosh x" by (simp add: abs_if) lemma tanh_real_abs [simp]: "tanh (abs x :: real) = abs (tanh x)" by (auto simp: abs_if) lemma sinh_real_eq_iff [simp]: "sinh x = sinh y \ x = (y :: real)" using sinh_real_strict_mono by (simp add: strict_mono_eq) lemma tanh_real_eq_iff [simp]: "tanh x = tanh y \ x = (y :: real)" using tanh_real_strict_mono by (simp add: strict_mono_eq) lemma cosh_real_eq_iff [simp]: "cosh x = cosh y \ abs x = abs (y :: real)" proof - have "cosh x = cosh y \ x = y" if "x \ 0" "y \ 0" for x y :: real using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] that by (cases x y rule: linorder_cases) auto from this[of "abs x" "abs y"] show ?thesis by simp qed lemma sinh_real_le_iff [simp]: "sinh x \ sinh y \ x \ (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less_eq) lemma cosh_real_nonneg_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_le_iff: "x \ 0 \ y \ 0 \ cosh x \ cosh y \ x \ (y::real)" using cosh_real_nonneg_le_iff[of "-x" "-y"] by simp lemma tanh_real_le_iff [simp]: "tanh x \ tanh y \ x \ (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less_eq) lemma sinh_real_less_iff [simp]: "sinh x < sinh y \ x < (y::real)" using sinh_real_strict_mono by (simp add: strict_mono_less) lemma cosh_real_nonneg_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x < (y::real)" using cosh_real_strict_mono[of x y] cosh_real_strict_mono[of y x] by (cases x y rule: linorder_cases) auto lemma cosh_real_nonpos_less_iff: "x \ 0 \ y \ 0 \ cosh x < cosh y \ x > (y::real)" using cosh_real_nonneg_less_iff[of "-x" "-y"] by simp lemma tanh_real_less_iff [simp]: "tanh x < tanh y \ x < (y::real)" using tanh_real_strict_mono by (simp add: strict_mono_less) subsubsection \Limits\ lemma sinh_real_at_top: "filterlim (sinh :: real \ real) at_top at_top" proof - have *: "((\x. - exp (- x)) \ (-0::real)) at_top" by (intro tendsto_minus filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) have "filterlim (\x. (1 / 2) * (-exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) also have "(\x. (1 / 2) * (-exp (-x) + exp x) :: real) = sinh" by (simp add: fun_eq_iff sinh_def) finally show ?thesis . qed lemma sinh_real_at_bot: "filterlim (sinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x. -sinh x :: real) at_bot at_top" by (simp add: filterlim_uminus_at_top [symmetric] sinh_real_at_top) also have "(\x. -sinh x :: real) = (\x. sinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma cosh_real_at_top: "filterlim (cosh :: real \ real) at_top at_top" proof - have *: "((\x. exp (- x)) \ (0::real)) at_top" by (intro filterlim_compose[OF exp_at_bot] filterlim_uminus_at_bot_at_top) have "filterlim (\x. (1 / 2) * (exp (-x) + exp x) :: real) at_top at_top" by (rule filterlim_tendsto_pos_mult_at_top[OF _ _ filterlim_tendsto_add_at_top[OF *]] tendsto_const)+ (auto simp: exp_at_top) also have "(\x. (1 / 2) * (exp (-x) + exp x) :: real) = cosh" by (simp add: fun_eq_iff cosh_def) finally show ?thesis . qed lemma cosh_real_at_bot: "filterlim (cosh :: real \ real) at_top at_bot" proof - have "filterlim (\x. cosh (-x) :: real) at_top at_top" by (simp add: cosh_real_at_top) thus ?thesis by (subst filterlim_at_bot_mirror) qed lemma tanh_real_at_top: "(tanh \ (1::real)) at_top" proof - have "((\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) \ (1 - 0) / (1 + 0)) at_top" by (intro tendsto_intros filterlim_compose[OF exp_at_bot] filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_ident) auto also have "(\x::real. (1 - exp (- 2 * x)) / (1 + exp (- 2 * x))) = tanh" by (rule ext) (simp add: tanh_real_altdef) finally show ?thesis by simp qed lemma tanh_real_at_bot: "(tanh \ (-1::real)) at_bot" proof - have "((\x::real. -tanh x) \ -1) at_top" by (intro tendsto_minus tanh_real_at_top) also have "(\x. -tanh x :: real) = (\x. tanh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed subsubsection \Properties of the inverse hyperbolic functions\ lemma isCont_sinh: "isCont sinh (x :: 'a :: {real_normed_field, banach})" unfolding sinh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_cosh: "isCont cosh (x :: 'a :: {real_normed_field, banach})" unfolding cosh_def [abs_def] by (auto intro!: continuous_intros) lemma isCont_tanh: "cosh x \ 0 \ isCont tanh (x :: 'a :: {real_normed_field, banach})" unfolding tanh_def [abs_def] by (auto intro!: continuous_intros isCont_divide isCont_sinh isCont_cosh) lemma continuous_on_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_on_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" shows "continuous_on A (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_sinh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. sinh (f x))" unfolding sinh_def using assms by (intro continuous_intros) lemma continuous_cosh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" shows "continuous F (\x. cosh (f x))" unfolding cosh_def using assms by (intro continuous_intros) lemma continuous_on_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous_on A f" "\x. x \ A \ cosh (f x) \ 0" shows "continuous_on A (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros) auto lemma continuous_at_within_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous (at x within A) f" "cosh (f x) \ 0" shows "continuous (at x within A) (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma continuous_tanh [continuous_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" assumes "continuous F f" "cosh (f (Lim F (\x. x))) \ 0" shows "continuous F (\x. tanh (f x))" unfolding tanh_def using assms by (intro continuous_intros continuous_divide) auto lemma tendsto_sinh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. sinh (f x)) \ sinh a) F" by (rule isCont_tendsto_compose [OF isCont_sinh]) lemma tendsto_cosh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ ((\x. cosh (f x)) \ cosh a) F" by (rule isCont_tendsto_compose [OF isCont_cosh]) lemma tendsto_tanh [tendsto_intros]: fixes f :: "_ \'a::{real_normed_field,banach}" shows "(f \ a) F \ cosh a \ 0 \ ((\x. tanh (f x)) \ tanh a) F" by (rule isCont_tendsto_compose [OF isCont_tanh]) lemma arsinh_real_has_field_derivative [derivative_intros]: fixes x :: real shows "(arsinh has_field_derivative (1 / (sqrt (x ^ 2 + 1)))) (at x within A)" proof - have pos: "1 + x ^ 2 > 0" by (intro add_pos_nonneg) auto from pos arsinh_real_aux[of x] show ?thesis unfolding arsinh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps) qed lemma arcosh_real_has_field_derivative [derivative_intros]: fixes x :: real assumes "x > 1" shows "(arcosh has_field_derivative (1 / (sqrt (x ^ 2 - 1)))) (at x within A)" proof - from assms have "x + sqrt (x\<^sup>2 - 1) > 0" by (simp add: add_pos_pos) thus ?thesis using assms unfolding arcosh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt field_split_simps power2_eq_1_iff) qed lemma artanh_real_has_field_derivative [derivative_intros]: "(artanh has_field_derivative (1 / (1 - x ^ 2))) (at x within A)" if "\x\ < 1" for x :: real proof - from that have "- 1 < x" "x < 1" by linarith+ hence "(artanh has_field_derivative (4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4)) (at x within A)" unfolding artanh_def [abs_def] by (auto intro!: derivative_eq_intros simp: powr_minus powr_half_sqrt) also have "(4 - 4 * x) / ((1 + x) * (1 - x) * (1 - x) * 4) = 1 / ((1 + x) * (1 - x))" using \-1 < x\ \x < 1\ by (simp add: frac_eq_eq) also have "(1 + x) * (1 - x) = 1 - x ^ 2" by (simp add: algebra_simps power2_eq_square) finally show ?thesis . qed lemma continuous_on_arsinh [continuous_intros]: "continuous_on A (arsinh :: real \ real)" by (rule DERIV_continuous_on derivative_intros)+ lemma continuous_on_arcosh [continuous_intros]: assumes "A \ {1..}" shows "continuous_on A (arcosh :: real \ real)" proof - have pos: "x + sqrt (x ^ 2 - 1) > 0" if "x \ 1" for x using that by (intro add_pos_nonneg) auto show ?thesis unfolding arcosh_def [abs_def] by (intro continuous_on_subset [OF _ assms] continuous_on_ln continuous_on_add continuous_on_id continuous_on_powr') (auto dest: pos simp: powr_half_sqrt intro!: continuous_intros) qed lemma continuous_on_artanh [continuous_intros]: assumes "A \ {-1<..<1}" shows "continuous_on A (artanh :: real \ real)" unfolding artanh_def [abs_def] by (intro continuous_on_subset [OF _ assms]) (auto intro!: continuous_intros) lemma continuous_on_arsinh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" shows "continuous_on A (\x. arsinh (f x))" by (rule continuous_on_compose2[OF continuous_on_arsinh assms]) auto lemma continuous_on_arcosh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ 1" shows "continuous_on A (\x. arcosh (f x))" by (rule continuous_on_compose2[OF continuous_on_arcosh assms(1) order.refl]) (use assms(2) in auto) lemma continuous_on_artanh' [continuous_intros]: fixes f :: "real \ real" assumes "continuous_on A f" "\x. x \ A \ f x \ {-1<..<1}" shows "continuous_on A (\x. artanh (f x))" by (rule continuous_on_compose2[OF continuous_on_artanh assms(1) order.refl]) (use assms(2) in auto) lemma isCont_arsinh [continuous_intros]: "isCont arsinh (x :: real)" using continuous_on_arsinh[of UNIV] by (auto simp: continuous_on_eq_continuous_at) lemma isCont_arcosh [continuous_intros]: assumes "x > 1" shows "isCont arcosh (x :: real)" proof - have "continuous_on {1::real<..} arcosh" by (rule continuous_on_arcosh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma isCont_artanh [continuous_intros]: assumes "x > -1" "x < 1" shows "isCont artanh (x :: real)" proof - have "continuous_on {-1<..<(1::real)} artanh" by (rule continuous_on_artanh) auto with assms show ?thesis by (auto simp: continuous_on_eq_continuous_at) qed lemma tendsto_arsinh [tendsto_intros]: "(f \ a) F \ ((\x. arsinh (f x)) \ arsinh a) F" for f :: "_ \ real" by (rule isCont_tendsto_compose [OF isCont_arsinh]) lemma tendsto_arcosh_strong [tendsto_intros]: fixes f :: "_ \ real" assumes "(f \ a) F" "a \ 1" "eventually (\x. f x \ 1) F" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule continuous_on_tendsto_compose[OF continuous_on_arcosh[OF order.refl]]) (use assms in auto) lemma tendsto_arcosh: fixes f :: "_ \ real" assumes "(f \ a) F" "a > 1" shows "((\x. arcosh (f x)) \ arcosh a) F" by (rule isCont_tendsto_compose [OF isCont_arcosh]) (use assms in auto) lemma tendsto_arcosh_at_left_1: "(arcosh \ 0) (at_right (1::real))" proof - have "(arcosh \ arcosh 1) (at_right (1::real))" by (rule tendsto_arcosh_strong) (auto simp: eventually_at intro!: exI[of _ 1]) thus ?thesis by simp qed lemma tendsto_artanh [tendsto_intros]: fixes f :: "'a \ real" assumes "(f \ a) F" "a > -1" "a < 1" shows "((\x. artanh (f x)) \ artanh a) F" by (rule isCont_tendsto_compose [OF isCont_artanh]) (use assms in auto) lemma continuous_arsinh [continuous_intros]: "continuous F f \ continuous F (\x. arsinh (f x :: real))" unfolding continuous_def by (rule tendsto_arsinh) (* TODO: This rule does not work for one-sided continuity at 1 *) lemma continuous_arcosh_strong [continuous_intros]: assumes "continuous F f" "eventually (\x. f x \ 1) F" shows "continuous F (\x. arcosh (f x :: real))" proof (cases "F = bot") case False show ?thesis unfolding continuous_def proof (intro tendsto_arcosh_strong) show "1 \ f (Lim F (\x. x))" using assms False unfolding continuous_def by (rule tendsto_lowerbound) qed (insert assms, auto simp: continuous_def) qed auto lemma continuous_arcosh: "continuous F f \ f (Lim F (\x. x)) > 1 \ continuous F (\x. arcosh (f x :: real))" unfolding continuous_def by (rule tendsto_arcosh) auto lemma continuous_artanh [continuous_intros]: "continuous F f \ f (Lim F (\x. x)) \ {-1<..<1} \ continuous F (\x. artanh (f x :: real))" unfolding continuous_def by (rule tendsto_artanh) auto lemma arsinh_real_at_top: "filterlim (arsinh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arsinh_real_def add_ac) lemma arsinh_real_at_bot: "filterlim (arsinh :: real \ real) at_bot at_bot" proof - have "filterlim (\x::real. -arsinh x) at_bot at_top" by (subst filterlim_uminus_at_top [symmetric]) (rule arsinh_real_at_top) also have "(\x::real. -arsinh x) = (\x. arsinh (-x))" by simp finally show ?thesis by (subst filterlim_at_bot_mirror) qed lemma arcosh_real_at_top: "filterlim (arcosh :: real \ real) at_top at_top" proof (subst filterlim_cong[OF refl refl]) show "filterlim (\x. ln (x + sqrt (-1 + x\<^sup>2))) at_top at_top" by (intro filterlim_compose[OF ln_at_top filterlim_at_top_add_at_top] filterlim_ident filterlim_compose[OF sqrt_at_top] filterlim_tendsto_add_at_top[OF tendsto_const] filterlim_pow_at_top) auto qed (auto intro!: eventually_mono[OF eventually_ge_at_top[of 1]] simp: arcosh_real_def) lemma artanh_real_at_left_1: "filterlim (artanh :: real \ real) at_top (at_left 1)" proof - have *: "filterlim (\x::real. (1 + x) / (1 - x)) at_top (at_left 1)" by (rule LIM_at_top_divide) (auto intro!: tendsto_eq_intros eventually_mono[OF eventually_at_left_real[of 0]]) have "filterlim (\x::real. (1/2) * ln ((1 + x) / (1 - x))) at_top (at_left 1)" by (intro filterlim_tendsto_pos_mult_at_top[OF tendsto_const] * filterlim_compose[OF ln_at_top]) auto also have "(\x::real. (1/2) * ln ((1 + x) / (1 - x))) = artanh" by (simp add: artanh_def [abs_def]) finally show ?thesis . qed lemma artanh_real_at_right_1: "filterlim (artanh :: real \ real) at_bot (at_right (-1))" proof - have "?thesis \ filterlim (\x::real. -artanh x) at_top (at_right (-1))" by (simp add: filterlim_uminus_at_bot) also have "\ \ filterlim (\x::real. artanh (-x)) at_top (at_right (-1))" by (intro filterlim_cong refl eventually_mono[OF eventually_at_right_real[of "-1" "1"]]) auto also have "\ \ filterlim (artanh :: real \ real) at_top (at_left 1)" by (simp add: filterlim_at_left_to_right) also have \ by (rule artanh_real_at_left_1) finally show ?thesis . qed subsection \Simprocs for root and power literals\ lemma numeral_powr_numeral_real [simp]: "numeral m powr numeral n = (numeral m ^ numeral n :: real)" by (simp add: powr_numeral) context begin private lemma sqrt_numeral_simproc_aux: assumes "m * m \ n" shows "sqrt (numeral n :: real) \ numeral m" proof - have "numeral n \ numeral m * (numeral m :: real)" by (simp add: assms [symmetric]) moreover have "sqrt \ \ numeral m" by (subst real_sqrt_abs2) simp ultimately show "sqrt (numeral n :: real) \ numeral m" by simp qed private lemma root_numeral_simproc_aux: assumes "Num.pow m n \ x" shows "root (numeral n) (numeral x :: real) \ numeral m" by (subst assms [symmetric], subst numeral_pow, subst real_root_pos2) simp_all private lemma powr_numeral_simproc_aux: assumes "Num.pow y n = x" shows "numeral x powr (m / numeral n :: real) \ numeral y powr m" by (subst assms [symmetric], subst numeral_pow, subst powr_numeral [symmetric]) (simp, subst powr_powr, simp_all) private lemma numeral_powr_inverse_eq: "numeral x powr (inverse (numeral n)) = numeral x powr (1 / numeral n :: real)" by simp ML \ signature ROOT_NUMERAL_SIMPROC = sig val sqrt : int option -> int -> int option val sqrt' : int option -> int -> int option val nth_root : int option -> int -> int -> int option val nth_root' : int option -> int -> int -> int option val sqrt_simproc : Proof.context -> cterm -> thm option val root_simproc : int * int -> Proof.context -> cterm -> thm option val powr_simproc : int * int -> Proof.context -> cterm -> thm option end structure Root_Numeral_Simproc : ROOT_NUMERAL_SIMPROC = struct fun iterate NONE p f x = let fun go x = if p x then x else go (f x) in SOME (go x) end | iterate (SOME threshold) p f x = let fun go (threshold, x) = if p x then SOME x else if threshold = 0 then NONE else go (threshold - 1, f x) in go (threshold, x) end fun nth_root _ 1 x = SOME x | nth_root _ _ 0 = SOME 0 | nth_root _ _ 1 = SOME 1 | nth_root threshold n x = let fun newton_step y = ((n - 1) * y + x div Integer.pow (n - 1) y) div n fun is_root y = Integer.pow n y <= x andalso x < Integer.pow n (y + 1) in if x < n then SOME 1 else if x < Integer.pow n 2 then SOME 1 else let val y = Real.floor (Math.pow (Real.fromInt x, Real.fromInt 1 / Real.fromInt n)) in if is_root y then SOME y else iterate threshold is_root newton_step ((x + n - 1) div n) end end fun nth_root' _ 1 x = SOME x | nth_root' _ _ 0 = SOME 0 | nth_root' _ _ 1 = SOME 1 | nth_root' threshold n x = if x < n then NONE else if x < Integer.pow n 2 then NONE else case nth_root threshold n x of NONE => NONE | SOME y => if Integer.pow n y = x then SOME y else NONE fun sqrt _ 0 = SOME 0 | sqrt _ 1 = SOME 1 | sqrt threshold n = let fun aux (a, b) = if n >= b * b then aux (b, b * b) else (a, b) val (lower_root, lower_n) = aux (1, 2) fun newton_step x = (x + n div x) div 2 fun is_sqrt r = r*r <= n andalso n < (r+1)*(r+1) val y = Real.floor (Math.sqrt (Real.fromInt n)) in if is_sqrt y then SOME y else Option.mapPartial (iterate threshold is_sqrt newton_step o (fn x => x * lower_root)) (sqrt threshold (n div lower_n)) end fun sqrt' threshold x = case sqrt threshold x of NONE => NONE | SOME y => if y * y = x then SOME y else NONE fun sqrt_simproc ctxt ct = let val n = ct |> Thm.term_of |> dest_comb |> snd |> dest_comb |> snd |> HOLogic.dest_numeral in case sqrt' (SOME 10000) n of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n]) @{thm sqrt_numeral_simproc_aux}) end handle TERM _ => NONE fun root_simproc (threshold1, threshold2) ctxt ct = let val [n, x] = ct |> Thm.term_of |> strip_comb |> snd |> map (dest_comb #> snd #> HOLogic.dest_numeral) in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME m => SOME (Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt o HOLogic.mk_numeral) [m, n, x]) @{thm root_numeral_simproc_aux}) end handle TERM _ => NONE | Match => NONE fun powr_simproc (threshold1, threshold2) ctxt ct = let val eq_thm = Conv.try_conv (Conv.rewr_conv @{thm numeral_powr_inverse_eq}) ct val ct = Thm.dest_equals_rhs (Thm.cprop_of eq_thm) val (_, [x, t]) = strip_comb (Thm.term_of ct) val (_, [m, n]) = strip_comb t val [x, n] = map (dest_comb #> snd #> HOLogic.dest_numeral) [x, n] in if n > threshold1 orelse x > threshold2 then NONE else case nth_root' (SOME 100) n x of NONE => NONE | SOME y => let val [y, n, x] = map HOLogic.mk_numeral [y, n, x] val thm = Thm.instantiate' [] (map (SOME o Thm.cterm_of ctxt) [y, n, x, m]) @{thm powr_numeral_simproc_aux} in SOME (@{thm transitive} OF [eq_thm, thm]) end end handle TERM _ => NONE | Match => NONE end \ end simproc_setup sqrt_numeral ("sqrt (numeral n)") = \K Root_Numeral_Simproc.sqrt_simproc\ simproc_setup root_numeral ("root (numeral n) (numeral x)") = \K (Root_Numeral_Simproc.root_simproc (200, Integer.pow 200 2))\ simproc_setup powr_divide_numeral ("numeral x powr (m / numeral n :: real)" | "numeral x powr (inverse (numeral n) :: real)") = \K (Root_Numeral_Simproc.powr_simproc (200, Integer.pow 200 2))\ lemma "root 100 1267650600228229401496703205376 = 2" by simp lemma "sqrt 196 = 14" by simp lemma "256 powr (7 / 4 :: real) = 16384" by simp lemma "27 powr (inverse 3) = (3::real)" by simp end