src/HOL/Analysis/Complex_Transcendental.thy
author Manuel Eberl <manuel@pruvisto.org>
Tue, 15 Apr 2025 17:38:20 +0200
changeset 82518 da14e77a48b2
parent 82459 a1de627d417a
permissions -rw-r--r--
lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
     1
section \<open>Complex Transcendental Functions\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     2
61711
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
     3
text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2015)\<close>
21d7910d6816 Theory of homotopic paths (from HOL Light), plus comments and minor refinements
paulson <lp15@cam.ac.uk>
parents: 61694
diff changeset
     4
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     5
theory Complex_Transcendental
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
     6
imports
70196
b7ef9090feed Added embedding_map_into_euclideanreal; reduced dependence on Equivalence_Lebesgue_Henstock_Integration in Analysis theories by moving a few lemmas
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
     7
  Complex_Analysis_Basics Summation_Tests "HOL-Library.Periodic_Fun"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
begin
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     9
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    10
subsection\<open>Möbius transformations\<close>
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    11
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    12
(* TODO: Figure out what to do with Möbius transformations *)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
    13
definition\<^marker>\<open>tag important\<close> "moebius a b c d \<equiv> (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))"
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    14
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    15
theorem moebius_inverse:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    16
  assumes "a * d \<noteq> b * c" "c * z + d \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    17
  shows   "moebius d (-b) (-c) a (moebius a b c d z) = z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    18
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    19
  from assms have "(-c) * moebius a b c d z + a \<noteq> 0" unfolding moebius_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    20
    by (simp add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    21
  with assms show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    22
    unfolding moebius_def by (simp add: moebius_def divide_simps) (simp add: algebra_simps)?
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    23
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    24
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    25
lemma moebius_inverse':
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    26
  assumes "a * d \<noteq> b * c" "c * z - a \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    27
  shows   "moebius a b c d (moebius d (-b) (-c) a z) = z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    28
  using assms moebius_inverse[of d a "-b" "-c" z]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    29
  by (auto simp: algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    30
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    31
lemma cmod_add_real_less:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    32
  assumes "Im z \<noteq> 0" "r\<noteq>0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    33
    shows "cmod (z + r) < cmod z + \<bar>r\<bar>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    34
proof (cases z)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    35
  case (Complex x y)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
    36
  then have "0 < y * y"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
    37
    using assms mult_neg_neg by force
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
    38
  with assms have "r * x / \<bar>r\<bar> < sqrt (x*x + y*y)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
    39
    by (simp add: real_less_rsqrt power2_eq_square)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    40
  then show ?thesis using assms Complex
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    41
    apply (simp add: cmod_def)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    42
    apply (rule power2_less_imp_less, auto)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    43
    apply (simp add: power2_eq_square field_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    44
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    45
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    46
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    47
lemma cmod_diff_real_less: "Im z \<noteq> 0 \<Longrightarrow> x\<noteq>0 \<Longrightarrow> cmod (z - x) < cmod z + \<bar>x\<bar>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    48
  using cmod_add_real_less [of z "-x"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    49
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    50
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    51
lemma cmod_square_less_1_plus:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    52
  assumes "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    53
    shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    54
proof (cases "Im z = 0 \<or> Re z = 0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    55
  case True
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
    56
  with assms abs_square_less_1 show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    57
    by (force simp add: Re_power2 Im_power2 cmod_def)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    58
next
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    59
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    60
  with cmod_diff_real_less [of "1 - z\<^sup>2" "1"] show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    61
    by (simp add: norm_power Im_power2)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    62
qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    63
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
    64
subsection\<^marker>\<open>tag unimportant\<close>\<open>The Exponential Function\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
80241
92a66f1df06e Simplification of sin, cos, exp of multiples of pi
paulson <lp15@cam.ac.uk>
parents: 79857
diff changeset
    66
lemma exp_npi_numeral: "exp (\<i> * pi * Num.numeral n)  = (-1) ^ Num.numeral n"
92a66f1df06e Simplification of sin, cos, exp of multiples of pi
paulson <lp15@cam.ac.uk>
parents: 79857
diff changeset
    67
  by (metis exp_of_nat2_mult exp_pi_i' of_nat_numeral)
92a66f1df06e Simplification of sin, cos, exp of multiples of pi
paulson <lp15@cam.ac.uk>
parents: 79857
diff changeset
    68
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    69
lemma norm_exp_i_times [simp]: "norm (exp(\<i> * of_real y)) = 1"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    70
  by simp
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    71
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    72
lemma norm_exp_imaginary: "norm(exp z) = 1 \<Longrightarrow> Re z = 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    73
  by simp
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    74
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    75
lemma field_differentiable_within_exp: "exp field_differentiable (at z within s)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    76
  using DERIV_exp field_differentiable_at_within field_differentiable_def by blast
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
lemma continuous_within_exp:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  fixes z::"'a::{real_normed_field,banach}"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  shows "continuous (at z within s) exp"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77324
diff changeset
    81
  by (simp add: continuous_at_imp_continuous_within)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
    83
lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    84
  by (simp add: field_differentiable_within_exp holomorphic_on_def)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    86
lemma holomorphic_on_exp' [holomorphic_intros]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    87
  "f holomorphic_on s \<Longrightarrow> (\<lambda>x. exp (f x)) holomorphic_on s"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    88
  using holomorphic_on_compose[OF _ holomorphic_on_exp] by (simp add: o_def)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    89
77200
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    90
lemma exp_analytic_on [analytic_intros]:
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    91
  assumes "f analytic_on A"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    92
  shows   "(\<lambda>z. exp (f z)) analytic_on A"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    93
  by (metis analytic_on_holomorphic assms holomorphic_on_exp')
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    94
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    95
lemma
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    96
  assumes "\<And>w. w \<in> A \<Longrightarrow> exp (f w) = w"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    97
  assumes "f holomorphic_on A" "z \<in> A" "open A"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    98
  shows   deriv_complex_logarithm: "deriv f z = 1 / z"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
    99
    and   has_field_derivative_complex_logarithm: "(f has_field_derivative 1 / z) (at z)"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   100
proof -
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   101
  have [simp]: "z \<noteq> 0"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   102
    using assms(1)[of z] assms(3) by auto
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   103
  have deriv [derivative_intros]: "(f has_field_derivative deriv f z) (at z)"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   104
    using assms holomorphic_derivI by blast
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   105
  have "((\<lambda>w. w) has_field_derivative 1) (at z)"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   106
    by (intro derivative_intros)
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   107
  also have "?this \<longleftrightarrow> ((\<lambda>w. exp (f w)) has_field_derivative 1) (at z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   108
    by (smt (verit, best) assms has_field_derivative_transform_within_open)
77200
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   109
  finally have "((\<lambda>w. exp (f w)) has_field_derivative 1) (at z)" .
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   110
  moreover have "((\<lambda>w. exp (f w)) has_field_derivative exp (f z) * deriv f z) (at z)"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   111
    by (rule derivative_eq_intros refl)+
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   112
  ultimately have "exp (f z) * deriv f z = 1"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   113
    using DERIV_unique by blast
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   114
  with assms show "deriv f z = 1 / z"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   115
    by (simp add: field_simps)
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   116
  with deriv show "(f has_field_derivative 1 / z) (at z)"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   117
    by simp
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   118
qed
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
   119
  
67968
a5ad4c015d1c removed dots at the end of (sub)titles
nipkow
parents: 67706
diff changeset
   120
subsection\<open>Euler and de Moivre formulas\<close>
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   121
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69566
diff changeset
   122
text\<open>The sine series times \<^term>\<open>i\<close>\<close>
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
   123
lemma sin_i_eq: "(\<lambda>n. (\<i> * sin_coeff n) * z^n) sums (\<i> * sin z)"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   125
  have "(\<lambda>n. \<i> * sin_coeff n *\<^sub>R z^n) sums (\<i> * sin z)"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
    using sin_converges sums_mult by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
  then show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
    by (simp add: scaleR_conv_of_real field_simps)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   131
theorem exp_Euler: "exp(\<i> * z) = cos(z) + \<i> * sin(z)"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
proof -
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   133
  have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) = (\<lambda>n. (\<i> * z) ^ n /\<^sub>R (fact n))"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   134
    by (force simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   135
  also have "\<dots> sums (exp (\<i> * z))"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
    by (rule exp_converges)
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   137
  finally have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) sums (exp (\<i> * z))" .
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   138
  moreover have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) sums (cos z + \<i> * sin z)"
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
   139
    using sums_add [OF cos_converges [of z] sin_i_eq [of z]]
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
    by (simp add: field_simps scaleR_conv_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  ultimately show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
    using sums_unique2 by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   145
corollary\<^marker>\<open>tag unimportant\<close> exp_minus_Euler: "exp(-(\<i> * z)) = cos(z) - \<i> * sin(z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   146
  using exp_Euler [of "-z"] by simp
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   148
lemma sin_exp_eq: "sin z = (exp(\<i> * z) - exp(-(\<i> * z))) / (2*\<i>)"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   151
lemma sin_exp_eq': "sin z = \<i> * (exp(-(\<i> * z)) - exp(\<i> * z)) / 2"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   154
lemma cos_exp_eq:  "cos z = (exp(\<i> * z) + exp(-(\<i> * z))) / 2"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   157
theorem Euler: "exp(z) = of_real(exp(Re z)) *
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   158
              (of_real(cos(Im z)) + \<i> * of_real(sin(Im z)))"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77324
diff changeset
   159
  by (simp add: Complex_eq cis.code exp_eq_polar)
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   160
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   161
lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   162
  by (simp add: sin_exp_eq field_simps Re_divide Im_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   163
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   164
lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   165
  by (simp add: sin_exp_eq field_simps Im_divide Re_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   166
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   167
lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   168
  by (simp add: cos_exp_eq field_simps Re_divide Re_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   169
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   170
lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   171
  by (simp add: cos_exp_eq field_simps Im_divide Im_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   172
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   173
lemma Re_sin_pos: "0 < Re z \<Longrightarrow> Re z < pi \<Longrightarrow> Re (sin z) > 0"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   174
  by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   175
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   176
lemma Im_sin_nonneg: "Re z = 0 \<Longrightarrow> 0 \<le> Im z \<Longrightarrow> 0 \<le> Im (sin z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   177
  by (simp add: Re_sin Im_sin algebra_simps)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   178
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   179
lemma Im_sin_nonneg2: "Re z = pi \<Longrightarrow> Im z \<le> 0 \<Longrightarrow> 0 \<le> Im (sin z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   180
  by (simp add: Re_sin Im_sin algebra_simps)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   181
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   182
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationships between real and complex trigonometric and hyperbolic functions\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   184
lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
  by (simp add: sin_of_real)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   186
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   187
lemma real_cos_eq [simp]: "Re(cos(of_real x)) = cos x"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
  by (simp add: cos_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   190
lemma DeMoivre: "(cos z + \<i> * sin z) ^ n = cos(n * z) + \<i> * sin(n * z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   191
  by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   192
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   193
lemma exp_cnj: "cnj (exp z) = exp (cnj z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   194
  by (simp add: cis_cnj exp_eq_polar)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   196
lemma cnj_sin: "cnj(sin z) = sin(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   197
  by (simp add: sin_exp_eq exp_cnj field_simps)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   198
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   199
lemma cnj_cos: "cnj(cos z) = cos(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   200
  by (simp add: cos_exp_eq exp_cnj field_simps)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   201
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   202
lemma field_differentiable_at_sin: "sin field_differentiable at z"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   203
  using DERIV_sin field_differentiable_def by blast
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   204
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   205
lemma field_differentiable_within_sin: "sin field_differentiable (at z within S)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   206
  by (simp add: field_differentiable_at_sin field_differentiable_at_within)
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   207
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   208
lemma field_differentiable_at_cos: "cos field_differentiable at z"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   209
  using DERIV_cos field_differentiable_def by blast
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   210
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   211
lemma field_differentiable_within_cos: "cos field_differentiable (at z within S)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   212
  by (simp add: field_differentiable_at_cos field_differentiable_at_within)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   213
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   214
lemma holomorphic_on_sin: "sin holomorphic_on S"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   215
  by (simp add: field_differentiable_within_sin holomorphic_on_def)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   216
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   217
lemma holomorphic_on_cos: "cos holomorphic_on S"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   218
  by (simp add: field_differentiable_within_cos holomorphic_on_def)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   219
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   220
lemma holomorphic_on_sin' [holomorphic_intros]:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   221
  assumes "f holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   222
  shows   "(\<lambda>x. sin (f x)) holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   223
  using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   224
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   225
lemma holomorphic_on_cos' [holomorphic_intros]:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   226
  assumes "f holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   227
  shows   "(\<lambda>x. cos (f x)) holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   228
  using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   229
79857
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   230
lemma analytic_on_sin [analytic_intros]: "f analytic_on A \<Longrightarrow> (\<lambda>w. sin (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   231
  and analytic_on_cos [analytic_intros]: "f analytic_on A \<Longrightarrow> (\<lambda>w. cos (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   232
  and analytic_on_sinh [analytic_intros]: "f analytic_on A \<Longrightarrow> (\<lambda>w. sinh (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   233
  and analytic_on_cosh [analytic_intros]: "f analytic_on A \<Longrightarrow> (\<lambda>w. cosh (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   234
  unfolding sin_exp_eq cos_exp_eq sinh_def cosh_def
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   235
  by (auto intro!: analytic_intros)
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   236
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   237
lemma analytic_on_tan [analytic_intros]:
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   238
        "f analytic_on A \<Longrightarrow> (\<And>z. z \<in> A \<Longrightarrow> cos (f z) \<noteq> 0) \<Longrightarrow> (\<lambda>w. tan (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   239
  and analytic_on_cot [analytic_intros]:
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   240
        "f analytic_on A \<Longrightarrow> (\<And>z. z \<in> A \<Longrightarrow> sin (f z) \<noteq> 0) \<Longrightarrow> (\<lambda>w. cot (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   241
  and analytic_on_tanh [analytic_intros]:
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   242
        "f analytic_on A \<Longrightarrow> (\<And>z. z \<in> A \<Longrightarrow> cosh (f z) \<noteq> 0) \<Longrightarrow> (\<lambda>w. tanh (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
   243
  unfolding tan_def cot_def tanh_def by (auto intro!: analytic_intros)
77221
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
   244
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   245
subsection\<^marker>\<open>tag unimportant\<close>\<open>More on the Polar Representation of Complex Numbers\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   246
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   247
lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   248
  using Complex_eq Euler complex.sel by presburger
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   249
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   250
lemma exp_eq_1: "exp z = 1 \<longleftrightarrow> Re(z) = 0 \<and> (\<exists>n::int. Im(z) = of_int (2 * n) * pi)"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   251
                 (is "?lhs = ?rhs")
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   252
proof
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   253
  assume "exp z = 1"
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   254
  then have "Re z = 0"
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   255
    by (metis exp_eq_one_iff norm_exp_eq_Re norm_one)
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   256
  with \<open>?lhs\<close> show ?rhs
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   257
    by (metis Re_exp cos_one_2pi_int exp_zero mult.commute mult_1 of_int_mult of_int_numeral one_complex.simps(1))
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   258
next
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   259
  assume ?rhs then show ?lhs
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   260
    using Im_exp Re_exp complex_eq_iff
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   261
    by (simp add: cos_one_2pi_int cos_one_sin_zero mult.commute)
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   262
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   263
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   264
lemma exp_eq: "exp w = exp z \<longleftrightarrow> (\<exists>n::int. w = z + (of_int (2 * n) * pi) * \<i>)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   265
                (is "?lhs = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   266
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   267
  have "exp w = exp z \<longleftrightarrow> exp (w-z) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   268
    by (simp add: exp_diff)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   269
  also have "\<dots> \<longleftrightarrow> (Re w = Re z \<and> (\<exists>n::int. Im w - Im z = of_int (2 * n) * pi))"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   270
    by (simp add: exp_eq_1)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   271
  also have "\<dots> \<longleftrightarrow> ?rhs"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   272
    by (auto simp: algebra_simps intro!: complex_eqI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   273
  finally show ?thesis .
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   274
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   275
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   276
lemma exp_complex_eqI: "\<bar>Im w - Im z\<bar> < 2*pi \<Longrightarrow> exp w = exp z \<Longrightarrow> w = z"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   277
  by (auto simp: exp_eq abs_mult)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   278
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   279
lemma exp_integer_2pi:
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   280
  assumes "n \<in> \<int>"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   281
  shows "exp((2 * n * pi) * \<i>) = 1"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   282
  by (metis assms cis_conv_exp cis_multiple_2pi mult.assoc mult.commute)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   283
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   284
lemma exp_plus_2pin [simp]: "exp (z + \<i> * (of_int n * (of_real pi * 2))) = exp z"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   285
  by (simp add: exp_eq)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   286
66466
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   287
lemma exp_integer_2pi_plus1:
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   288
  assumes "n \<in> \<int>"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   289
  shows "exp(((2 * n + 1) * pi) * \<i>) = - 1"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   290
  using exp_integer_2pi [OF assms]
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   291
  by (metis cis_conv_exp cis_mult cis_pi distrib_left mult.commute mult.right_neutral)
66466
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   292
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   293
lemma inj_on_exp_pi:
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   294
  fixes z::complex shows "inj_on exp (ball z pi)"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   295
proof (clarsimp simp: inj_on_def exp_eq)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   296
  fix y n
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   297
  assume "dist z (y + 2 * of_int n * of_real pi * \<i>) < pi"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   298
         "dist z y < pi"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   299
  then have "dist y (y + 2 * of_int n * of_real pi * \<i>) < pi+pi"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   300
    using dist_commute_lessI dist_triangle_less_add by blast
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   301
  then have "norm (2 * of_int n * of_real pi * \<i>) < 2*pi"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   302
    by (simp add: dist_norm)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   303
  then show "n = 0"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   304
    by (auto simp: norm_mult)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   305
qed
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   306
68585
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   307
lemma cmod_add_squared:
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   308
  fixes r1 r2::real
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   309
  shows "(cmod (r1 * exp (\<i> * \<theta>1) + r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\<theta>1 - \<theta>2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs")
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   310
proof -
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   311
  have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   312
    by (rule complex_norm_square)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   313
  also have "\<dots> = (?z1 * cnj ?z1 + ?z2 * cnj ?z2) + (?z1 * cnj ?z2 + cnj ?z1 * ?z2)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   314
    by (simp add: algebra_simps)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   315
  also have "\<dots> = (norm ?z1)\<^sup>2 + (norm ?z2)\<^sup>2 + 2 * Re (?z1 * cnj ?z2)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   316
    unfolding complex_norm_square [symmetric] cnj_add_mult_eq_Re by simp
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   317
  also have "\<dots> = ?rhs"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   318
    by (simp add: norm_mult) (simp add: exp_Euler complex_is_Real_iff [THEN iffD1] cos_diff algebra_simps)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   319
  finally show ?thesis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   320
    using of_real_eq_iff by blast
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   321
qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   322
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   323
lemma cmod_diff_squared:
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   324
  fixes r1 r2::real
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   325
  shows "(cmod (r1 * exp (\<i> * \<theta>1) - r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\<theta>1 - \<theta>2)" 
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   326
  using cmod_add_squared [of r1 _ "-r2"] by simp
68585
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   327
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   328
lemma polar_convergence:
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   329
  fixes R::real
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   330
  assumes "\<And>j. r j > 0" "R > 0"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   331
  shows "((\<lambda>j. r j * exp (\<i> * \<theta> j)) \<longlonglongrightarrow> (R * exp (\<i> * \<Theta>))) \<longleftrightarrow>
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   332
         (r \<longlonglongrightarrow> R) \<and> (\<exists>k. (\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>)"    (is "(?z \<longlonglongrightarrow> ?Z) = ?rhs")
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   333
proof
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   334
  assume L: "?z \<longlonglongrightarrow> ?Z"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   335
  have rR: "r \<longlonglongrightarrow> R"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   336
    using tendsto_norm [OF L] assms by (auto simp: norm_mult abs_of_pos)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   337
  moreover obtain k where "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   338
  proof -
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   339
    have "cos (\<theta> j - \<Theta>) = ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)" for j
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   340
      using assms by (auto simp: cmod_diff_squared less_le)
68585
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   341
    moreover have "(\<lambda>j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \<longlonglongrightarrow> ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   342
      by (intro L rR tendsto_intros) (use \<open>R > 0\<close> in force)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   343
    moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
   344
      using \<open>R > 0\<close> by (simp add: power2_eq_square field_split_simps)
68585
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   345
    ultimately have "(\<lambda>j. cos (\<theta> j - \<Theta>)) \<longlonglongrightarrow> 1"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   346
      by auto
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   347
    then show ?thesis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   348
      using that cos_diff_limit_1 by blast
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   349
  qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   350
  ultimately show ?rhs
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   351
    by metis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   352
next
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   353
  assume R: ?rhs
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   354
  show "?z \<longlonglongrightarrow> ?Z"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   355
  proof (rule tendsto_mult)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   356
    show "(\<lambda>x. complex_of_real (r x)) \<longlonglongrightarrow> of_real R"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   357
      using R by (auto simp: tendsto_of_real_iff)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   358
    obtain k where "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   359
      using R by metis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   360
    then have "(\<lambda>j. complex_of_real (\<theta> j - of_int (k j) * (2 * pi))) \<longlonglongrightarrow> of_real \<Theta>"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   361
      using tendsto_of_real_iff by force
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   362
    then have "(\<lambda>j.  exp (\<i> * of_real (\<theta> j - of_int (k j) * (2 * pi)))) \<longlonglongrightarrow> exp (\<i> * \<Theta>)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   363
      using tendsto_mult [OF tendsto_const] isCont_exp isCont_tendsto_compose by blast
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   364
    moreover have "exp (\<i> * of_real (\<theta> j - of_int (k j) * (2 * pi))) = exp (\<i> * \<theta> j)" for j
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   365
      unfolding exp_eq
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   366
      by (rule_tac x="- k j" in exI) (auto simp: algebra_simps)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   367
    ultimately show "(\<lambda>j. exp (\<i> * \<theta> j)) \<longlonglongrightarrow> exp (\<i> * \<Theta>)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   368
      by auto
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   369
  qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   370
qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   371
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   372
lemma exp_i_ne_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   373
  assumes "0 < x" "x < 2*pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   374
  shows "exp(\<i> * of_real x) \<noteq> 1"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   375
  by (smt (verit) Im_i_times Re_complex_of_real assms exp_complex_eqI exp_zero zero_complex.sel(2))
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   376
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   377
lemma sin_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   378
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   379
  shows "sin z = 0 \<longleftrightarrow> (\<exists>n::int. z = of_real(n * pi))"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   380
  by (simp add: sin_exp_eq exp_eq)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   381
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   382
lemma cos_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   383
  fixes z::complex
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   384
  shows "cos z = 0 \<longleftrightarrow> (\<exists>n::int. z = complex_of_real(n * pi) + of_real pi/2)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   385
  using sin_eq_0 [of "z - of_real pi/2"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   386
  by (simp add: sin_diff algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   387
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   388
lemma cos_eq_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   389
  fixes z::complex
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   390
  shows "cos z = 1 \<longleftrightarrow> (\<exists>n::int. z = complex_of_real(2 * n * pi))"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   391
  by (metis Re_complex_of_real cos_of_real cos_one_2pi_int cos_one_sin_zero mult.commute of_real_1 sin_eq_0)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   392
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   393
lemma csin_eq_1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   394
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   395
  shows "sin z = 1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi) + of_real pi/2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   396
  using cos_eq_1 [of "z - of_real pi/2"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   397
  by (simp add: cos_diff algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   398
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   399
lemma csin_eq_minus1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   400
  fixes z::complex
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   401
  shows "sin z = -1 \<longleftrightarrow> (\<exists>n::int. z = complex_of_real(2 * n * pi) + 3/2*pi)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   402
        (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   403
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   404
  have "sin z = -1 \<longleftrightarrow> sin (-z) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   405
    by (simp add: equation_minus_iff)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   406
  also have "\<dots> \<longleftrightarrow> (\<exists>n::int. z = - of_real(2 * n * pi) - of_real pi/2)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   407
    by (metis (mono_tags, lifting) add_uminus_conv_diff csin_eq_1 equation_minus_iff minus_add_distrib)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   408
  also have "\<dots> = ?rhs"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   409
    apply safe
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   410
    apply (rule_tac [2] x="-(x+1)" in exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   411
    apply (rule_tac x="-(x+1)" in exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   412
    apply (simp_all add: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   413
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   414
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   415
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   416
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   417
lemma ccos_eq_minus1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   418
  fixes z::complex
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   419
  shows "cos z = -1 \<longleftrightarrow> (\<exists>n::int. z = complex_of_real(2 * n * pi) + pi)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   420
  using csin_eq_1 [of "z - of_real pi/2"]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   421
  by (simp add: sin_diff algebra_simps equation_minus_iff)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   422
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   423
lemma sin_eq_1: "sin x = 1 \<longleftrightarrow> (\<exists>n::int. x = (2 * n + 1 / 2) * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   424
                (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   425
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   426
  have "sin x = 1 \<longleftrightarrow> sin (complex_of_real x) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   427
    by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   428
  also have "\<dots> \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + of_real pi/2)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   429
    by (metis csin_eq_1 Re_complex_of_real id_apply of_real_add of_real_divide of_real_eq_id of_real_numeral)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   430
  also have "\<dots> = ?rhs"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   431
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   432
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   433
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   434
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   435
lemma sin_eq_minus1: "sin x = -1 \<longleftrightarrow> (\<exists>n::int. x = (2*n + 3/2) * pi)"  (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   436
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   437
  have "sin x = -1 \<longleftrightarrow> sin (complex_of_real x) = -1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   438
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   439
  also have "\<dots> \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + 3/2*pi)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   440
    by (metis Re_complex_of_real csin_eq_minus1 id_apply of_real_add of_real_eq_id)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   441
  also have "\<dots> = ?rhs"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   442
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   443
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   444
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   445
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   446
lemma cos_eq_minus1: "cos x = -1 \<longleftrightarrow> (\<exists>n::int. x = (2*n + 1) * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   447
                      (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   448
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   449
  have "cos x = -1 \<longleftrightarrow> cos (complex_of_real x) = -1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   450
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   451
  also have "\<dots> \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + pi)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   452
    by (metis ccos_eq_minus1 id_apply of_real_add of_real_eq_id of_real_eq_iff)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   453
  also have "\<dots> = ?rhs"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   454
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   455
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   456
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   457
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   458
lemma cos_gt_neg1:
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   459
  assumes "(t::real) \<in> {-pi<..<pi}"
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   460
  shows   "cos t > -1"
77103
11d844d21f5c Shortened a messy proof
paulson <lp15@cam.ac.uk>
parents: 77089
diff changeset
   461
  using assms
11d844d21f5c Shortened a messy proof
paulson <lp15@cam.ac.uk>
parents: 77089
diff changeset
   462
  by simp (metis cos_minus cos_monotone_0_pi cos_monotone_minus_pi_0 cos_pi linorder_le_cases)
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   463
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
   464
lemma dist_exp_i_1: "norm(exp(\<i> * of_real t) - 1) = 2 * \<bar>sin(t / 2)\<bar>"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   465
proof -
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   466
  have "sqrt (2 - cos t * 2) = 2 * \<bar>sin (t / 2)\<bar>"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   467
    using cos_double_sin [of "t/2"] by (simp add: real_sqrt_mult)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   468
  then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   469
    by (simp add: exp_Euler cmod_def power2_diff sin_of_real cos_of_real algebra_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   470
qed
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   471
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   472
lemma sin_cx_2pi [simp]: "\<lbrakk>z = of_int m; even m\<rbrakk> \<Longrightarrow> sin (z * complex_of_real pi) = 0"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   473
  by (simp add: sin_eq_0)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   474
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   475
lemma cos_cx_2pi [simp]: "\<lbrakk>z = of_int m; even m\<rbrakk> \<Longrightarrow> cos (z * complex_of_real pi) = 1"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   476
  using cos_eq_1 by auto
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   477
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   478
lemma complex_sin_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   479
  fixes w :: complex
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   480
  shows "sin w = sin z \<longleftrightarrow> (\<exists>n \<in> \<int>. w = z + of_real(2*n*pi) \<or> w = -z + of_real((2*n + 1)*pi))"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   481
        (is "?lhs = ?rhs")
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   482
proof
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   483
  assume ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   484
  then consider "sin ((w - z) / 2) = 0" | "cos ((w + z) / 2) = 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   485
    by (metis divide_eq_0_iff nonzero_eq_divide_eq right_minus_eq sin_diff_sin zero_neq_numeral)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   486
  then show ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   487
  proof cases
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   488
    case 1
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   489
    then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   490
      by (simp add: sin_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   491
  next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   492
    case 2
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   493
    then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   494
      by (simp add: cos_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   495
  qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   496
next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   497
  assume ?rhs
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   498
  then consider n::int where "w = z + of_real (2 * of_int n * pi)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   499
              | n::int where  " w = -z + of_real ((2 * of_int n + 1) * pi)"
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   500
    using Ints_cases by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   501
  then show ?lhs
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   502
  proof cases
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   503
    case 1
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   504
    then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   505
      using Periodic_Fun.sin.plus_of_int [of z n]
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   506
      by (auto simp: algebra_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   507
  next
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   508
    case 2
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   509
    then show ?thesis
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   510
      using Periodic_Fun.sin.plus_of_int [of "-z" "n"]
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   511
      apply (simp add: algebra_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   512
      by (metis add.commute add.inverse_inverse add_diff_cancel_left diff_add_cancel sin_plus_pi)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   513
  qed
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   514
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   515
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   516
lemma complex_cos_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   517
  fixes w :: complex
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   518
  shows "cos w = cos z \<longleftrightarrow> (\<exists>n \<in> \<int>. w = z + of_real(2*n*pi) \<or> w = -z + of_real(2*n*pi))"
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   519
        (is "?lhs = ?rhs")
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   520
proof 
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   521
  assume ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   522
  then consider "sin ((w + z) / 2) = 0" | "sin ((z - w) / 2) = 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   523
    by (metis mult_eq_0_iff cos_diff_cos right_minus_eq zero_neq_numeral)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   524
  then show ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   525
  proof cases
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   526
    case 1
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   527
    then obtain n where "w + z = of_int n * (complex_of_real pi * 2)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   528
      by (auto simp: sin_eq_0 algebra_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   529
    then have "w = -z + of_real(2 * of_int n * pi)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   530
      by (auto simp: algebra_simps)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   531
    then show ?thesis
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   532
      using Ints_of_int by blast
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   533
  next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   534
    case 2
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   535
    then obtain n where "z = w + of_int n * (complex_of_real pi * 2)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   536
      by (auto simp: sin_eq_0 algebra_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   537
    then have "w = z + complex_of_real (2 * of_int(-n) * pi)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   538
      by (auto simp: algebra_simps)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   539
    then show ?thesis
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   540
      using Ints_of_int by blast
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   541
  qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   542
next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   543
  assume ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   544
  then obtain n::int where w: "w = z + of_real (2* of_int n*pi) \<or>
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   545
                               w = -z + of_real(2*n*pi)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   546
    using Ints_cases  by (metis of_int_mult of_int_numeral)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   547
  then show ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   548
    using Periodic_Fun.cos.plus_of_int [of z n]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   549
    apply (simp add: algebra_simps)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   550
    by (metis cos.plus_of_int cos_minus minus_add_cancel mult.commute)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   551
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   552
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   553
lemma sin_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   554
   "sin x = sin y \<longleftrightarrow> (\<exists>n \<in> \<int>. x = y + 2*n*pi \<or> x = -y + (2*n + 1)*pi)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   555
  using complex_sin_eq [of x y]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   556
  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)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   557
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   558
lemma cos_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   559
   "cos x = cos y \<longleftrightarrow> (\<exists>n \<in> \<int>. x = y + 2*n*pi \<or> x = -y + 2*n*pi)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77324
diff changeset
   560
  using complex_cos_eq [of x y] unfolding cos_of_real 
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77324
diff changeset
   561
  by (metis Re_complex_of_real of_real_add of_real_minus)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   562
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   563
lemma sinh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   564
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   565
  shows "(exp z - inverse (exp z)) / 2 = -\<i> * sin(\<i> * z)"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
   566
  by (simp add: sin_exp_eq field_split_simps exp_minus)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   567
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
   568
lemma sin_i_times:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   569
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   570
  shows "sin(\<i> * z) = \<i> * ((exp z - inverse (exp z)) / 2)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   571
  using sinh_complex by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   572
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   573
lemma sinh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   574
  fixes x :: real
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   575
  shows "of_real((exp x - inverse (exp x)) / 2) = -\<i> * sin(\<i> * of_real x)"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   576
  by (simp add: exp_of_real sin_i_times)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   577
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   578
lemma cosh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   579
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   580
  shows "(exp z + inverse (exp z)) / 2 = cos(\<i> * z)"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
   581
  by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   582
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   583
lemma cosh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   584
  fixes x :: real
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   585
  shows "of_real((exp x + inverse (exp x)) / 2) = cos(\<i> * of_real x)"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
   586
  by (simp add: cos_exp_eq field_split_simps exp_minus exp_of_real)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   587
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
   588
lemmas cos_i_times = cosh_complex [symmetric]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   589
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   590
lemma norm_cos_squared:
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   591
  "norm(cos z) ^ 2 = cos(Re z) ^ 2 + (exp(Im z) - inverse(exp(Im z))) ^ 2 / 4"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   592
proof (cases z)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   593
  case (Complex x1 x2)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   594
  then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   595
    apply (simp only: cos_add cmod_power2 cos_of_real sin_of_real Complex_eq)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   596
    apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   597
    apply (simp only: left_diff_distrib [symmetric] power_mult_distrib sin_squared_eq)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   598
    apply (simp add: power2_eq_square field_split_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   599
    done
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   600
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   601
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   602
lemma norm_sin_squared:
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   603
  "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   604
  using cos_double_sin [of "Re z"]
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   605
  apply (simp add: sin_cos_eq norm_cos_squared exp_minus mult.commute [of _ 2] exp_double)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   606
  apply (simp add: algebra_simps power2_eq_square)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   607
  done
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   608
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   609
lemma exp_uminus_Im: "exp (- Im z) \<le> exp (cmod z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   610
  using abs_Im_le_cmod linear order_trans by fastforce
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   611
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   612
lemma norm_cos_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   613
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   614
  shows "norm(cos z) \<le> exp(norm z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   615
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   616
  have "Im z \<le> cmod z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   617
    using abs_Im_le_cmod abs_le_D1 by auto
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   618
  then have "exp (- Im z) + exp (Im z) \<le> exp (cmod z) * 2"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   619
    by (metis exp_uminus_Im add_mono exp_le_cancel_iff mult_2_right)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   620
  then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   621
    by (force simp add: cos_exp_eq norm_divide intro: order_trans [OF norm_triangle_ineq])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   622
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   623
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   624
lemma norm_cos_plus1_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   625
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   626
  shows "norm(1 + cos z) \<le> 2 * exp(norm z)"
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77324
diff changeset
   627
  by (metis mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono one_le_exp_iff)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   628
67578
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   629
lemma sinh_conv_sin: "sinh z = -\<i> * sin (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   630
  by (simp add: sinh_field_def sin_i_times exp_minus)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   631
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   632
lemma cosh_conv_cos: "cosh z = cos (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   633
  by (simp add: cosh_field_def cos_i_times exp_minus)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   634
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   635
lemma tanh_conv_tan: "tanh z = -\<i> * tan (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   636
  by (simp add: tanh_def sinh_conv_sin cosh_conv_cos tan_def)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   637
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   638
lemma sin_conv_sinh: "sin z = -\<i> * sinh (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   639
  by (simp add: sinh_conv_sin)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   640
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   641
lemma cos_conv_cosh: "cos z = cosh (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   642
  by (simp add: cosh_conv_cos)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   643
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   644
lemma tan_conv_tanh: "tan z = -\<i> * tanh (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   645
  by (simp add: tan_def sin_conv_sinh cos_conv_cosh tanh_def)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   646
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   647
lemma sinh_complex_eq_iff:
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   648
  "sinh (z :: complex) = sinh w \<longleftrightarrow>
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   649
      (\<exists>n\<in>\<int>. z = w - 2 * \<i> * of_real n * of_real pi \<or>
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   650
              z = -(2 * complex_of_real n + 1) * \<i> * complex_of_real pi - w)" (is "_ = ?rhs")
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   651
proof -
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   652
  have "sinh z = sinh w \<longleftrightarrow> sin (\<i> * z) = sin (\<i> * w)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   653
    by (simp add: sinh_conv_sin)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   654
  also have "\<dots> \<longleftrightarrow> ?rhs"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   655
    by (subst complex_sin_eq) (force simp: field_simps complex_eq_iff)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   656
  finally show ?thesis .
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   657
qed
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   658
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   659
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   660
subsection\<^marker>\<open>tag unimportant\<close>\<open>Taylor series for complex exponential, sine and cosine\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   661
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   662
declare power_Suc [simp del]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   663
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   664
lemma Taylor_exp_field:
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   665
  fixes z::"'a::{banach,real_normed_field}"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   666
  shows "norm (exp z - (\<Sum>i\<le>n. z ^ i / fact i)) \<le> exp (norm z) * (norm z ^ Suc n) / fact n"
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
   667
proof (rule field_Taylor[of _ n "\<lambda>k. exp" "exp (norm z)" 0 z, simplified])
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   668
  show "convex (closed_segment 0 z)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   669
    by (rule convex_closed_segment [of 0 z])
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   670
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   671
  fix k x
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   672
  assume "x \<in> closed_segment 0 z" "k \<le> n"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   673
  show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   674
    using DERIV_exp DERIV_subset by blast
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   675
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   676
  fix x
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   677
  assume x: "x \<in> closed_segment 0 z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   678
  have "norm (exp x) \<le> exp (norm x)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   679
    by (rule norm_exp)
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   680
  also have "norm x \<le> norm z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   681
    using x by (auto simp: closed_segment_def intro!: mult_left_le_one_le)
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   682
  finally show "norm (exp x) \<le> exp (norm z)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   683
    by simp
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   684
qed auto
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   685
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   686
text \<open>For complex @{term z}, a tighter bound than in the previous result\<close>
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   687
lemma Taylor_exp:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   688
  "norm(exp z - (\<Sum>k\<le>n. z ^ k / (fact k))) \<le> exp\<bar>Re z\<bar> * (norm z) ^ (Suc n) / (fact n)"
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
   689
proof (rule complex_Taylor [of _ n "\<lambda>k. exp" "exp\<bar>Re z\<bar>" 0 z, simplified])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   690
  show "convex (closed_segment 0 z)"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   691
    by (rule convex_closed_segment [of 0 z])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   692
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   693
  fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   694
  assume "x \<in> closed_segment 0 z" "k \<le> n"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   695
  show "(exp has_field_derivative exp x) (at x within closed_segment 0 z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   696
    using DERIV_exp DERIV_subset by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   697
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   698
  fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   699
  assume "x \<in> closed_segment 0 z"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   700
  then obtain u where u: "x = complex_of_real u * z" "0 \<le> u" "u \<le> 1"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   701
    by (auto simp: closed_segment_def scaleR_conv_of_real)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   702
  then have "u * Re z \<le> \<bar>Re z\<bar>"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   703
    by (metis abs_ge_self abs_ge_zero mult.commute mult.right_neutral mult_mono)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   704
  then show "Re x \<le> \<bar>Re z\<bar>"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   705
    by (simp add: u)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   706
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   707
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   708
lemma
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   709
  assumes "0 \<le> u" "u \<le> 1"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   710
  shows cmod_sin_le_exp: "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   711
    and cmod_cos_le_exp: "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   712
proof -
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   713
  have mono: "\<And>u w z::real. w \<le> u \<Longrightarrow> z \<le> u \<Longrightarrow> (w + z)/2 \<le> u"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   714
    by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   715
  have *: "(cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2 \<le> exp \<bar>Im z\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   716
  proof (rule mono)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   717
    show "cmod (exp (\<i> * (u * z))) \<le> exp \<bar>Im z\<bar>"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   718
      using assms
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   719
      by (auto simp: abs_if mult_left_le_one_le not_less intro: order_trans [of _ 0])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   720
    show "cmod (exp (- (\<i> * (u * z)))) \<le> exp \<bar>Im z\<bar>"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   721
      using assms
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   722
      by (auto simp: abs_if mult_left_le_one_le mult_nonneg_nonpos intro: order_trans [of _ 0])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   723
  qed
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   724
  have "cmod (sin (u *\<^sub>R z)) = cmod (exp (\<i> * (u * z)) - exp (- (\<i> * (u * z)))) / 2"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   725
    by (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   726
  also have "\<dots> \<le> (cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   727
    by (intro divide_right_mono norm_triangle_ineq4) simp
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   728
  also have "\<dots> \<le> exp \<bar>Im z\<bar>"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   729
    by (rule *)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   730
  finally show "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" .
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   731
  have "cmod (cos (u *\<^sub>R z)) = cmod (exp (\<i> * (u * z)) + exp (- (\<i> * (u * z)))) / 2"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   732
    by (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   733
  also have "\<dots> \<le> (cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   734
    by (intro divide_right_mono norm_triangle_ineq) simp
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   735
  also have "\<dots> \<le> exp \<bar>Im z\<bar>"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   736
    by (rule *)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   737
  finally show "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" .
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   738
qed
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   739
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   740
lemma Taylor_sin:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   741
  "norm(sin z - (\<Sum>k\<le>n. complex_of_real (sin_coeff k) * z ^ k))
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   742
   \<le> exp\<bar>Im z\<bar> * (norm z) ^ (Suc n) / (fact n)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   743
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   744
  have mono: "\<And>u w z::real. w \<le> u \<Longrightarrow> z \<le> u \<Longrightarrow> w + z \<le> u*2"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   745
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   746
  have *: "cmod (sin z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   747
                 (\<Sum>i\<le>n. (-1) ^ (i div 2) * (if even i then sin 0 else cos 0) * z ^ i / (fact i)))
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   748
           \<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact n)"
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
   749
  proof (rule complex_Taylor [of "closed_segment 0 z" n
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
   750
                                 "\<lambda>k x. (-1)^(k div 2) * (if even k then sin x else cos x)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   751
                                 "exp\<bar>Im z\<bar>" 0 z, simplified])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   752
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   753
    show "((\<lambda>x. (- 1) ^ (k div 2) * (if even k then sin x else cos x)) has_field_derivative
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   754
            (- 1) ^ (Suc k div 2) * (if odd k then sin x else cos x))
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   755
            (at x within closed_segment 0 z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   756
      by (cases "even k") (intro derivative_eq_intros | simp add: power_Suc)+
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   757
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   758
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   759
    assume "x \<in> closed_segment 0 z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   760
    then show "cmod ((- 1) ^ (Suc n div 2) * (if odd n then sin x else cos x)) \<le> exp \<bar>Im z\<bar>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   761
      by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   762
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   763
  have **: "\<And>k. complex_of_real (sin_coeff k) * z ^ k
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   764
            = (-1)^(k div 2) * (if even k then sin 0 else cos 0) * z^k / of_nat (fact k)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   765
    by (auto simp: sin_coeff_def elim!: oddE)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   766
  show ?thesis
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   767
    by (simp add: ** order_trans [OF _ *])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   768
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   769
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   770
lemma Taylor_cos:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   771
  "norm(cos z - (\<Sum>k\<le>n. complex_of_real (cos_coeff k) * z ^ k))
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   772
   \<le> exp\<bar>Im z\<bar> * (norm z) ^ Suc n / (fact n)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   773
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   774
  have mono: "\<And>u w z::real. w \<le> u \<Longrightarrow> z \<le> u \<Longrightarrow> w + z \<le> u*2"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   775
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   776
  have *: "cmod (cos z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   777
                 (\<Sum>i\<le>n. (-1) ^ (Suc i div 2) * (if even i then cos 0 else sin 0) * z ^ i / (fact i)))
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   778
           \<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact n)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   779
  proof (rule complex_Taylor [of "closed_segment 0 z" n 
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   780
                                 "\<lambda>k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" 
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   781
                                 "exp\<bar>Im z\<bar>" 0 z, simplified])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   782
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   783
    assume "x \<in> closed_segment 0 z" "k \<le> n"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   784
    show "((\<lambda>x. (- 1) ^ (Suc k div 2) * (if even k then cos x else sin x)) has_field_derivative
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   785
            (- 1) ^ Suc (k div 2) * (if odd k then cos x else sin x))
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   786
             (at x within closed_segment 0 z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   787
      by (cases "even k") (intro derivative_eq_intros | simp add: power_Suc)+
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   788
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   789
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   790
    assume "x \<in> closed_segment 0 z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   791
    then show "cmod ((- 1) ^ Suc (n div 2) * (if odd n then cos x else sin x)) \<le> exp \<bar>Im z\<bar>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   792
      by (auto simp: closed_segment_def norm_mult norm_power cmod_sin_le_exp cmod_cos_le_exp)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   793
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   794
  have **: "\<And>k. complex_of_real (cos_coeff k) * z ^ k
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   795
            = (-1)^(Suc k div 2) * (if even k then cos 0 else sin 0) * z^k / of_nat (fact k)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   796
    by (auto simp: cos_coeff_def elim!: evenE)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   797
  show ?thesis
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   798
    by (simp add: ** order_trans [OF _ *])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   799
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   800
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
   801
declare power_Suc [simp]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   802
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   803
text\<open>32-bit Approximation to e\<close>
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   804
lemma e_approx_32: "\<bar>exp(1) - 5837465777 / 2147483648\<bar> \<le> (inverse(2 ^ 32)::real)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   805
  using Taylor_exp [of 1 14] exp_le
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   806
  apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
66611
c375b64a6c24 adapted to better linear arith
nipkow
parents: 66480
diff changeset
   807
  apply (simp only: pos_le_divide_eq [symmetric])
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   808
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   809
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   810
lemma e_less_272: "exp 1 < (272/100::real)"
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   811
  using e_approx_32
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
   812
  by (simp add: abs_if split: if_split_asm)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   813
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   814
lemma ln_272_gt_1: "ln (272/100) > (1::real)"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   815
  by (metis e_less_272 exp_less_cancel_iff exp_ln_iff less_trans ln_exp)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   816
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   817
text\<open>Apparently redundant. But many arguments involve integers.\<close>
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   818
lemma ln3_gt_1: "ln 3 > (1::real)"
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   819
  by (simp add: less_trans [OF ln_272_gt_1])
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   820
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   821
subsection\<open>The argument of a complex number (HOL Light version)\<close>
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   822
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   823
definition\<^marker>\<open>tag important\<close> is_Arg :: "[complex,real] \<Rightarrow> bool"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   824
  where "is_Arg z r \<equiv> z = of_real(norm z) * exp(\<i> * of_real r)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   825
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   826
definition\<^marker>\<open>tag important\<close> Arg2pi :: "complex \<Rightarrow> real"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   827
  where "Arg2pi z \<equiv> if z = 0 then 0 else THE t. 0 \<le> t \<and> t < 2*pi \<and> is_Arg z t"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   828
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   829
lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \<longleftrightarrow> is_Arg z r"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   830
  by (simp add: algebra_simps is_Arg_def)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   831
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   832
lemma is_Arg_eqI:
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   833
  assumes "is_Arg z r" and "is_Arg z s" and "abs (r-s) < 2*pi" and "z \<noteq> 0"
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   834
  shows "r = s"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   835
  using assms unfolding is_Arg_def
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   836
  by (metis Im_i_times Re_complex_of_real exp_complex_eqI mult_cancel_left mult_eq_0_iff)
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   837
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   838
text\<open>This function returns the angle of a complex number from its representation in polar coordinates.
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69566
diff changeset
   839
Due to periodicity, its range is arbitrary. \<^term>\<open>Arg2pi\<close> follows HOL Light in adopting the interval \<open>[0,2\<pi>)\<close>.
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   840
But we have the same periodicity issue with logarithms, and it is usual to adopt the same interval
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69529
diff changeset
   841
for the complex logarithm and argument functions. Further on down, we shall define both functions for the interval \<open>(-\<pi>,\<pi>]\<close>.
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   842
The present version is provided for compatibility.\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   843
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   844
lemma Arg2pi_0 [simp]: "Arg2pi(0) = 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   845
  by (simp add: Arg2pi_def)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   846
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   847
lemma Arg2pi_unique_lemma:
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   848
  assumes "is_Arg z t"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   849
      and "is_Arg z t'"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   850
      and "0 \<le> t"  "t < 2*pi"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   851
      and "0 \<le> t'" "t' < 2*pi"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   852
      and "z \<noteq> 0"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   853
  shows "t' = t"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   854
  using is_Arg_eqI assms by force
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   855
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   856
lemma Arg2pi: "0 \<le> Arg2pi z \<and> Arg2pi z < 2*pi \<and> is_Arg z (Arg2pi z)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   857
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   858
  case True then show ?thesis
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   859
    by (simp add: Arg2pi_def is_Arg_def)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   860
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   861
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   862
  obtain t where t: "0 \<le> t" "t < 2*pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   863
             and ReIm: "Re z / cmod z = cos t" "Im z / cmod z = sin t"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   864
    using sincos_total_2pi [OF complex_unit_circle [OF False]]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   865
    by blast
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   866
  then have z: "is_Arg z t"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   867
    unfolding is_Arg_def
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   868
    using t False ReIm
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   869
    by (intro complex_eqI) (auto simp: exp_Euler sin_of_real cos_of_real field_split_simps)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   870
  show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   871
    apply (simp add: Arg2pi_def False)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   872
    apply (rule theI [where a=t])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   873
    using t z False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   874
    apply (auto intro: Arg2pi_unique_lemma)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   875
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   876
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   877
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   878
corollary\<^marker>\<open>tag unimportant\<close>
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   879
  shows Arg2pi_ge_0: "0 \<le> Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   880
    and Arg2pi_lt_2pi: "Arg2pi z < 2*pi"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   881
    and Arg2pi_eq: "z = of_real(norm z) * exp(\<i> * of_real(Arg2pi z))"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   882
  using Arg2pi is_Arg_def by auto
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   883
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   884
lemma complex_norm_eq_1_exp: "norm z = 1 \<longleftrightarrow> exp(\<i> * of_real (Arg2pi z)) = z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   885
  by (metis Arg2pi_eq cis_conv_exp mult.left_neutral norm_cis of_real_1)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   886
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   887
lemma Arg2pi_unique: "\<lbrakk>of_real r * exp(\<i> * of_real a) = z; 0 < r; 0 \<le> a; a < 2*pi\<rbrakk> \<Longrightarrow> Arg2pi z = a"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   888
  by (rule Arg2pi_unique_lemma [unfolded is_Arg_def, OF _ Arg2pi_eq]) (use Arg2pi [of z] in \<open>auto simp: norm_mult\<close>)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   889
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   890
lemma cos_Arg2pi: "cmod z * cos (Arg2pi z) = Re z" and sin_Arg2pi: "cmod z * sin (Arg2pi z) = Im z"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   891
  using Arg2pi_eq [of z] cis_conv_exp Re_rcis Im_rcis unfolding rcis_def by metis+
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   892
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   893
lemma Arg2pi_minus:
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   894
  assumes "z \<noteq> 0" shows "Arg2pi (-z) = (if Arg2pi z < pi then Arg2pi z + pi else Arg2pi z - pi)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   895
  apply (rule Arg2pi_unique [of "norm z", OF complex_eqI])
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   896
  using cos_Arg2pi sin_Arg2pi Arg2pi_ge_0 Arg2pi_lt_2pi [of z] assms
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
   897
  by (auto simp: Re_exp Im_exp)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   898
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   899
lemma Arg2pi_times_of_real [simp]:
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   900
  assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   901
  by (metis (no_types, lifting) Arg2pi Arg2pi_eq Arg2pi_unique assms mult.assoc 
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   902
      mult_eq_0_iff mult_pos_pos of_real_mult zero_less_norm_iff)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   903
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   904
lemma Arg2pi_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg2pi (z * of_real r) = Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   905
  by (metis Arg2pi_times_of_real mult.commute)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   906
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   907
lemma Arg2pi_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg2pi (z / of_real r) = Arg2pi z"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   908
  by (metis Arg2pi_times_of_real2 less_irrefl nonzero_eq_divide_eq of_real_eq_0_iff)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   909
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   910
lemma Arg2pi_le_pi: "Arg2pi z \<le> pi \<longleftrightarrow> 0 \<le> Im z"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   911
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   912
  case False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   913
  have "0 \<le> Im z \<longleftrightarrow> 0 \<le> Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg2pi z)))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   914
    by (metis Arg2pi_eq)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   915
  also have "\<dots> = (0 \<le> Im (exp (\<i> * complex_of_real (Arg2pi z))))"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   916
    using False  by (simp add: zero_le_mult_iff)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   917
  also have "\<dots> \<longleftrightarrow> Arg2pi z \<le> pi"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   918
    by (simp add: Im_exp) (metis Arg2pi_ge_0 Arg2pi_lt_2pi sin_lt_zero sin_ge_zero not_le)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   919
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   920
    by blast
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   921
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   922
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   923
lemma Arg2pi_lt_pi: "0 < Arg2pi z \<and> Arg2pi z < pi \<longleftrightarrow> 0 < Im z"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   924
  using Arg2pi_le_pi [of z]
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   925
  by (smt (verit, del_insts) Arg2pi_0 Arg2pi_le_pi Arg2pi_minus uminus_complex.simps(2) zero_complex.simps(2))
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   926
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   927
lemma Arg2pi_eq_0: "Arg2pi z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   928
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   929
  case False
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   930
  then have "z \<in> \<real> \<and> 0 \<le> Re z \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg2pi z)))"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   931
    by (metis cis.sel(1) cis_conv_exp cos_Arg2pi norm_ge_zero norm_le_zero_iff zero_le_mult_iff)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
   932
  also have "\<dots> \<longleftrightarrow> Arg2pi z = 0"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   933
  proof -
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   934
    have [simp]: "Arg2pi z = 0 \<Longrightarrow> z \<in> \<real>"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   935
      using Arg2pi_eq [of z] by (auto simp: Reals_def)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   936
    moreover have "\<lbrakk>z \<in> \<real>; 0 \<le> cos (Arg2pi z)\<rbrakk> \<Longrightarrow> Arg2pi z = 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   937
      by (smt (verit, ccfv_SIG) Arg2pi_ge_0 Arg2pi_le_pi Arg2pi_lt_pi complex_is_Real_iff cos_pi)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   938
    ultimately show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   939
      by (auto simp: Re_exp)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   940
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   941
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   942
    by blast
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   943
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   944
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
   945
corollary\<^marker>\<open>tag unimportant\<close> Arg2pi_gt_0:
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   946
  assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   947
    shows "Arg2pi z > 0"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   948
  using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   949
  unfolding nonneg_Reals_def by fastforce
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   950
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   951
lemma Arg2pi_eq_pi: "Arg2pi z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
   952
    using Arg2pi_le_pi [of z] Arg2pi_lt_pi [of z] Arg2pi_eq_0 [of z] Arg2pi_ge_0 [of z]
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   953
    by (fastforce simp: complex_is_Real_iff)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   954
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   955
lemma Arg2pi_eq_0_pi: "Arg2pi z = 0 \<or> Arg2pi z = pi \<longleftrightarrow> z \<in> \<real>"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   956
  using Arg2pi_eq_0 Arg2pi_eq_pi not_le by auto
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   957
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   958
lemma Arg2pi_of_real: "Arg2pi (of_real r) = (if r<0 then pi else 0)"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   959
  using Arg2pi_eq_0_pi Arg2pi_eq_pi by fastforce
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   960
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   961
lemma Arg2pi_real: "z \<in> \<real> \<Longrightarrow> Arg2pi z = (if 0 \<le> Re z then 0 else pi)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   962
  using Arg2pi_eq_0 Arg2pi_eq_0_pi by auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   963
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   964
lemma Arg2pi_inverse: "Arg2pi(inverse z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   965
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   966
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   967
  show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   968
    apply (rule Arg2pi_unique [of "inverse (norm z)"])
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   969
    using Arg2pi_eq False Arg2pi_ge_0 [of z] Arg2pi_lt_2pi [of z] Arg2pi_eq_0 [of z]
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   970
    by (auto simp: Arg2pi_real in_Reals_norm exp_diff field_simps)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   971
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   972
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   973
lemma Arg2pi_eq_iff:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   974
  assumes "w \<noteq> 0" "z \<noteq> 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
   975
  shows "Arg2pi w = Arg2pi z \<longleftrightarrow> (\<exists>x. 0 < x \<and> w = of_real x * z)" (is "?lhs = ?rhs")
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   976
proof
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   977
  assume ?lhs
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   978
  then have "(cmod w) * (z / cmod z) = w"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   979
    by (metis Arg2pi_eq assms(2) mult_eq_0_iff nonzero_mult_div_cancel_left)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   980
  then show ?rhs
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   981
    by (metis assms divide_pos_pos of_real_divide times_divide_eq_left times_divide_eq_right zero_less_norm_iff)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
   982
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   983
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   984
lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \<longleftrightarrow> Arg2pi z = 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   985
  by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   986
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   987
lemma Arg2pi_divide:
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   988
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg2pi w \<le> Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   989
    shows "Arg2pi(z / w) = Arg2pi z - Arg2pi w"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   990
  apply (rule Arg2pi_unique [of "norm(z / w)"])
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   991
  using assms Arg2pi_eq Arg2pi_ge_0 [of w] Arg2pi_lt_2pi [of z]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   992
  apply (auto simp: exp_diff norm_divide field_simps)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   993
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   994
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   995
lemma Arg2pi_le_div_sum:
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   996
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg2pi w \<le> Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   997
    shows "Arg2pi z = Arg2pi w + Arg2pi(z / w)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   998
  by (simp add: Arg2pi_divide assms)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   999
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1000
lemma Arg2pi_le_div_sum_eq:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1001
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1002
    shows "Arg2pi w \<le> Arg2pi z \<longleftrightarrow> Arg2pi z = Arg2pi w + Arg2pi(z / w)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1003
  using assms by (auto simp: Arg2pi_ge_0 intro: Arg2pi_le_div_sum)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1004
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1005
lemma Arg2pi_diff:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1006
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1007
    shows "Arg2pi w - Arg2pi z = (if Arg2pi z \<le> Arg2pi w then Arg2pi(w / z) else Arg2pi(w/z) - 2*pi)"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1008
  using assms Arg2pi_divide Arg2pi_inverse [of "w/z"] Arg2pi_eq_0_pi
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1009
  by (force simp add: Arg2pi_ge_0 Arg2pi_divide not_le split: if_split_asm)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1010
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1011
lemma Arg2pi_add:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1012
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1013
    shows "Arg2pi w + Arg2pi z = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi(w * z) else Arg2pi(w * z) + 2*pi)"
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1014
  using assms Arg2pi_diff [of "w*z" z] Arg2pi_le_div_sum_eq [of z "w*z"] Arg2pi [of "w * z"]
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1015
  by auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1016
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1017
lemma Arg2pi_times:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1018
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1019
    shows "Arg2pi (w * z) = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi w + Arg2pi z
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1020
                            else (Arg2pi w + Arg2pi z) - 2*pi)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1021
  using Arg2pi_add [OF assms] by auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1022
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1023
lemma Arg2pi_cnj_eq_inverse:
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1024
  assumes "z \<noteq> 0" shows "Arg2pi (cnj z) = Arg2pi (inverse z)"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1025
proof -
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1026
  have "\<exists>r>0. of_real r / z = cnj z"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1027
    by (metis assms complex_norm_square nonzero_mult_div_cancel_left zero_less_norm_iff zero_less_power)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1028
  then show ?thesis
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1029
    by (metis Arg2pi_times_of_real2 divide_inverse_commute)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1030
qed
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1031
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1032
lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1033
  by (metis Arg2pi_cnj_eq_inverse Arg2pi_inverse Reals_cnj_iff complex_cnj_zero)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1034
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1035
lemma Arg2pi_exp: "0 \<le> Im z \<Longrightarrow> Im z < 2*pi \<Longrightarrow> Arg2pi(exp z) = Im z"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1036
  by (simp add: Arg2pi_unique exp_eq_polar)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1037
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1038
lemma complex_split_polar:
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1039
  obtains r a::real where "z = complex_of_real r * (cos a + \<i> * sin a)" "0 \<le> r" "0 \<le> a" "a < 2*pi"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1040
  using Arg2pi cis.ctr cis_conv_exp unfolding Complex_eq is_Arg_def by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1041
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1042
lemma Re_Im_le_cmod: "Im w * sin \<phi> + Re w * cos \<phi> \<le> cmod w"
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1043
proof (cases w rule: complex_split_polar)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1044
  case (1 r a) with sin_cos_le1 [of a \<phi>] show ?thesis
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1045
    apply (simp add: norm_mult cmod_unit_one)
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 72301
diff changeset
  1046
    by (metis (no_types, opaque_lifting) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le)
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1047
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1048
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1049
subsection\<^marker>\<open>tag unimportant\<close>\<open>Analytic properties of tangent function\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1050
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1051
lemma cnj_tan: "cnj(tan z) = tan(cnj z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1052
  by (simp add: cnj_cos cnj_sin tan_def)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1053
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1054
lemma field_differentiable_at_tan: "cos z \<noteq> 0 \<Longrightarrow> tan field_differentiable at z"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1055
  unfolding field_differentiable_def
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1056
  using DERIV_tan by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1057
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1058
lemma field_differentiable_within_tan: "cos z \<noteq> 0
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1059
         \<Longrightarrow> tan field_differentiable (at z within s)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1060
  using field_differentiable_at_tan field_differentiable_at_within by blast
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1061
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1062
lemma continuous_within_tan: "cos z \<noteq> 0 \<Longrightarrow> continuous (at z within s) tan"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1063
  using continuous_at_imp_continuous_within isCont_tan by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1064
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1065
lemma continuous_on_tan [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> cos z \<noteq> 0) \<Longrightarrow> continuous_on s tan"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1066
  by (simp add: continuous_at_imp_continuous_on)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1067
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1068
lemma holomorphic_on_tan: "(\<And>z. z \<in> s \<Longrightarrow> cos z \<noteq> 0) \<Longrightarrow> tan holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1069
  by (simp add: field_differentiable_within_tan holomorphic_on_def)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1070
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1071
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1072
subsection\<open>The principal branch of the Complex logarithm\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1073
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1074
instantiation complex :: ln
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1075
begin
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1076
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1077
definition\<^marker>\<open>tag important\<close> ln_complex :: "complex \<Rightarrow> complex"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1078
  where "ln_complex \<equiv> \<lambda>z. THE w. exp w = z & -pi < Im(w) & Im(w) \<le> pi"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1079
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1080
text\<open>NOTE: within this scope, the constant Ln is not yet available!\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1081
lemma
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1082
  assumes "z \<noteq> 0"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1083
    shows exp_Ln [simp]:  "exp(ln z) = z"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1084
      and mpi_less_Im_Ln: "-pi < Im(ln z)"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1085
      and Im_Ln_le_pi:    "Im(ln z) \<le> pi"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1086
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1087
  obtain \<psi> where z: "z / (cmod z) = Complex (cos \<psi>) (sin \<psi>)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1088
    using complex_unimodular_polar [of "z / (norm z)"] assms
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  1089
    by (auto simp: norm_divide field_split_simps)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1090
  obtain \<phi> where \<phi>: "- pi < \<phi>" "\<phi> \<le> pi" "sin \<phi> = sin \<psi>" "cos \<phi> = cos \<psi>"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1091
    using sincos_principal_value [of "\<psi>"] assms
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  1092
    by (auto simp: norm_divide field_split_simps)
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1093
  have "exp(ln z) = z & -pi < Im(ln z) & Im(ln z) \<le> pi" unfolding ln_complex_def
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1094
    apply (rule theI [where a = "Complex (ln(norm z)) \<phi>"])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1095
    using z assms \<phi>
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1096
    apply (auto simp: field_simps exp_complex_eqI exp_eq_polar cis.code)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1097
    done
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1098
  then show "exp(ln z) = z" "-pi < Im(ln z)" "Im(ln z) \<le> pi"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1099
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1100
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1101
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1102
lemma Ln_exp [simp]:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1103
  assumes "-pi < Im(z)" "Im(z) \<le> pi"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1104
    shows "ln(exp z) = z"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1105
proof (rule exp_complex_eqI)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1106
  show "\<bar>Im (ln (exp z)) - Im z\<bar> < 2 * pi"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1107
    using assms mpi_less_Im_Ln  [of "exp z"] Im_Ln_le_pi [of "exp z"] by auto
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1108
qed auto
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1109
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  1110
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1111
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation to Real Logarithm\<close>
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1112
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1113
lemma Ln_of_real:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1114
  assumes "0 < z"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1115
    shows "ln(of_real z::complex) = of_real(ln z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1116
  by (smt (verit) Im_complex_of_real Ln_exp assms exp_ln of_real_exp pi_ge_two)
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1117
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1118
corollary\<^marker>\<open>tag unimportant\<close> Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<real>"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1119
  by (auto simp: Ln_of_real elim: Reals_cases)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1120
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1121
corollary\<^marker>\<open>tag unimportant\<close> Im_Ln_of_real [simp]: "r > 0 \<Longrightarrow> Im (ln (of_real r)) = 0"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1122
  by (simp add: Ln_of_real)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1123
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
  1124
lemma cmod_Ln_Reals [simp]: "z \<in> \<real> \<Longrightarrow> 0 < Re z \<Longrightarrow> cmod (ln z) = norm (ln (Re z))"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1125
  using Ln_of_real by force
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1126
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1127
lemma Ln_Reals_eq: "\<lbrakk>x \<in> \<real>; Re x > 0\<rbrakk> \<Longrightarrow> ln x = of_real (ln (Re x))"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1128
  using Ln_of_real by force
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1129
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1130
lemma Ln_1 [simp]: "ln 1 = (0::complex)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1131
  by (smt (verit, best) Ln_of_real ln_one of_real_0 of_real_1)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1132
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1133
lemma Ln_eq_zero_iff [simp]: "x \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln x = 0 \<longleftrightarrow> x = 1" for x::complex
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1134
  by (metis (mono_tags, lifting) Ln_1 exp_Ln exp_zero nonpos_Reals_zero_I)
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1135
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1136
instance
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1137
  by intro_classes (rule ln_complex_def Ln_1)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1138
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1139
end
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1140
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1141
abbreviation Ln :: "complex \<Rightarrow> complex"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1142
  where "Ln \<equiv> ln"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1143
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1144
lemma Ln_eq_iff: "w \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> (Ln w = Ln z \<longleftrightarrow> w = z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1145
  by (metis exp_Ln)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1146
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1147
lemma Ln_unique: "exp(z) = w \<Longrightarrow> -pi < Im(z) \<Longrightarrow> Im(z) \<le> pi \<Longrightarrow> Ln w = z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1148
  using Ln_exp by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1149
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1150
lemma Re_Ln [simp]: "z \<noteq> 0 \<Longrightarrow> Re(Ln z) = ln(norm z)"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1151
  by (metis exp_Ln ln_exp norm_exp_eq_Re)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1152
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1153
corollary\<^marker>\<open>tag unimportant\<close> ln_cmod_le:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1154
  assumes z: "z \<noteq> 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1155
    shows "ln (cmod z) \<le> cmod (Ln z)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1156
  by (metis Re_Ln complex_Re_le_cmod z)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1157
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1158
proposition\<^marker>\<open>tag unimportant\<close> exists_complex_root:
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1159
  fixes z :: complex
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1160
  assumes "n \<noteq> 0"  obtains w where "z = w ^ n"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1161
  by (metis assms exp_Ln exp_of_nat_mult nonzero_mult_div_cancel_left of_nat_eq_0_iff power_0_left times_divide_eq_right)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1162
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1163
corollary\<^marker>\<open>tag unimportant\<close> exists_complex_root_nonzero:
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1164
  fixes z::complex
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1165
  assumes "z \<noteq> 0" "n \<noteq> 0"
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1166
  obtains w where "w \<noteq> 0" "z = w ^ n"
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1167
  by (metis exists_complex_root [of n z] assms power_0_left)
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1168
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1169
subsection\<^marker>\<open>tag unimportant\<close>\<open>Derivative of Ln away from the branch cut\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1170
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1171
lemma Im_Ln_less_pi: 
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1172
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"shows "Im (Ln z) < pi"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1173
proof -
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1174
  have znz [simp]: "z \<noteq> 0"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1175
    using assms by auto
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1176
  with Im_Ln_le_pi [of z] show ?thesis
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1177
    by (smt (verit, best) Arg2pi_eq_0_pi Arg2pi_exp Ln_in_Reals assms complex_is_Real_iff complex_nonpos_Reals_iff exp_Ln pi_ge_two)
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1178
qed
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1179
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1180
lemma has_field_derivative_Ln: 
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1181
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1182
  shows "(Ln has_field_derivative inverse(z)) (at z)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1183
proof -
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1184
  have znz [simp]: "z \<noteq> 0"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1185
    using assms by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1186
  then have "Im (Ln z) \<noteq> pi"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1187
    by (smt (verit, best) Arg2pi_eq_0_pi Arg2pi_exp Ln_in_Reals assms complex_is_Real_iff complex_nonpos_Reals_iff exp_Ln pi_ge_two)
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1188
  let ?U = "{w. -pi < Im(w) \<and> Im(w) < pi}"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1189
  have 1: "open ?U"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1190
    by (simp add: open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1191
  have 2: "\<And>x. x \<in> ?U \<Longrightarrow> (exp has_derivative blinfun_apply(Blinfun ((*) (exp x)))) (at x)"
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
  1192
    by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right has_field_derivative_imp_has_derivative)
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1193
  have 3: "continuous_on ?U (\<lambda>x. Blinfun ((*) (exp x)))"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1194
    unfolding blinfun_mult_right.abs_eq [symmetric] by (intro continuous_intros)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1195
  have 4: "Ln z \<in> ?U"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1196
    by (simp add: Im_Ln_less_pi assms mpi_less_Im_Ln)
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1197
  have 5: "Blinfun ((*) (inverse z)) o\<^sub>L Blinfun ((*) (exp (Ln z))) = id_blinfun"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1198
    by (rule blinfun_eqI) (simp add: bounded_linear_mult_right bounded_linear_Blinfun_apply)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1199
  obtain U' V g g' where "open U'" and sub: "U' \<subseteq> ?U"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1200
    and "Ln z \<in> U'" "open V" "z \<in> V"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1201
    and hom: "homeomorphism U' V exp g"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1202
    and g: "\<And>y. y \<in> V \<Longrightarrow> (g has_derivative (g' y)) (at y)"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1203
    and g': "\<And>y. y \<in> V \<Longrightarrow> g' y = inv ((*) (exp (g y)))"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1204
    and bij: "\<And>y. y \<in> V \<Longrightarrow> bij ((*) (exp (g y)))"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1205
    using inverse_function_theorem [OF 1 2 3 4 5]
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1206
    by (simp add: bounded_linear_Blinfun_apply bounded_linear_mult_right) blast
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1207
  show "(Ln has_field_derivative inverse(z)) (at z)"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1208
    unfolding has_field_derivative_def
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1209
  proof (rule has_derivative_transform_within_open)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1210
    show g_eq_Ln: "g y = Ln y" if "y \<in> V" for y
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1211
      by (smt (verit, ccfv_threshold) Ln_exp hom homeomorphism_def imageI mem_Collect_eq sub subset_iff that)
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1212
    have "0 \<notin> V"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1213
      by (meson exp_not_eq_zero hom homeomorphism_def)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1214
    then have "\<And>y. y \<in> V \<Longrightarrow> g' y = inv ((*) y)"
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1215
      by (metis exp_Ln g' g_eq_Ln)
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1216
    then have g': "g' z = (\<lambda>x. x/z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1217
      by (metis \<open>z \<in> V\<close> bij bij_inv_eq_iff exp_Ln g_eq_Ln nonzero_mult_div_cancel_left znz)
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1218
    show "(g has_derivative (*) (inverse z)) (at z)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1219
      using g [OF \<open>z \<in> V\<close>] g' by (simp add: divide_inverse_commute)
70999
5b753486c075 Inverse function theorem + lemmas
paulson <lp15@cam.ac.uk>
parents: 70817
diff changeset
  1220
  qed (auto simp: \<open>z \<in> V\<close> \<open>open V\<close>)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1221
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1222
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1223
declare has_field_derivative_Ln [derivative_intros]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1224
declare has_field_derivative_Ln [THEN DERIV_chain2, derivative_intros]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1225
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1226
lemma field_differentiable_at_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Ln field_differentiable at z"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1227
  using field_differentiable_def has_field_derivative_Ln by blast
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1228
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1229
lemma field_differentiable_within_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1230
         \<Longrightarrow> Ln field_differentiable (at z within S)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1231
  using field_differentiable_at_Ln field_differentiable_within_subset by blast
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1232
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1233
lemma continuous_at_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z) Ln"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1234
  by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Ln)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1235
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70196
diff changeset
  1236
lemma isCont_Ln' [simp,continuous_intros]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1237
   "\<lbrakk>isCont f z; f z \<notin> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. Ln (f x)) z"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1238
  by (blast intro: isCont_o2 [OF _ continuous_at_Ln])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1239
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70196
diff changeset
  1240
lemma continuous_within_Ln [continuous_intros]: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within S) Ln"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1241
  using continuous_at_Ln continuous_at_imp_continuous_within by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1242
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1243
lemma continuous_on_Ln [continuous_intros]: "(\<And>z. z \<in> S \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on S Ln"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1244
  by (simp add: continuous_at_imp_continuous_on continuous_within_Ln)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1245
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1246
lemma continuous_on_Ln' [continuous_intros]:
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1247
  "continuous_on S f \<Longrightarrow> (\<And>z. z \<in> S \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on S (\<lambda>x. Ln (f x))"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1248
  by (rule continuous_on_compose2[OF continuous_on_Ln, of "UNIV - nonpos_Reals" S f]) auto
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1249
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1250
lemma holomorphic_on_Ln [holomorphic_intros]: "S \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Ln holomorphic_on S"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1251
  by (simp add: disjoint_iff field_differentiable_within_Ln holomorphic_on_def)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1252
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1253
lemma holomorphic_on_Ln' [holomorphic_intros]:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1254
  "(\<And>z. z \<in> A \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> f holomorphic_on A \<Longrightarrow> (\<lambda>z. Ln (f z)) holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1255
  using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \<real>\<^sub>\<le>\<^sub>0"]
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1256
  by (auto simp: o_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1257
79857
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1258
lemma analytic_on_ln [analytic_intros]:
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1259
  assumes "f analytic_on A" "f ` A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {}"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1260
  shows   "(\<lambda>w. ln (f w)) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1261
proof -
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1262
  have *: "ln analytic_on (-\<real>\<^sub>\<le>\<^sub>0)"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1263
    by (subst analytic_on_open) (auto intro!: holomorphic_intros)
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1264
  have "(ln \<circ> f) analytic_on A"
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1265
    by (rule analytic_on_compose_gen[OF assms(1) *]) (use assms(2) in auto)
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1266
  thus ?thesis
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1267
    by (simp add: o_def)
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1268
qed
819c28a7280f New material by Wenda Li and Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 79670
diff changeset
  1269
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1270
lemma tendsto_Ln [tendsto_intros]:
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1271
  assumes "(f \<longlongrightarrow> L) F" "L \<notin> \<real>\<^sub>\<le>\<^sub>0"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1272
  shows   "((\<lambda>x. Ln (f x)) \<longlongrightarrow> Ln L) F"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1273
  by (simp add: assms isCont_tendsto_compose)
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1274
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1275
lemma divide_ln_mono:
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1276
  fixes x y::real
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1277
  assumes "3 \<le> x" "x \<le> y"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1278
  shows "x / ln x \<le> y / ln y"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1279
proof -
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1280
  have "\<And>u. \<lbrakk>x \<le> u; u \<le> y\<rbrakk> \<Longrightarrow> ((\<lambda>z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1281
    using \<open>3 \<le> x\<close> by (force intro!: derivative_eq_intros simp: field_simps power_eq_if)
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1282
  moreover
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1283
  have "x / ln x \<le> y / ln y"
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1284
    if "Re (y / Ln y) - Re (x / Ln x) = (Re (1 / Ln u) - Re (1 / (Ln u)\<^sup>2)) * (y - x)"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1285
    and x: "x \<le> u" "u \<le> y" for u
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1286
  proof -
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1287
    have eq: "y / ln y = (1 / ln u - 1 / (ln u)\<^sup>2) * (y - x) + x / ln x"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1288
      using that \<open>3 \<le> x\<close> by (auto simp: Ln_Reals_eq in_Reals_norm group_add_class.diff_eq_eq)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1289
    show ?thesis
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1290
      using exp_le \<open>3 \<le> x\<close> x by (simp add: eq) (simp add: power_eq_if divide_simps ln_ge_iff)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1291
  qed
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1292
  ultimately show ?thesis
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1293
    using complex_mvt_line [of x y "\<lambda>z. z / Ln z" "\<lambda>z. 1/(Ln z) - 1/(Ln z)^2"] assms
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1294
    by (force simp add: closed_segment_Reals closed_segment_eq_real_ivl)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1295
qed
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1296
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1297
theorem Ln_series:
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1298
  fixes z :: complex
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1299
  assumes "norm z < 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1300
  shows   "(\<lambda>n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\<lambda>n. ?f n * z^n) sums _")
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1301
proof -
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1302
  let ?F = "\<lambda>z. \<Sum>n. ?f n * z^n" and ?F' = "\<lambda>z. \<Sum>n. diffs ?f n * z^n"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1303
  have r: "conv_radius ?f = 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1304
    by (intro conv_radius_ratio_limit_nonzero[of _ 1])
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1305
       (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1306
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1307
  have "\<exists>c. \<forall>z\<in>ball 0 1. ln (1 + z) - ?F z = c"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1308
  proof (rule has_field_derivative_zero_constant)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1309
    fix z :: complex assume z': "z \<in> ball 0 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1310
    hence z: "norm z < 1" by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1311
    define t :: complex where "t = of_real (1 + norm z) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1312
    from z have t: "norm z < norm t" "norm t < 1" unfolding t_def
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1313
      by (simp_all add: field_simps norm_divide del: of_real_add)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1314
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1315
    have "Re (-z) \<le> norm (-z)" by (rule complex_Re_le_cmod)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1316
    also from z have "\<dots> < 1" by simp
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1317
    finally have "((\<lambda>z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1318
      by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1319
    moreover have "(?F has_field_derivative ?F' z) (at z)" using t r
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1320
      by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1321
    ultimately have "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z))
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1322
                       (at z within ball 0 1)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1323
      by (intro derivative_intros) (simp_all add: at_within_open[OF z'])
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1324
    also have "(\<lambda>n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1325
      by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1326
    from sums_split_initial_segment[OF this, of 1]
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1327
      have "(\<lambda>i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1328
    hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1329
    also have "inverse (1 + z) - inverse (1 + z) = 0" by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1330
    finally show "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" .
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1331
  qed simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1332
  then obtain c where c: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> ln (1 + z) - ?F z = c" by blast
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1333
  from c[of 0] have "c = 0" by (simp only: powser_zero) simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1334
  with c[of z] assms have "ln (1 + z) = ?F z" by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1335
  moreover have "summable (\<lambda>n. ?f n * z^n)" using assms r
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1336
    by (intro summable_in_conv_radius) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1337
  ultimately show ?thesis by (simp add: sums_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1338
qed
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1339
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1340
lemma Ln_series': "cmod z < 1 \<Longrightarrow> (\<lambda>n. - ((-z)^n) / of_nat n) sums ln (1 + z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1341
  by (drule Ln_series) (simp add: power_minus')
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1342
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1343
lemma ln_series':
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1344
  fixes x::real
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1345
  assumes "\<bar>x\<bar> < 1"
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1346
  shows   "(\<lambda>n. - ((-x)^n) / of_nat n) sums ln (1 + x)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1347
proof -
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1348
  from assms have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1349
    by (intro Ln_series') simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1350
  also have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) = (\<lambda>n. complex_of_real (- ((-x)^n) / of_nat n))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1351
    by (rule ext) simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1352
  also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1353
    by (smt (verit) Ln_of_real of_real_1 of_real_add)
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1354
  finally show ?thesis by (subst (asm) sums_of_real_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1355
qed
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1356
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1357
lemma Ln_approx_linear:
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1358
  fixes z :: complex
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1359
  assumes "norm z < 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1360
  shows   "norm (ln (1 + z) - z) \<le> norm z^2 / (1 - norm z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1361
proof -
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1362
  let ?f = "\<lambda>n. (-1)^Suc n / of_nat n"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1363
  from assms have "(\<lambda>n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1364
  moreover have "(\<lambda>n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1365
  ultimately have "(\<lambda>n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1366
    by (subst left_diff_distrib, intro sums_diff) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1367
  from sums_split_initial_segment[OF this, of "Suc 1"]
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1368
    have "(\<lambda>i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1369
    by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1370
  hence "(Ln (1 + z) - z) = (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1371
    by (simp add: sums_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1372
  also have A: "summable (\<lambda>n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1373
    by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]])
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1374
       (auto simp: assms field_simps intro!: always_eventually)
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
  1375
  hence "norm (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1376
      \<le> (\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))"
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1377
    by (intro summable_norm)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1378
       (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1379
  also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \<le> norm ((-z)^2 * (-z)^i) * 1" for i
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  1380
    by (intro mult_left_mono) (simp_all add: field_split_simps)
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1381
  hence "(\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1382
       \<le> (\<Sum>i. norm (-(z^2) * (-z)^i))"
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1383
    using A assms
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1384
    unfolding norm_power norm_inverse norm_divide norm_mult
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1385
    apply (intro suminf_le summable_mult summable_geometric)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1386
    apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1387
    done
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1388
  also have "\<dots> = norm z^2 * (\<Sum>i. norm z^i)" using assms
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1389
    by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1390
  also have "(\<Sum>i. norm z^i) = inverse (1 - norm z)" using assms
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1391
    by (subst suminf_geometric) (simp_all add: divide_inverse)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1392
  also have "norm z^2 * \<dots> = norm z^2 / (1 - norm z)" by (simp add: divide_inverse)
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1393
  finally show ?thesis .
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1394
qed
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1395
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1396
76722
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1397
lemma norm_Ln_le:
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1398
  fixes z :: complex
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1399
  assumes "norm z < 1/2"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1400
  shows   "norm (Ln(1+z)) \<le> 2 * norm z"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1401
proof -
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1402
  have sums: "(\<lambda>n. - ((- z) ^ n) / of_nat n) sums ln (1 + z)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1403
    by (intro Ln_series') (use assms in auto)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1404
  have summable: "summable (\<lambda>n. norm (- ((- z) ^ n / of_nat n)))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1405
    using ln_series'[of "-norm z"] assms
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1406
    by (simp add: sums_iff summable_minus_iff norm_power norm_divide)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1407
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1408
  have "norm (ln (1 + z)) = norm (\<Sum>n. -((-z) ^ n / of_nat n))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1409
    using sums by (simp add: sums_iff)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1410
  also have "\<dots> \<le> (\<Sum>n. norm (-((-z) ^ n / of_nat n)))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1411
    using summable by (rule summable_norm)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1412
  also have "\<dots> = (\<Sum>n. norm (-((-z) ^ Suc n / of_nat (Suc n))))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1413
    using summable by (subst suminf_split_head) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1414
  also have "\<dots> \<le> (\<Sum>n. norm z * (1 / 2) ^ n)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1415
  proof (rule suminf_le)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1416
    show "summable (\<lambda>n. norm z * (1 / 2) ^ n)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1417
      by (intro summable_mult summable_geometric) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1418
  next
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1419
    show "summable (\<lambda>n. norm (- ((- z) ^ Suc n / of_nat (Suc n))))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1420
      using summable by (subst summable_Suc_iff)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1421
  next
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1422
    fix n
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1423
    have "norm (-((-z) ^ Suc n / of_nat (Suc n))) = norm z * (norm z ^ n / real (Suc n))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1424
      by (simp add: norm_power norm_divide norm_mult del: of_nat_Suc)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1425
    also have "\<dots> \<le> norm z * ((1 / 2) ^ n / 1)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1426
      using assms by (intro mult_left_mono frac_le power_mono) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1427
    finally show "norm (- ((- z) ^ Suc n / of_nat (Suc n))) \<le> norm z * (1 / 2) ^ n"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1428
      by simp
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1429
  qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1430
  also have "\<dots> = norm z * (\<Sum>n. (1 / 2) ^ n)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1431
    by (subst suminf_mult) (auto intro: summable_geometric)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1432
  also have "(\<Sum>n. (1 / 2 :: real) ^ n) = 2"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1433
    using geometric_sums[of "1 / 2 :: real"] by (simp add: sums_iff)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1434
  finally show ?thesis
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1435
    by (simp add: mult_ac)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1436
qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1437
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1438
subsection\<^marker>\<open>tag unimportant\<close>\<open>Quadrant-type results for Ln\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1439
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1440
lemma cos_lt_zero_pi: "pi/2 < x \<Longrightarrow> x < 3*pi/2 \<Longrightarrow> cos x < 0"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1441
  using cos_minus_pi cos_gt_zero_pi [of "x-pi"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1442
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1443
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1444
lemma Re_Ln_pos_le:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1445
  assumes "z \<noteq> 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1446
  shows "\<bar>Im(Ln z)\<bar> \<le> pi/2 \<longleftrightarrow> 0 \<le> Re(z)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1447
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1448
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1449
    assume "w = Ln z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1450
    then have w: "Im w \<le> pi" "- pi < Im w"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1451
      using Im_Ln_le_pi [of z]  mpi_less_Im_Ln [of z]  assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1452
      by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1453
    then have "\<bar>Im w\<bar> \<le> pi/2 \<longleftrightarrow> 0 \<le> Re(exp w)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1454
      using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1455
      by (auto simp: Re_exp zero_le_mult_iff abs_if intro: cos_ge_zero)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1456
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1457
  then show ?thesis using assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1458
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1459
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1460
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1461
lemma Re_Ln_pos_lt:
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1462
  assumes "z \<noteq> 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1463
  shows "\<bar>Im(Ln z)\<bar> < pi/2 \<longleftrightarrow> 0 < Re(z)"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1464
  using Re_Ln_pos_le assms
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1465
  by (smt (verit) Re_exp arccos_cos cos_minus cos_pi_half exp_Ln exp_gt_zero field_sum_of_halves mult_eq_0_iff)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1466
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1467
lemma Im_Ln_pos_le:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1468
  assumes "z \<noteq> 0"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1469
    shows "0 \<le> Im(Ln z) \<and> Im(Ln z) \<le> pi \<longleftrightarrow> 0 \<le> Im(z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1470
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1471
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1472
    assume "w = Ln z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1473
    then have w: "Im w \<le> pi" "- pi < Im w"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1474
      using Im_Ln_le_pi [of z]  mpi_less_Im_Ln [of z]  assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1475
      by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1476
    then have "0 \<le> Im w \<and> Im w \<le> pi \<longleftrightarrow> 0 \<le> Im(exp w)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1477
      using sin_ge_zero [of "- (Im w)"] sin_ge_zero [of "abs(Im w)"] sin_zero_pi_iff [of "Im w"]
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1478
      by (force simp: Im_exp zero_le_mult_iff sin_ge_zero) }
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1479
  then show ?thesis using assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1480
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1481
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1482
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1483
lemma Im_Ln_pos_lt:
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1484
  assumes "z \<noteq> 0"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1485
  shows "0 < Im(Ln z) \<and> Im(Ln z) < pi \<longleftrightarrow> 0 < Im(z)"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1486
  using Im_Ln_pos_le [OF assms] assms
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1487
  by (smt (verit, best) Arg2pi_exp Arg2pi_lt_pi exp_Ln)
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1488
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1489
lemma Re_Ln_pos_lt_imp: "0 < Re(z) \<Longrightarrow> \<bar>Im(Ln z)\<bar> < pi/2"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1490
  by (metis Re_Ln_pos_lt less_irrefl zero_complex.simps(1))
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1491
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1492
lemma Im_Ln_pos_lt_imp: "0 < Im(z) \<Longrightarrow> 0 < Im(Ln z) \<and> Im(Ln z) < pi"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1493
  by (metis Im_Ln_pos_lt not_le order_refl zero_complex.simps(2))
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1494
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1495
text\<open>A reference to the set of positive real numbers\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1496
lemma Im_Ln_eq_0: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = 0 \<longleftrightarrow> 0 < Re(z) \<and> Im(z) = 0)"
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1497
  using Im_Ln_pos_le Im_Ln_pos_lt Re_Ln_pos_lt by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1498
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1499
lemma Im_Ln_eq_pi: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = pi \<longleftrightarrow> Re(z) < 0 \<and> Im(z) = 0)"
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1500
  using Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt complex.expand by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1501
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1502
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1503
subsection\<^marker>\<open>tag unimportant\<close>\<open>More Properties of Ln\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1504
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1505
lemma cnj_Ln: assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1506
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1507
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1508
  show ?thesis
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1509
    by (smt (verit) False Im_Ln_less_pi Ln_exp assms cnj.sel(2) exp_Ln exp_cnj mpi_less_Im_Ln)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1510
qed (use assms in auto)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1511
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1512
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1513
lemma Ln_inverse: assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" shows "Ln(inverse z) = -(Ln z)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1514
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1515
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1516
  show ?thesis
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1517
    by (smt (verit) False Im_Ln_less_pi Ln_exp assms exp_Ln exp_minus mpi_less_Im_Ln uminus_complex.sel(2))
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1518
qed (use assms in auto)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1519
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1520
lemma Ln_minus1 [simp]: "Ln(-1) = \<i> * pi"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1521
proof (rule exp_complex_eqI)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1522
  show "\<bar>Im (Ln (- 1)) - Im (\<i> * pi)\<bar> < 2 * pi"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1523
    using Im_Ln_le_pi [of "-1"] mpi_less_Im_Ln [of "-1"] by auto
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1524
qed auto
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1525
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1526
lemma Ln_ii [simp]: "Ln \<i> = \<i> * of_real pi/2"
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1527
  using Ln_exp [of "\<i> * (of_real pi/2)"]
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1528
  unfolding exp_Euler
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1529
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1530
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1531
lemma Ln_minus_ii [simp]: "Ln(-\<i>) = - (\<i> * pi/2)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1532
  using Ln_inverse by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1533
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1534
lemma Ln_times:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1535
  assumes "w \<noteq> 0" "z \<noteq> 0"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1536
    shows "Ln(w * z) =
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1537
           (if Im(Ln w + Ln z) \<le> -pi then (Ln(w) + Ln(z)) + \<i> * of_real(2*pi)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1538
            else if Im(Ln w + Ln z) > pi then (Ln(w) + Ln(z)) - \<i> * of_real(2*pi)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1539
            else Ln(w) + Ln(z))"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1540
  using pi_ge_zero Im_Ln_le_pi [of w] Im_Ln_le_pi [of z]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1541
  using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1542
  by (auto simp: exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1543
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1544
corollary\<^marker>\<open>tag unimportant\<close> Ln_times_simple:
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1545
    "\<lbrakk>w \<noteq> 0; z \<noteq> 0; -pi < Im(Ln w) + Im(Ln z); Im(Ln w) + Im(Ln z) \<le> pi\<rbrakk>
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1546
         \<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1547
  by (simp add: Ln_times)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1548
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1549
corollary\<^marker>\<open>tag unimportant\<close> Ln_times_of_real:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1550
    "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(of_real r * z) = ln r + Ln(z)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1551
  using mpi_less_Im_Ln Im_Ln_le_pi
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1552
  by (force simp: Ln_times)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1553
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1554
corollary\<^marker>\<open>tag unimportant\<close> Ln_times_of_nat:
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1555
    "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(of_nat r * z :: complex) = ln (of_nat r) + Ln(z)"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1556
  using Ln_times_of_real[of "of_nat r" z] by simp
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1557
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1558
corollary\<^marker>\<open>tag unimportant\<close> Ln_times_Reals:
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1559
    "\<lbrakk>r \<in> Reals; Re r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(r * z) = ln (Re r) + Ln(z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1560
  using Ln_Reals_eq Ln_times_of_real by fastforce
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1561
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1562
corollary\<^marker>\<open>tag unimportant\<close> Ln_divide_of_real:
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1563
  "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(z / of_real r) = Ln(z) - ln r"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1564
  using Ln_times_of_real [of "inverse r" z]
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1565
  by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse flip: of_real_inverse)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1566
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  1567
corollary\<^marker>\<open>tag unimportant\<close> Ln_prod:
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1568
  fixes f :: "'a \<Rightarrow> complex"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1569
  assumes "finite A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1570
  shows "\<exists>n. Ln (prod f A) = (\<Sum>x \<in> A. Ln (f x) + (of_int (n x) * (2 * pi)) * \<i>)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1571
  using assms
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1572
proof (induction A)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1573
  case (insert x A)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1574
  then obtain n where n: "Ln (prod f A) = (\<Sum>x\<in>A. Ln (f x) + of_real (of_int (n x) * (2 * pi)) * \<i>)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1575
    by auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1576
  define D where "D \<equiv> Im (Ln (f x)) + Im (Ln (prod f A))"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1577
  define q::int where "q \<equiv> (if D \<le> -pi then 1 else if D > pi then -1 else 0)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1578
  have "prod f A \<noteq> 0" "f x \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1579
    by (auto simp: insert.hyps insert.prems)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1580
  with insert.hyps pi_ge_zero show ?case
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1581
    by (rule_tac x="n(x:=q)" in exI) (force simp: Ln_times q_def D_def n intro!: sum.cong)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1582
qed auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1583
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1584
lemma Ln_minus:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1585
  assumes "z \<noteq> 0"
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1586
    shows "Ln(-z) = (if Im(z) \<le> 0 \<and> \<not>(Re(z) < 0 \<and> Im(z) = 0)
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1587
                     then Ln(z) + \<i> * pi
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1588
                     else Ln(z) - \<i> * pi)" 
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1589
  using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1590
        Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z]
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1591
  by (intro Ln_unique) (auto simp: exp_add exp_diff)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1592
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1593
lemma Ln_inverse_if:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1594
  assumes "z \<noteq> 0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1595
    shows "Ln (inverse z) = (if z \<in> \<real>\<^sub>\<le>\<^sub>0 then -(Ln z) + \<i> * 2 * complex_of_real pi else -(Ln z))"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1596
proof (cases "z \<in> \<real>\<^sub>\<le>\<^sub>0")
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1597
  case False then show ?thesis
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1598
    by (simp add: Ln_inverse)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1599
next
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1600
  case True
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1601
  then have z: "Im z = 0" "Re z < 0" "- z \<notin> \<real>\<^sub>\<le>\<^sub>0"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1602
    using assms complex_eq_iff complex_nonpos_Reals_iff by auto
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1603
  have "Ln(inverse z) = Ln(- (inverse (-z)))"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1604
    by simp
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1605
  also have "\<dots> = Ln (inverse (-z)) + \<i> * complex_of_real pi"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1606
    using assms z by (simp add: Ln_minus divide_less_0_iff)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1607
  also have "\<dots> = - Ln (- z) + \<i> * complex_of_real pi"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1608
    using z Ln_inverse by presburger
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1609
  also have "\<dots> = - (Ln z) + \<i> * 2 * complex_of_real pi"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  1610
    using Ln_minus assms z by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1611
  finally show ?thesis by (simp add: True)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1612
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1613
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1614
lemma Ln_times_ii:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1615
  assumes "z \<noteq> 0"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1616
    shows  "Ln(\<i> * z) = (if 0 \<le> Re(z) | Im(z) < 0
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1617
                          then Ln(z) + \<i> * of_real pi/2
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1618
                          else Ln(z) - \<i> * of_real(3 * pi/2))"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1619
  using Im_Ln_le_pi [of z] mpi_less_Im_Ln [of z] assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1620
        Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] Re_Ln_pos_le [of z]
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  1621
  by (simp add: Ln_times) auto
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1622
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65585
diff changeset
  1623
lemma Ln_of_nat [simp]: "0 < n \<Longrightarrow> Ln (of_nat n) = of_real (ln (of_nat n))"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1624
  by (metis Ln_of_real of_nat_0_less_iff of_real_of_nat_eq)
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1625
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1626
lemma Ln_of_nat_over_of_nat:
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1627
  assumes "m > 0" "n > 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1628
  shows   "Ln (of_nat m / of_nat n) = of_real (ln (of_nat m) - ln (of_nat n))"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1629
proof -
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1630
  have "of_nat m / of_nat n = (of_real (of_nat m / of_nat n) :: complex)" by simp
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1631
  also from assms have "Ln \<dots> = of_real (ln (of_nat m / of_nat n))"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1632
    by (simp add: Ln_of_real[symmetric])
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  1633
  also from assms have "\<dots> = of_real (ln (of_nat m) - ln (of_nat n))"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1634
    by (simp add: ln_div)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1635
  finally show ?thesis .
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1636
qed
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1637
76722
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1638
lemma norm_Ln_times_le:
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1639
  assumes "w \<noteq> 0" "z \<noteq> 0"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1640
  shows  "cmod (Ln(w * z)) \<le> cmod (Ln(w) + Ln(z))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1641
proof (cases "- pi < Im(Ln w + Ln z) \<and> Im(Ln w + Ln z) \<le> pi")
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1642
  case True
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1643
  then show ?thesis
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1644
    by (simp add: Ln_times_simple assms)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1645
next
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1646
  case False
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1647
  then show ?thesis
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1648
    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)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1649
qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1650
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1651
corollary norm_Ln_prod_le:
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1652
  fixes f :: "'a \<Rightarrow> complex"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1653
  assumes "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> 0"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1654
  shows "cmod (Ln (prod f A)) \<le> (\<Sum>x \<in> A. cmod (Ln (f x)))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1655
  using assms
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1656
proof (induction A rule: infinite_finite_induct)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1657
  case (insert x A)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1658
  then show ?case
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1659
    by simp (smt (verit) norm_Ln_times_le norm_triangle_ineq prod_zero_iff)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1660
qed auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1661
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1662
lemma norm_Ln_exp_le: "norm (Ln (exp z)) \<le> norm z"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1663
  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)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1664
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1665
subsection\<^marker>\<open>tag unimportant\<close>\<open>Uniform convergence and products\<close>
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1666
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1667
(* TODO: could be generalised perhaps, but then one would have to do without the ln *)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1668
lemma uniformly_convergent_on_prod_aux:
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1669
  fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1670
  assumes norm_f: "\<And>n x. x \<in> A \<Longrightarrow> norm (f n x) < 1"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1671
  assumes cont: "\<And>n. continuous_on A (f n)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1672
  assumes conv: "uniformly_convergent_on A (\<lambda>N x. \<Sum>n<N. ln (1 + f n x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1673
  assumes A: "compact A"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1674
  shows "uniformly_convergent_on A (\<lambda>N x. \<Prod>n<N. 1 + f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1675
proof -
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1676
  from conv obtain S where S: "uniform_limit A (\<lambda>N x. \<Sum>n<N. ln (1 + f n x)) S sequentially"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1677
    by (auto simp: uniformly_convergent_on_def)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1678
  have cont': "continuous_on A S"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1679
  proof (rule uniform_limit_theorem[OF _ S])
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1680
    have "f n x + 1 \<notin> \<real>\<^sub>\<le>\<^sub>0" if "x \<in> A" for n x
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1681
    proof
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1682
      assume "f n x + 1 \<in> \<real>\<^sub>\<le>\<^sub>0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1683
      then obtain t where t: "t \<le> 0" "f n x = of_real (t - 1)"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1684
        by (metis add_diff_cancel nonpos_Reals_cases of_real_1 of_real_diff)
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1685
      moreover have "norm \<dots> \<ge> 1"
76722
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1686
        using t by (subst norm_of_real) auto
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1687
      ultimately show False
76722
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1688
        using norm_f[of x n] that by auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1689
    qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1690
    thus "\<forall>\<^sub>F n in sequentially. continuous_on A (\<lambda>x. \<Sum>n<n. Ln (1 + f n x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1691
      by (auto intro!: always_eventually continuous_intros cont simp: add_ac)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1692
  qed auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1693
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1694
  define B where "B = {x + y |x y. x \<in> S ` A \<and> y \<in> cball 0 1}"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1695
  have "compact B"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1696
    unfolding B_def by (intro compact_sums compact_continuous_image cont' A) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1697
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1698
  have "uniformly_convergent_on A (\<lambda>N x. exp ((\<Sum>n<N. ln (1 + f n x))))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1699
    using conv
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1700
  proof (rule uniformly_convergent_on_compose_uniformly_continuous_on)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1701
    show "closed B"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1702
      using \<open>compact B\<close> by (auto dest: compact_imp_closed)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1703
    show "uniformly_continuous_on B exp"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1704
      by (intro compact_uniformly_continuous continuous_intros \<open>compact B\<close>)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1705
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1706
    have "eventually (\<lambda>N. \<forall>x\<in>A. dist (\<Sum>n<N. Ln (1 + f n x)) (S x) < 1) sequentially"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1707
      using S unfolding uniform_limit_iff by simp
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1708
    thus "eventually (\<lambda>N. \<forall>x\<in>A. (\<Sum>n<N. Ln (1 + f n x)) \<in> B) sequentially"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1709
    proof eventually_elim
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1710
      case (elim N)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1711
      show "\<forall>x\<in>A. (\<Sum>n<N. Ln (1 + f n x)) \<in> B"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1712
      proof safe
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1713
        fix x assume x: "x \<in> A"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1714
        have "(\<Sum>n<N. Ln (1 + f n x)) = S x + ((\<Sum>n<N. Ln (1 + f n x)) - S x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1715
          by auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1716
        moreover have "((\<Sum>n<N. Ln (1 + f n x)) - S x) \<in> ball 0 1"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1717
          using elim x by (auto simp: dist_norm norm_minus_commute)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1718
        ultimately show "(\<Sum>n<N. Ln (1 + f n x)) \<in> B"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1719
          unfolding B_def using x by fastforce
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1720
      qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1721
    qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1722
  qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1723
  also have "?this \<longleftrightarrow> uniformly_convergent_on A (\<lambda>N x. \<Prod>n<N. 1 + f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1724
  proof (intro uniformly_convergent_cong refl always_eventually allI ballI)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1725
    fix N :: nat and x assume x: "x \<in> A"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1726
    have "exp (\<Sum>n<N. ln (1 + f n x)) = (\<Prod>n<N. exp (ln (1 + f n x)))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1727
      by (simp add: exp_sum)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1728
    also have "\<dots> = (\<Prod>n<N. 1 + f n x)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1729
      using norm_f[of x] x
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1730
      by (smt (verit, best) add.right_neutral add_diff_cancel exp_Ln norm_minus_commute norm_one prod.cong)
76722
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1731
    finally show "exp (\<Sum>n<N. ln (1 + f n x)) = (\<Prod>n<N. 1 + f n x)" .
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1732
  qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1733
  finally show ?thesis .
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1734
qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1735
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1736
text \<open>Theorem 17.6 by Bak and Newman, Complex Analysis [roughly]\<close>
76722
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1737
lemma uniformly_convergent_on_prod:
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1738
  fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1739
  assumes cont: "\<And>n. continuous_on A (f n)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1740
  assumes A: "compact A"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1741
  assumes conv_sum: "uniformly_convergent_on A (\<lambda>N x. \<Sum>n<N. norm (f n x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1742
  shows   "uniformly_convergent_on A (\<lambda>N x. \<Prod>n<N. 1 + f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1743
proof -
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1744
  obtain M where M: "\<And>n x. n \<ge> M \<Longrightarrow> x \<in> A \<Longrightarrow> norm (f n x) < 1 / 2"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1745
  proof -
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1746
    from conv_sum have "uniformly_Cauchy_on A (\<lambda>N x. \<Sum>n<N. norm (f n x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1747
      using uniformly_convergent_Cauchy by blast
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1748
    moreover have "(1 / 2 :: real) > 0"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1749
      by simp
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1750
    ultimately obtain M where M:
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1751
      "\<And>x m n. x \<in> A \<Longrightarrow> m \<ge> M \<Longrightarrow> n \<ge> M \<Longrightarrow> dist (\<Sum>k<m. norm (f k x)) (\<Sum>k<n. norm (f k x)) < 1 / 2"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1752
      unfolding uniformly_Cauchy_on_def by fast
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1753
    show ?thesis
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1754
    proof (rule that[of M])
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1755
      fix n x assume nx: "n \<ge> M" "x \<in> A"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1756
      have "dist (\<Sum>k<Suc n. norm (f k x)) (\<Sum>k<n. norm (f k x)) < 1 / 2"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1757
        by (rule M) (use nx in auto)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1758
      also have "dist (\<Sum>k<Suc n. norm (f k x)) (\<Sum>k<n. norm (f k x)) = norm (f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1759
        by (simp add: dist_norm)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1760
      finally show "norm (f n x) < 1 / 2" .
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1761
    qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1762
  qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1763
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1764
  have conv: "uniformly_convergent_on A (\<lambda>N x. \<Sum>n<N. ln (1 + f (n + M) x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1765
  proof (rule uniformly_summable_comparison_test)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1766
    show "norm (ln (1 + f (n + M) x)) \<le> 2 * norm (f (n + M) x)" if "x \<in> A" for n x
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1767
      by (rule norm_Ln_le) (use M[of "n + M" x] that in auto)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1768
    have *: "filterlim (\<lambda>n. n + M) at_top at_top"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1769
      by (rule filterlim_add_const_nat_at_top)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1770
    have "uniformly_convergent_on A (\<lambda>N x. 2 * ((\<Sum>n<N+M. norm (f n x)) - (\<Sum>n<M. norm (f n x))))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1771
      by (intro uniformly_convergent_mult uniformly_convergent_minus
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1772
                uniformly_convergent_on_compose[OF conv_sum *] uniformly_convergent_on_const)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1773
    also have "(\<lambda>N x. 2 * ((\<Sum>n<N+M. norm (f n x)) - (\<Sum>n<M. norm (f n x)))) =
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1774
               (\<lambda>N x. \<Sum>n<N. 2 * norm (f (n + M) x))" (is "?lhs = ?rhs")
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1775
    proof (intro ext)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1776
      fix N x
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1777
      have "(\<Sum>n<N+M. norm (f n x)) - (\<Sum>n<M. norm (f n x)) = (\<Sum>n\<in>{..<N+M}-{..<M}. norm (f n x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1778
        by (subst sum_diff) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1779
      also have "\<dots> = (\<Sum>n<N. norm (f (n + M) x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1780
        by (intro sum.reindex_bij_witness[of _ "\<lambda>n. n + M" "\<lambda>n. n - M"]) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1781
      finally show "?lhs N x = ?rhs N x"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1782
        by (simp add: sum_distrib_left)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1783
    qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1784
    finally show "uniformly_convergent_on A (\<lambda>N x. \<Sum>n<N. 2 * cmod (f (n + M) x))" .
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1785
  qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1786
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1787
  have conv': "uniformly_convergent_on A (\<lambda>N x. \<Prod>n<N. 1 + f (n + M) x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1788
  proof (rule uniformly_convergent_on_prod_aux)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1789
    show "norm (f (n + M) x) < 1" if "x \<in> A" for n x
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1790
      using M[of "n + M" x] that by simp
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1791
  qed (use M assms conv in auto)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1792
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1793
  then obtain S where S: "uniform_limit A (\<lambda>N x. \<Prod>n<N. 1 + f (n + M) x) S sequentially"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1794
    by (auto simp: uniformly_convergent_on_def)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1795
  have cont':  "continuous_on A S"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1796
    by (intro uniform_limit_theorem[OF _ S] always_eventually ballI allI continuous_intros cont) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1797
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1798
  have "uniform_limit A (\<lambda>N x. (\<Prod>n<M. 1 + f n x) * (\<Prod>n<N. 1 + f (n + M) x)) (\<lambda>x. (\<Prod>n<M. 1 + f n x) * S x) sequentially"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1799
  proof (rule uniform_lim_mult[OF uniform_limit_const S])
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1800
    show "bounded (S ` A)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1801
      by (intro compact_imp_bounded compact_continuous_image A cont')
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1802
    show "bounded ((\<lambda>x. \<Prod>n<M. 1 + f n x) ` A)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1803
      by (intro compact_imp_bounded compact_continuous_image A continuous_intros cont)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1804
  qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1805
  hence "uniformly_convergent_on A (\<lambda>N x. (\<Prod>n<M. 1 + f n x) * (\<Prod>n<N. 1 + f (n + M) x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1806
    by (auto simp: uniformly_convergent_on_def)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1807
  also have "(\<lambda>N x. (\<Prod>n<M. 1 + f n x) * (\<Prod>n<N. 1 + f (n + M) x)) = (\<lambda>N x. (\<Prod>n<M+N. 1 + f n x))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1808
  proof (intro ext)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1809
    fix N :: nat and x :: complex
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1810
    have "(\<Prod>n<N. 1 + f (n + M) x) = (\<Prod>n\<in>{M..<M+N}. 1 + f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1811
      by (intro prod.reindex_bij_witness[of _ "\<lambda>n. n - M" "\<lambda>n. n + M"]) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1812
    also have "(\<Prod>n<M. 1 + f n x) * \<dots> = (\<Prod>n\<in>{..<M}\<union>{M..<M+N}. 1 + f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1813
      by (subst prod.union_disjoint) auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1814
    also have "{..<M} \<union> {M..<M+N} = {..<M+N}"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1815
      by auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1816
    finally show "(\<Prod>n<M. 1 + f n x) * (\<Prod>n<N. 1 + f (n + M) x) = (\<Prod>n<M+N. 1 + f n x)" .
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1817
  qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1818
  finally have "uniformly_convergent_on A (\<lambda>N x. \<Prod>n<M + N. 1 + f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1819
    by simp
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1820
  hence "uniformly_convergent_on A (\<lambda>N x. \<Prod>n<M + (N - M). 1 + f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1821
    by (rule uniformly_convergent_on_compose) (rule filterlim_minus_const_nat_at_top)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1822
  also have "?this \<longleftrightarrow> ?thesis"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1823
  proof (rule uniformly_convergent_cong)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1824
    show "eventually (\<lambda>x. \<forall>y\<in>A. (\<Prod>n<M + (x - M). 1 + f n y) = (\<Prod>n<x. 1 + f n y)) at_top"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1825
      using eventually_ge_at_top[of M] by eventually_elim auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1826
  qed auto
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1827
  finally show ?thesis .
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1828
qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1829
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1830
lemma uniformly_convergent_on_prod':
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1831
  fixes f :: "nat \<Rightarrow> complex \<Rightarrow> complex"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1832
  assumes cont: "\<And>n. continuous_on A (f n)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1833
  assumes A: "compact A"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1834
  assumes conv_sum: "uniformly_convergent_on A (\<lambda>N x. \<Sum>n<N. norm (f n x - 1))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1835
  shows "uniformly_convergent_on A (\<lambda>N x. \<Prod>n<N. f n x)"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1836
proof -
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1837
  have "uniformly_convergent_on A (\<lambda>N x. \<Prod>n<N. 1 + (f n x - 1))"
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1838
    by (rule uniformly_convergent_on_prod) (use assms in \<open>auto intro!: continuous_intros\<close>)
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1839
  thus ?thesis
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1840
    by simp
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1841
qed
b1d57dd345e1 First round of moving material from the number theory development
paulson <lp15@cam.ac.uk>
parents: 76137
diff changeset
  1842
76724
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1843
text\<open>Prop 17.6 of Bak and Newman, Complex Analysis, p. 243. 
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1844
     Only this version is for the reals. Can the two proofs be consolidated?\<close>
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1845
lemma uniformly_convergent_on_prod_real:
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1846
  fixes u :: "nat \<Rightarrow> real \<Rightarrow> real"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1847
  assumes contu: "\<And>k. continuous_on K (u k)" 
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1848
     and uconv: "uniformly_convergent_on K (\<lambda>n x. \<Sum>k<n. \<bar>u k x\<bar>)"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1849
     and K: "compact K"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1850
   shows "uniformly_convergent_on K (\<lambda>n x. \<Prod>k<n. 1 + u k x)"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1851
proof -
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1852
  define f where "f \<equiv> \<lambda>k. complex_of_real \<circ> u k \<circ> Re"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1853
  define L where "L \<equiv> complex_of_real ` K"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1854
  have "compact L"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1855
    by (simp add: \<open>compact K\<close> L_def compact_continuous_image)
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1856
  have "Re ` complex_of_real ` X = X" for X
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1857
    by (auto simp: image_iff)
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1858
  with contu have contf: "\<And>k. continuous_on L (f k)"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1859
    unfolding f_def L_def by (intro continuous_intros) auto
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1860
  obtain S where S: "\<And>\<epsilon>. \<epsilon>>0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>x\<in>K. dist (\<Sum>k<n. \<bar>u k x\<bar>) (S x) < \<epsilon>"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1861
    using uconv unfolding uniformly_convergent_on_def uniform_limit_iff by presburger 
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1862
  have "\<forall>\<^sub>F n in sequentially. \<forall>z\<in>L. dist (\<Sum>k<n. cmod (f k z)) ((of_real \<circ> S \<circ> Re) z) < \<epsilon>" 
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1863
    if "\<epsilon>>0" for \<epsilon>
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1864
    using S [OF that] by eventually_elim (simp add: L_def f_def)
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1865
  then have uconvf: "uniformly_convergent_on L (\<lambda>n z. \<Sum>k<n. norm (f k z))"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1866
    unfolding uniformly_convergent_on_def uniform_limit_iff by blast
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1867
  obtain P where P: "\<And>\<epsilon>. \<epsilon>>0 \<Longrightarrow> \<forall>\<^sub>F n in sequentially. \<forall>z\<in>L. dist (\<Prod>k<n. 1 + f k z) (P z) < \<epsilon>"
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1868
    using uniformly_convergent_on_prod [OF contf \<open>compact L\<close> uconvf]
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1869
    unfolding uniformly_convergent_on_def uniform_limit_iff by blast
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1870
  have \<section>: "\<bar>(\<Prod>k<n. 1 + u k x) - Re (P x)\<bar> \<le> cmod ((\<Prod>k<n. 1 + of_real (u k x)) - P x)" for n x
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1871
  proof -
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1872
    have "(\<Prod>k\<in>N. of_real (1 + u k x)) = (\<Prod>k\<in>N. 1 + of_real (u k x))" for N
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1873
      by force
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1874
    then show ?thesis
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1875
      by (metis Re_complex_of_real abs_Re_le_cmod minus_complex.sel(1) of_real_prod)
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1876
  qed
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1877
  have "\<forall>\<^sub>F n in sequentially. \<forall>x\<in>K. dist (\<Prod>k<n. 1 + u k x) ((Re \<circ> P \<circ> of_real) x) < \<epsilon>" 
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1878
    if "\<epsilon>>0" for \<epsilon>
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1879
    using P [OF that] by eventually_elim (simp add: L_def f_def dist_norm le_less_trans [OF \<section>])
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1880
  then show ?thesis
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1881
    unfolding uniformly_convergent_on_def uniform_limit_iff by blast 
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1882
qed
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1883
7ff71bdcf731 Additional new material about infinite products, etc.
paulson <lp15@cam.ac.uk>
parents: 76722
diff changeset
  1884
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1885
subsection\<open>The Argument of a Complex Number\<close>
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1886
73885
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1887
text\<open>Unlike in HOL Light, it's defined for the same interval as the complex logarithm: \<open>(-\<pi>,\<pi>]\<close>.\<close>
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1888
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1889
lemma Arg_eq_Im_Ln:
73924
df893af36eb4 converting arg to Arg
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
  1890
  assumes "z \<noteq> 0" shows "Arg z = Im (Ln z)"
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1891
proof (rule cis_Arg_unique)
73885
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1892
  show "sgn z = cis (Im (Ln z))"
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
  1893
    by (metis assms exp_Ln exp_eq_polar nonzero_mult_div_cancel_left norm_eq_zero
73885
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1894
              norm_exp_eq_Re of_real_eq_0_iff sgn_eq)
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1895
  show "- pi < Im (Ln z)"
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1896
    by (simp add: assms mpi_less_Im_Ln)
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1897
  show "Im (Ln z) \<le> pi"
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1898
    by (simp add: Im_Ln_le_pi assms)
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1899
qed
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1900
26171a89466a A few useful lemmas about derivatives, colinearity and other topics
paulson <lp15@cam.ac.uk>
parents: 72301
diff changeset
  1901
text \<open>The 1990s definition of argument coincides with the more recent one\<close>
73924
df893af36eb4 converting arg to Arg
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
  1902
lemma\<^marker>\<open>tag important\<close> Arg_def:
df893af36eb4 converting arg to Arg
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
  1903
  shows "Arg z = (if z = 0 then 0 else Im (Ln z))"
df893af36eb4 converting arg to Arg
paulson <lp15@cam.ac.uk>
parents: 73885
diff changeset
  1904
  by (simp add: Arg_eq_Im_Ln Arg_zero)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1905
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  1906
lemma Arg_of_real [simp]: "Arg (of_real r) = (if r<0 then pi else 0)"
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1907
  by (simp add: Im_Ln_eq_pi Arg_def)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1908
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1909
lemma mpi_less_Arg: "-pi < Arg z" and Arg_le_pi: "Arg z \<le> pi"
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1910
  by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1911
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1912
lemma Arg_eq: 
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1913
  assumes "z \<noteq> 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1914
  shows "z = of_real(norm z) * exp(\<i> * Arg z)"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1915
  using cis_conv_exp rcis_cmod_Arg rcis_def by force
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1916
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1917
lemma is_Arg_Arg: "z \<noteq> 0 \<Longrightarrow> is_Arg z (Arg z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1918
  by (simp add: Arg_eq is_Arg_def)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1919
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1920
lemma Argument_exists:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1921
  assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1922
  obtains s where "is_Arg z s" "s\<in>R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1923
proof -
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1924
  let ?rp = "r - Arg z + pi"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1925
  define k where "k \<equiv> \<lfloor>?rp / (2 * pi)\<rfloor>"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1926
  have "(Arg z + of_int k * (2 * pi)) \<in> R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1927
    using floor_divide_lower [of "2*pi" ?rp] floor_divide_upper [of "2*pi" ?rp]
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1928
    by (auto simp: k_def algebra_simps R)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1929
  then show ?thesis
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1930
    using Arg_eq \<open>z \<noteq> 0\<close> is_Arg_2pi_iff is_Arg_def that by blast
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1931
qed
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1932
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1933
lemma Argument_exists_unique:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1934
  assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1935
  obtains s where "is_Arg z s" "s\<in>R" "\<And>t. \<lbrakk>is_Arg z t; t\<in>R\<rbrakk> \<Longrightarrow> s=t"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1936
proof -
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1937
  obtain s where s: "is_Arg z s" "s\<in>R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1938
    using Argument_exists [OF assms] .
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1939
  moreover have "\<And>t. \<lbrakk>is_Arg z t; t\<in>R\<rbrakk> \<Longrightarrow> s=t"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1940
    using assms s  by (auto simp: is_Arg_eqI)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1941
  ultimately show thesis
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1942
    using that by blast
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1943
qed
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1944
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1945
lemma Argument_Ex1:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1946
  assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1947
  shows "\<exists>!s. is_Arg z s \<and> s \<in> R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1948
  using Argument_exists_unique [OF assms]  by metis
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1949
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1950
lemma Arg_divide:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1951
  assumes "w \<noteq> 0" "z \<noteq> 0"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1952
  shows "is_Arg (z / w) (Arg z - Arg w)"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1953
  using Arg_eq [of z] Arg_eq [of w] Arg_eq [of "norm(z / w)"] assms
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1954
  by (auto simp: is_Arg_def norm_divide field_simps exp_diff Arg_of_real)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1955
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1956
lemma Arg_unique_lemma:
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1957
  assumes "is_Arg z t" "is_Arg z t'"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1958
      and "- pi < t"  "t \<le> pi"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1959
      and "- pi < t'" "t' \<le> pi"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1960
      and "z \<noteq> 0"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1961
    shows "t' = t"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1962
  using is_Arg_eqI assms by force
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1963
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1964
lemma complex_norm_eq_1_exp_eq: "norm z = 1 \<longleftrightarrow> exp(\<i> * (Arg z)) = z"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1965
  by (metis Arg2pi_eq Arg_eq complex_norm_eq_1_exp norm_eq_zero norm_exp_i_times)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1966
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1967
lemma Arg_unique: "\<lbrakk>of_real r * exp(\<i> * a) = z; 0 < r; -pi < a; a \<le> pi\<rbrakk> \<Longrightarrow> Arg z = a"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1968
  by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq])
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1969
     (use mpi_less_Arg Arg_le_pi in \<open>auto simp: norm_mult\<close>)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1970
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1971
lemma Arg_minus:
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1972
  assumes "z \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1973
  shows "Arg (-z) = (if Arg z \<le> 0 then Arg z + pi else Arg z - pi)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1974
proof -
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1975
  have [simp]: "cmod z * cos (Arg z) = Re z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1976
    using assms Arg_eq [of z] by (metis Re_exp exp_Ln norm_exp_eq_Re Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1977
  have [simp]: "cmod z * sin (Arg z) = Im z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1978
    using assms Arg_eq [of z] by (metis Im_exp exp_Ln norm_exp_eq_Re Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1979
  show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1980
    using mpi_less_Arg [of z] Arg_le_pi [of z] assms
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  1981
    by (intro Arg_unique [of "norm z", OF complex_eqI]) (auto simp: Re_exp Im_exp)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1982
qed
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1983
77140
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  1984
lemma Arg_1 [simp]: "Arg 1 = 0"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  1985
  by (rule Arg_unique[of 1]) auto
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  1986
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  1987
lemma Arg_numeral [simp]: "Arg (numeral n) = 0"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  1988
  by (rule Arg_unique[of "numeral n"]) auto
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  1989
  
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1990
lemma Arg_times_of_real [simp]:
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1991
  assumes "0 < r" shows "Arg (of_real r * z) = Arg z"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1992
  using Arg_def Ln_times_of_real assms by auto
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1993
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1994
lemma Arg_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg (z * of_real r) = Arg z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1995
  by (metis Arg_times_of_real mult.commute)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1996
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1997
lemma Arg_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg (z / of_real r) = Arg z"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  1998
  by (metis Arg_times_of_real2 less_irrefl nonzero_eq_divide_eq of_real_eq_0_iff)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1999
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2000
lemma Arg_less_0: "0 \<le> Arg z \<longleftrightarrow> 0 \<le> Im z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2001
  using Im_Ln_le_pi Im_Ln_pos_le
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2002
  by (simp add: Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2003
77140
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2004
text \<open>converse fails because the argument can equal $\pi$.\<close> 
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2005
lemma Arg_uminus: "Arg z < 0 \<Longrightarrow> Arg (-z) > 0"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2006
  by (smt (verit) Arg_bounded Arg_minus Complex.Arg_def)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2007
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2008
lemma Arg_eq_pi: "Arg z = pi \<longleftrightarrow> Re z < 0 \<and> Im z = 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2009
  by (auto simp: Arg_def Im_Ln_eq_pi)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2010
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2011
lemma Arg_lt_pi: "0 < Arg z \<and> Arg z < pi \<longleftrightarrow> 0 < Im z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2012
  using Arg_less_0 [of z] Im_Ln_pos_lt
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2013
  by (auto simp: order.order_iff_strict Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2014
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2015
lemma Arg_eq_0: "Arg z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2016
  using Arg_def Im_Ln_eq_0 complex_eq_iff complex_is_Real_iff by auto
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2017
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2018
lemma Arg_neg_iff: "Arg x < 0 \<longleftrightarrow> Im x < 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2019
  using Arg_less_0 linorder_not_le by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2020
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2021
lemma Arg_pos_iff: "Arg x > 0 \<longleftrightarrow> Im x > 0 \<or> (Im x = 0 \<and> Re x < 0)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2022
  by (metis Arg_eq_pi Arg_le_pi Arg_lt_pi order_less_le pi_gt_zero)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2023
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  2024
corollary\<^marker>\<open>tag unimportant\<close> Arg_ne_0: assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0" shows "Arg z \<noteq> 0"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2025
  using assms by (auto simp: nonneg_Reals_def Arg_eq_0)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2026
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2027
lemma Arg_eq_pi_iff: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2028
  using Arg_eq_pi complex_is_Real_iff by blast
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2029
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2030
lemma Arg_eq_0_pi: "Arg z = 0 \<or> Arg z = pi \<longleftrightarrow> z \<in> \<real>"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2031
  using Arg_eq_pi_iff Arg_eq_0 by force
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2032
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2033
lemma Arg_real: "z \<in> \<real> \<Longrightarrow> Arg z = (if 0 \<le> Re z then 0 else pi)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2034
  using Arg_eq_0 Arg_eq_0_pi by auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2035
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2036
lemma Arg_inverse: "Arg(inverse z) = (if z \<in> \<real> then Arg z else - Arg z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2037
proof (cases "z \<in> \<real>")
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2038
  case False
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2039
  then show ?thesis
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2040
    by (simp add: Arg_def Ln_inverse complex_is_Real_iff complex_nonpos_Reals_iff)
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2041
qed (use Arg_real Re_inverse in auto)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2042
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2043
lemma Arg_eq_iff:
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2044
  assumes "w \<noteq> 0" "z \<noteq> 0"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2045
  shows "Arg w = Arg z \<longleftrightarrow> (\<exists>x. 0 < x \<and> w = of_real x * z)" (is "?lhs = ?rhs")
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2046
proof
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2047
  assume ?lhs
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2048
  then have "w = (cmod w / cmod z) * z"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2049
    by (metis Arg_eq assms divide_divide_eq_right eq_divide_eq exp_not_eq_zero of_real_divide)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2050
  then show ?rhs
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2051
    using assms divide_pos_pos zero_less_norm_iff by blast
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2052
qed auto
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2053
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2054
lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \<longleftrightarrow> Arg z = 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2055
  by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2056
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2057
lemma Arg_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg (cnj z) = Arg (inverse z)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2058
  using Arg2pi_cnj_eq_inverse Arg2pi_eq_iff Arg_eq_iff by auto
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2059
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2060
lemma Arg_cnj: "Arg(cnj z) = (if z \<in> \<real> then Arg z else - Arg z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2061
  by (metis Arg_cnj_eq_inverse Arg_inverse Reals_0 complex_cnj_zero)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2062
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2063
lemma Arg_exp: "-pi < Im z \<Longrightarrow> Im z \<le> pi \<Longrightarrow> Arg(exp z) = Im z"
77140
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2064
  by (simp add: Arg_eq_Im_Ln)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2065
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2066
lemma Arg_cis: "x \<in> {-pi<..pi} \<Longrightarrow> Arg (cis x) = x"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2067
  unfolding cis_conv_exp by (subst Arg_exp) auto
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2068
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2069
lemma Arg_rcis: "x \<in> {-pi<..pi} \<Longrightarrow> r > 0 \<Longrightarrow> Arg (rcis r x) = x"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2070
  unfolding rcis_def by (subst Arg_times_of_real) (auto simp: Arg_cis)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2071
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2072
lemma Ln_Arg: "z\<noteq>0 \<Longrightarrow> Ln(z) = ln(norm z) + \<i> * Arg(z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2073
  by (metis Arg_def Re_Ln complex_eq)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2074
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2075
lemma Ln_of_real':
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2076
  assumes "x \<noteq> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2077
  shows   "Ln (of_real x) = of_real (ln \<bar>x\<bar>) + (if x < 0 then pi else 0) * \<i>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2078
  by (subst Ln_Arg) (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2079
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2080
lemma continuous_at_Arg:
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2081
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2082
    shows "continuous (at z) Arg"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2083
proof -
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2084
  have "(\<lambda>z. Im (Ln z)) \<midarrow>z\<rightarrow> Arg z"
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2085
    using Arg_def assms continuous_at by fastforce
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2086
  then show ?thesis
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2087
    unfolding continuous_at
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2088
    by (smt (verit, del_insts) Arg_eq_Im_Ln Lim_transform_away_at assms nonpos_Reals_zero_I)
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2089
qed
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2090
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2091
lemma continuous_within_Arg: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within S) Arg"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  2092
  using continuous_at_Arg continuous_at_imp_continuous_within by blast
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2093
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2094
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2095
lemma continuous_on_Arg: "continuous_on (-\<real>\<^sub>\<le>\<^sub>0) Arg"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2096
  using continuous_at_Arg by (simp add: continuous_at_imp_continuous_on)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2097
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2098
lemma continuous_on_Arg' [continuous_intros]:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2099
  assumes "continuous_on A f" "\<And>z. z \<in> A \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2100
  shows   "continuous_on A (\<lambda>x. Arg (f x))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2101
  by (rule continuous_on_compose2[OF continuous_on_Arg assms(1)]) (use assms(2) in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2102
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2103
lemma tendsto_Arg [tendsto_intros]:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2104
  assumes "(f \<longlongrightarrow> z) F" "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2105
  shows   "((\<lambda>x. Arg (f x)) \<longlongrightarrow> Arg z) F"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2106
  by (rule isCont_tendsto_compose[OF continuous_at_Arg]) (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2107
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2108
lemma Arg_Re_pos: "\<bar>Arg z\<bar> < pi / 2 \<longleftrightarrow> Re z > 0 \<or> z = 0"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2109
  using Arg_def Re_Ln_pos_lt by auto
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2110
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2111
lemma Arg_Re_nonneg: "\<bar>Arg z\<bar> \<le> pi / 2 \<longleftrightarrow> Re z \<ge> 0"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2112
  using Re_Ln_pos_le[of z] by (cases "z = 0") (auto simp: Arg_eq_Im_Ln Arg_zero)
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2113
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2114
lemma Arg_times:
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2115
  assumes "Arg z + Arg w \<in> {-pi<..pi}" "z \<noteq> 0" "w \<noteq> 0"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2116
  shows   "Arg (z * w) = Arg z + Arg w"
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2117
  using Arg_eq_Im_Ln Ln_times_simple assms by auto
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2118
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2119
lemma Arg_unique': "r > 0 \<Longrightarrow> \<phi> \<in> {-pi<..pi} \<Longrightarrow> x = rcis r \<phi> \<Longrightarrow> Arg x = \<phi>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2120
  by (rule Arg_unique[of r]) (auto simp: rcis_def cis_conv_exp)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2121
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2122
lemma Arg_times':
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2123
  assumes "w \<noteq> 0" "z \<noteq> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2124
  defines "x \<equiv> Arg w + Arg z"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2125
  shows   "Arg (w * z) = x + (if x \<in> {-pi<..pi} then 0 else if x > pi then -2*pi else 2*pi)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2126
proof (rule Arg_unique'[of "norm w * norm z"])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2127
  show "w * z = rcis (cmod w * cmod z)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2128
                  (x + (if x \<in> {- pi<..pi} then 0 else if x > pi then -2*pi else 2*pi))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2129
    by (subst (1 3) rcis_cmod_Arg [symmetric])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2130
       (use assms in \<open>auto simp: rcis_def simp flip: cis_mult cis_divide cis_inverse\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2131
  show "x + (if x \<in> {- pi<..pi} then 0 else if pi < x then - 2 * pi else 2 * pi) \<in> {- pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2132
    using Arg_bounded[of w] Arg_bounded[of z] by (auto simp: x_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2133
qed (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2134
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2135
lemma Arg_divide':
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2136
  assumes [simp]: "z \<noteq> 0" "w \<noteq> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2137
  shows   "Arg (z / w) = Arg z - Arg w +
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2138
             (if Arg z - Arg w > pi then -2 * pi else if Arg z - Arg w \<le> -pi then 2 * pi else 0)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2139
    (is "_ = ?rhs")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2140
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2141
  have "Arg (z * inverse w) = ?rhs"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2142
    by (subst Arg_times') 
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2143
       (use Arg_bounded[of w] Arg_bounded[of z] 
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2144
        in  \<open>auto simp: Arg_inverse elim!: Reals_cases split: if_splits\<close>)+
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2145
  also have "z * inverse w = z / w"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2146
    by (simp add: field_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2147
  finally show ?thesis .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2148
qed
77166
0fb350e7477b More new material thanks to Manuel
paulson <lp15@cam.ac.uk>
parents: 77140
diff changeset
  2149
  
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2150
subsection\<open>The Unwinding Number and the Ln product Formula\<close>
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2151
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2152
text\<open>Note that in this special case the unwinding number is -1, 0 or 1. But it's always an integer.\<close>
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2153
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2154
lemma is_Arg_exp_Im: "is_Arg (exp z) (Im z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2155
  using exp_eq_polar is_Arg_def norm_exp_eq_Re by auto
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2156
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2157
lemma is_Arg_exp_diff_2pi:
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2158
  assumes "is_Arg (exp z) \<theta>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2159
  shows "\<exists>k. Im z - of_int k * (2 * pi) = \<theta>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2160
proof (intro exI is_Arg_eqI)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2161
  let ?k = "\<lfloor>(Im z - \<theta>) / (2 * pi)\<rfloor>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2162
  show "is_Arg (exp z) (Im z - real_of_int ?k * (2 * pi))"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2163
    by (metis diff_add_cancel is_Arg_2pi_iff is_Arg_exp_Im)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2164
  show "\<bar>Im z - real_of_int ?k * (2 * pi) - \<theta>\<bar> < 2 * pi"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2165
    using floor_divide_upper [of "2*pi" "Im z - \<theta>"] floor_divide_lower [of "2*pi" "Im z - \<theta>"]
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2166
    by (auto simp: algebra_simps abs_if)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2167
qed (auto simp: is_Arg_exp_Im assms)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2168
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2169
lemma Arg_exp_diff_2pi: "\<exists>k. Im z - of_int k * (2 * pi) = Arg (exp z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2170
  using is_Arg_exp_diff_2pi [OF is_Arg_Arg] by auto
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2171
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2172
lemma unwinding_in_Ints: "(z - Ln(exp z)) / (of_real(2*pi) * \<i>) \<in> \<int>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2173
  using Arg_exp_diff_2pi [of z]
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2174
  by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2175
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  2176
definition\<^marker>\<open>tag important\<close> unwinding :: "complex \<Rightarrow> int" where
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2177
   "unwinding z \<equiv> THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2178
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2179
lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2180
  using unwinding_in_Ints [of z]
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2181
  unfolding unwinding_def Ints_def by force
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2182
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2183
lemma unwinding_2pi: "(2*pi) * \<i> * unwinding(z) = z - Ln(exp z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2184
  by (simp add: unwinding)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2185
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2186
lemma Ln_times_unwinding:
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2187
    "w \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \<i> * unwinding(Ln w + Ln z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2188
  using unwinding_2pi by (simp add: exp_add)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2189
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2190
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2191
lemma arg_conv_arctan:
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2192
  assumes "Re z > 0"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2193
  shows   "Arg z = arctan (Im z / Re z)"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2194
proof (rule cis_Arg_unique)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2195
  show "sgn z = cis (arctan (Im z / Re z))"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2196
  proof (rule complex_eqI)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2197
    have "Re (cis (arctan (Im z / Re z))) = 1 / sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2)"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2198
      by (simp add: cos_arctan power_divide)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2199
    also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2200
      using assms by (simp add: cmod_def field_simps)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2201
    also have "1 / sqrt \<dots> = Re z / norm z"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2202
      using assms by (simp add: real_sqrt_divide)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2203
    finally show "Re (sgn z) = Re (cis (arctan (Im z / Re z)))"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2204
      by simp
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2205
  next
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2206
    have "Im (cis (arctan (Im z / Re z))) = Im z / (Re z * sqrt (1 + (Im z)\<^sup>2 / (Re z)\<^sup>2))"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2207
      by (simp add: sin_arctan field_simps)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2208
    also have "1 + Im z ^ 2 / Re z ^ 2 = norm z ^ 2 / Re z ^ 2"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2209
      using assms by (simp add: cmod_def field_simps)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2210
    also have "Im z / (Re z * sqrt \<dots>) = Im z / norm z"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2211
      using assms by (simp add: real_sqrt_divide)
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2212
    finally show "Im (sgn z) = Im (cis (arctan (Im z / Re z)))"
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2213
      by simp
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2214
  qed
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2215
next
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2216
  show "arctan (Im z / Re z) > -pi"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2217
    by (smt (verit, ccfv_SIG) arctan field_sum_of_halves)
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2218
next
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2219
 show "arctan (Im z / Re z) \<le> pi"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2220
   by (smt (verit, best) arctan field_sum_of_halves)
73928
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2221
qed
3b76524f5a85 Imported lots of material from Stirling_Formula/Gamma_Asymptotics
paulson <lp15@cam.ac.uk>
parents: 73924
diff changeset
  2222
77089
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2223
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2224
subsection \<open>Characterisation of @{term "Im (Ln z)"} (Wenda Li)\<close>
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2225
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2226
lemma Im_Ln_eq_pi_half:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2227
    "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = pi/2 \<longleftrightarrow> 0 < Im(z) \<and> Re(z) = 0)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2228
    "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = -pi/2 \<longleftrightarrow> Im(z) < 0 \<and> Re(z) = 0)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2229
  using Im_Ln_pos_lt Im_Ln_pos_le Re_Ln_pos_le Re_Ln_pos_lt pi_ge_two by fastforce+
77089
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2230
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2231
lemma Im_Ln_eq:
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2232
  assumes "z\<noteq>0"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2233
  shows "Im (Ln z) = (if Re z\<noteq>0 then
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2234
                        if Re z>0 then
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2235
                           arctan (Im z/Re z)
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2236
                        else if Im z\<ge>0 then
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2237
                           arctan (Im z/Re z) + pi
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2238
                        else
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2239
                           arctan (Im z/Re z) - pi
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2240
                      else
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2241
                        if Im z>0 then pi/2 else -pi/2)"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2242
proof -
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2243
  have eq_arctan_pos: "Im (Ln z) = arctan (Im z/Re z)" when "Re z>0" for z
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2244
    by (metis Arg_eq_Im_Ln arg_conv_arctan order_less_irrefl that zero_complex.simps(1))
77089
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2245
  have ?thesis when "Re z=0"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2246
    using Im_Ln_eq_pi_half[OF \<open>z\<noteq>0\<close>] that
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2247
    using assms complex_eq_iff by auto
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2248
  moreover have ?thesis when "Re z>0"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2249
    using eq_arctan_pos[OF that] that by auto
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2250
  moreover have ?thesis when "Re z<0" "Im z\<ge>0"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2251
  proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2252
    have "Im (Ln (- z)) = arctan (Im (- z) / Re (- z))"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2253
      by (simp add: eq_arctan_pos that(1))
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2254
    moreover have "Ln (- z) = Ln z - \<i> * complex_of_real pi"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2255
      using Ln_minus assms that by fastforce
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2256
    ultimately show ?thesis using that by auto
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2257
  qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2258
  moreover have ?thesis when "Re z<0" "Im z<0"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2259
  proof -
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2260
    have "Im (Ln (- z)) = arctan (Im (- z) / Re (- z))"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2261
      by (simp add: eq_arctan_pos that(1))
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2262
    moreover have "Ln (- z) = Ln z + \<i> * complex_of_real pi"
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2263
      using Ln_minus assms that by auto
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2264
    ultimately show ?thesis using that by auto
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2265
  qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2266
  ultimately show ?thesis by linarith
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2267
qed
b4f892d0625d Some new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 76819
diff changeset
  2268
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  2269
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation between Ln and Arg2pi, and hence continuity of Arg2pi\<close>
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2270
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2271
lemma Arg2pi_Ln: "0 < Arg2pi z \<Longrightarrow> Arg2pi z = Im(Ln(-z)) + pi"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2272
  by (smt (verit, best) Arg2pi_0 Arg2pi_exp Arg2pi_minus Arg_exp Arg_minus Im_Ln_le_pi 
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2273
      exp_Ln mpi_less_Im_Ln neg_equal_0_iff_equal)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2274
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2275
lemma continuous_at_Arg2pi:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2276
  assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2277
    shows "continuous (at z) Arg2pi"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2278
proof -
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2279
  have "isCont (\<lambda>z. Im (Ln (- z)) + pi) z"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2280
    by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2281
  moreover consider "Re z < 0" | "Im z \<noteq> 0" using assms
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2282
    using complex_nonneg_Reals_iff not_le by blast
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2283
  ultimately have "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg2pi z"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2284
    by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within)
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2285
  then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2286
    unfolding continuous_at
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2287
    by (metis (mono_tags, lifting) Arg2pi_Ln Arg2pi_gt_0 Compl_iff Lim_transform_within_open assms 
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2288
        closed_nonneg_Reals_complex open_Compl)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2289
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2290
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2291
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2292
text\<open>Relation between Arg2pi and arctangent in upper halfplane\<close>
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2293
lemma Arg2pi_arctan_upperhalf:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2294
  assumes "0 < Im z"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2295
    shows "Arg2pi z = pi/2 - arctan(Re z / Im z)"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2296
proof (cases "z = 0")
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2297
  case False
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2298
  show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2299
  proof (rule Arg2pi_unique [of "norm z"])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2300
    show "(cmod z) * exp (\<i> * (pi / 2 - arctan (Re z / Im z))) = z"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2301
      apply (rule complex_eqI)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2302
      using assms norm_complex_def [of z, symmetric]
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2303
      unfolding exp_Euler cos_diff sin_diff sin_of_real cos_of_real
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2304
      by (simp_all add: field_simps real_sqrt_divide sin_arctan cos_arctan)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2305
  qed (use False arctan [of "Re z / Im z"] in auto)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2306
qed (use assms in auto)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2307
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2308
lemma Arg2pi_eq_Im_Ln:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  2309
  assumes "0 \<le> Im z" "0 < Re z"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2310
    shows "Arg2pi z = Im (Ln z)"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2311
  by (smt (verit, ccfv_SIG) Arg2pi_exp Im_Ln_pos_le assms exp_Ln pi_neq_zero zero_complex.simps(1))
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2312
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2313
lemma continuous_within_upperhalf_Arg2pi:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2314
  assumes "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2315
    shows "continuous (at z within {z. 0 \<le> Im z}) Arg2pi"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2316
proof (cases "z \<in> \<real>\<^sub>\<ge>\<^sub>0")
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2317
  case False then show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2318
    using continuous_at_Arg2pi continuous_at_imp_continuous_within by auto
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2319
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2320
  case True
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2321
  then have z: "z \<in> \<real>" "0 < Re z"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2322
    using assms  by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2323
  then have [simp]: "Arg2pi z = 0" "Im (Ln z) = 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2324
    by (auto simp: Arg2pi_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  2325
  show ?thesis
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2326
  proof (clarsimp simp add: continuous_within Lim_within dist_norm)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2327
    fix e::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2328
    assume "0 < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2329
    moreover have "continuous (at z) (\<lambda>x. Im (Ln x))"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2330
      using z by (simp add: continuous_at_Ln complex_nonpos_Reals_iff)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2331
    ultimately
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2332
    obtain d where d: "d>0" "\<And>x. x \<noteq> z \<Longrightarrow> cmod (x - z) < d \<Longrightarrow> \<bar>Im (Ln x)\<bar> < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2333
      by (auto simp: continuous_within Lim_within dist_norm)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2334
    { fix x
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2335
      assume "cmod (x - z) < Re z / 2"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2336
      then have "\<bar>Re x - Re z\<bar> < Re z / 2"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2337
        by (metis le_less_trans abs_Re_le_cmod minus_complex.simps(1))
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2338
      then have "0 < Re x"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2339
        using z by linarith
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2340
    }
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2341
    then show "\<exists>d>0. \<forall>x. 0 \<le> Im x \<longrightarrow> x \<noteq> z \<and> cmod (x - z) < d \<longrightarrow> \<bar>Arg2pi x\<bar> < e"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2342
      apply (rule_tac x="min d (Re z / 2)" in exI)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2343
      using z d by (auto simp: Arg2pi_eq_Im_Ln)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2344
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2345
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2346
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2347
lemma continuous_on_upperhalf_Arg2pi: "continuous_on ({z. 0 \<le> Im z} - {0}) Arg2pi"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2348
  unfolding continuous_on_eq_continuous_within
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2349
  by (metis DiffE Diff_subset continuous_within_subset continuous_within_upperhalf_Arg2pi insertCI)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2350
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2351
lemma open_Arg2pi2pi_less_Int:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2352
  assumes "0 \<le> s" "t \<le> 2*pi"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2353
    shows "open ({y. s < Arg2pi y} \<inter> {y. Arg2pi y < t})"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2354
proof -
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2355
  have 1: "continuous_on (UNIV - \<real>\<^sub>\<ge>\<^sub>0) Arg2pi"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2356
    using continuous_at_Arg2pi continuous_at_imp_continuous_within
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2357
    by (auto simp: continuous_on_eq_continuous_within)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2358
  have 2: "open (UNIV - \<real>\<^sub>\<ge>\<^sub>0 :: complex set)"  by (simp add: open_Diff)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2359
  have "open ({z. s < z} \<inter> {z. z < t})"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2360
    using open_lessThan [of t] open_greaterThan [of s]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2361
    by (metis greaterThan_def lessThan_def open_Int)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2362
  moreover have "{y. s < Arg2pi y} \<inter> {y. Arg2pi y < t} \<subseteq> - \<real>\<^sub>\<ge>\<^sub>0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2363
    using assms by (auto simp: Arg2pi_real complex_nonneg_Reals_iff complex_is_Real_iff)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2364
  ultimately show ?thesis
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  2365
    using continuous_imp_open_vimage [OF 1 2, of  "{z. Re z > s} \<inter> {z. Re z < t}"]
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2366
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2367
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2368
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2369
lemma open_Arg2pi2pi_gt: "open {z. t < Arg2pi z}"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2370
proof (cases "t < 0")
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2371
  case True then have "{z. t < Arg2pi z} = UNIV"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2372
    using Arg2pi_ge_0 less_le_trans by auto
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2373
  then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2374
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2375
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2376
  case False then show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2377
    using open_Arg2pi2pi_less_Int [of t "2*pi"] Arg2pi_lt_2pi
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2378
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2379
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2380
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2381
lemma closed_Arg2pi2pi_le: "closed {z. Arg2pi z \<le> t}"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2382
  using open_Arg2pi2pi_gt [of t]
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2383
  by (simp add: closed_def Set.Collect_neg_eq [symmetric] not_le)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2384
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  2385
subsection\<^marker>\<open>tag unimportant\<close>\<open>Complex Powers\<close>
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2386
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2387
lemma powr_to_1 [simp]: "z powr 1 = (z::complex)"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2388
  by (simp add: powr_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2389
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2390
lemma powr_nat:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2391
  fixes n::nat and z::complex shows "z powr n = (if z = 0 then 0 else z^n)"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2392
  by (simp add: exp_of_nat_mult powr_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2393
77179
6d2ca97a8f46 More of Manuel's material, and some changes
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
  2394
lemma powr_nat': "(z :: complex) \<noteq> 0 \<or> n \<noteq> 0 \<Longrightarrow> z powr of_nat n = z ^ n"
6d2ca97a8f46 More of Manuel's material, and some changes
paulson <lp15@cam.ac.uk>
parents: 77166
diff changeset
  2395
  by (cases "z = 0") (auto simp: powr_nat)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 77221
diff changeset
  2396
    
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2397
lemma norm_powr_real: "w \<in> \<real> \<Longrightarrow> 0 < Re w \<Longrightarrow> norm(w powr z) = exp(Re z * ln(Re w))"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2398
  using Ln_Reals_eq norm_exp_eq_Re by (auto simp: Im_Ln_eq_0 powr_def norm_complex_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2399
77200
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
  2400
lemma norm_powr_real_powr': "w \<in> \<real> \<Longrightarrow> norm (z powr w) = norm z powr Re w"
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
  2401
  by (auto simp: powr_def Reals_def)
8f2e6186408f Some more new material and some tidying of existing proofs
paulson <lp15@cam.ac.uk>
parents: 77179
diff changeset
  2402
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2403
lemma powr_complexpow [simp]:
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2404
  fixes x::complex shows "x \<noteq> 0 \<Longrightarrow> x powr (of_nat n) = x^n"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2405
  by (simp add: powr_nat')
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2406
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2407
lemma powr_complexnumeral [simp]:
74513
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73933
diff changeset
  2408
  fixes x::complex shows "x powr (numeral n) = x ^ (numeral n)"
67d87d224e00 A few new lemmas plus some refinements
paulson <lp15@cam.ac.uk>
parents: 73933
diff changeset
  2409
  by (metis of_nat_numeral power_zero_numeral powr_nat)
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2410
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2411
lemma cnj_powr:
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2412
  assumes "Im a = 0 \<Longrightarrow> Re a \<ge> 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2413
  shows   "cnj (a powr b) = cnj a powr cnj b"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2414
proof (cases "a = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2415
  case False
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2416
  with assms have "a \<notin> \<real>\<^sub>\<le>\<^sub>0" by (auto simp: complex_eq_iff complex_nonpos_Reals_iff)
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2417
  with False show ?thesis by (simp add: powr_def exp_cnj cnj_Ln)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2418
qed simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2419
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2420
lemma powr_real_real:
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2421
  assumes "w \<in> \<real>" "z \<in> \<real>" "0 < Re w"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2422
  shows "w powr z = exp(Re z * ln(Re w))"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2423
proof -
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2424
  have "w \<noteq> 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2425
    using assms by auto
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2426
  with assms show ?thesis
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2427
    by (simp add: powr_def Ln_Reals_eq of_real_exp)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2428
qed
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2429
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2430
lemma powr_of_real:
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2431
  fixes x::real and y::real
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2432
  shows "0 \<le> x \<Longrightarrow> of_real x powr (of_real y::complex) = of_real (x powr y)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2433
  by (simp_all add: powr_def exp_eq_polar)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2434
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2435
lemma powr_of_int:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2436
  fixes z::complex and n::int
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2437
  assumes "z\<noteq>(0::complex)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2438
  shows "z powr of_int n = (if n\<ge>0 then z^nat n else inverse (z^nat (-n)))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2439
  by (metis assms not_le of_int_of_nat powr_complexpow powr_minus)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2440
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 77221
diff changeset
  2441
lemma complex_powr_of_int: "z \<noteq> 0 \<or> n \<noteq> 0 \<Longrightarrow> z powr of_int n = (z :: complex) powi n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 77221
diff changeset
  2442
  by (cases "z = 0 \<or> n = 0")
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 77221
diff changeset
  2443
     (auto simp: power_int_def powr_minus powr_nat powr_of_int power_0_left power_inverse)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 77221
diff changeset
  2444
  
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2445
lemma powr_Reals_eq: "\<lbrakk>x \<in> \<real>; y \<in> \<real>; Re x \<ge> 0\<rbrakk> \<Longrightarrow> x powr y = of_real (Re x powr Re y)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2446
  by (metis of_real_Re powr_of_real)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2447
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2448
lemma norm_powr_real_mono:
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 77221
diff changeset
  2449
    "\<lbrakk>w \<in> \<real>; 1 < Re w\<rbrakk> \<Longrightarrow> cmod(w powr z1) \<le> cmod(w powr z2) \<longleftrightarrow> Re z1 \<le> Re z2"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2450
  by (auto simp: powr_def algebra_simps Reals_def Ln_of_real)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2451
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2452
lemma powr_times_real:
79670
f471e1715fc4 A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
  2453
    "\<lbrakk>x \<in> \<real>; y \<in> \<real>; 0 \<le> Re x; 0 \<le> Re y\<rbrakk> \<Longrightarrow> (x * y) powr z = x powr z * y powr z"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2454
  by (auto simp: Reals_def powr_def Ln_times exp_add algebra_simps less_eq_real_def Ln_of_real)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2455
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2456
lemma Re_powr_le: "r \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> Re (r powr z) \<le> Re r powr Re z"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2457
  by (auto simp: powr_def nonneg_Reals_def order_trans [OF complex_Re_le_cmod])
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2458
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2459
lemma
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2460
  fixes w::complex
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2461
  assumes "w \<in> \<real>\<^sub>\<ge>\<^sub>0" "z \<in> \<real>"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2462
  shows Reals_powr [simp]: "w powr z \<in> \<real>" and nonneg_Reals_powr [simp]: "w powr z \<in> \<real>\<^sub>\<ge>\<^sub>0"
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2463
  using assms by (auto simp: nonneg_Reals_def Reals_def powr_of_real)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2464
79670
f471e1715fc4 A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
  2465
lemma exp_powr_complex:
f471e1715fc4 A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
  2466
  fixes x::complex 
f471e1715fc4 A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
  2467
  assumes "-pi < Im(x)" "Im(x) \<le> pi"
f471e1715fc4 A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
  2468
  shows "exp x powr y = exp (x*y)"
f471e1715fc4 A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
  2469
  using assms by (simp add: powr_def mult.commute)
f471e1715fc4 A small collection of new and useful facts, including the AM-GM inequality
paulson <lp15@cam.ac.uk>
parents: 78890
diff changeset
  2470
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2471
lemma powr_neg_real_complex:
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2472
  fixes w::complex
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2473
  shows "(- of_real x) powr w = (-1) powr (of_real (sgn x) * w) * of_real x powr w"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2474
proof (cases "x = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2475
  assume x: "x \<noteq> 0"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2476
  hence "(-x) powr w = exp (w * ln (-of_real x))" by (simp add: powr_def)
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2477
  also from x have "ln (-of_real x) = Ln (of_real x) + of_real (sgn x) * pi * \<i>"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2478
    by (simp add: Ln_minus Ln_of_real)
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2479
  also from x have "exp (w * \<dots>) = cis pi powr (of_real (sgn x) * w) * of_real x powr w"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2480
    by (simp add: powr_def exp_add algebra_simps Ln_of_real cis_conv_exp)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2481
  also note cis_pi
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2482
  finally show ?thesis by simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2483
qed simp_all
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2484
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2485
lemma has_field_derivative_powr:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2486
  fixes z :: complex
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2487
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2488
  shows "((\<lambda>z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2489
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2490
  case False
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2491
  then have \<section>: "exp (s * Ln z) * inverse z = exp ((s - 1) * Ln z)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2492
    by (simp add: divide_complex_def exp_diff left_diff_distrib')
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2493
  show ?thesis
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2494
    unfolding powr_def
71029
934e0044e94b Moved or deleted some out of place material, also eliminating obsolete naming conventions
paulson <lp15@cam.ac.uk>
parents: 71001
diff changeset
  2495
  proof (rule has_field_derivative_transform_within)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2496
    show "((\<lambda>z. exp (s * Ln z)) has_field_derivative s * (if z = 0 then 0 else exp ((s - 1) * Ln z)))
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2497
           (at z)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2498
      by (intro derivative_eq_intros | simp add: assms False \<section>)+
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2499
  qed (use False in auto)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2500
qed (use assms in auto)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2501
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2502
declare has_field_derivative_powr[THEN DERIV_chain2, derivative_intros]
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2503
77324
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2504
(*Seemingly impossible to use DERIV_power_int without introducing the assumption z\<in>S*)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2505
lemma has_field_derivative_powr_of_int:
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2506
  fixes z :: complex
77324
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2507
  assumes gderiv: "(g has_field_derivative gd) (at z within S)" and "g z \<noteq> 0"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2508
  shows "((\<lambda>z. g z powr of_int n) has_field_derivative (n * g z powr (of_int n - 1) * gd)) (at z within S)"
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2509
proof -
77324
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2510
  obtain e where "e>0" and e_dist: "\<forall>y\<in>S. dist z y < e \<longrightarrow> g y \<noteq> 0"
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2511
    using DERIV_continuous assms continuous_within_avoid gderiv by blast
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2512
  define D where "D = of_int n * g z powr (of_int (n - 1)) * gd"
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2513
  define E where "E = of_int n * g z powi (n - 1) * gd"
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2514
  have "((\<lambda>z. g z powr of_int n) has_field_derivative D) (at z within S)
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2515
    \<longleftrightarrow> ((\<lambda>z. g z powr of_int n) has_field_derivative E) (at z within S)"
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2516
    using assms complex_powr_of_int D_def E_def by presburger
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2517
  also have "\<dots> \<longleftrightarrow> ((\<lambda>z. g z powi n) has_field_derivative E) (at z within S)"
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2518
  proof (rule has_field_derivative_cong_eventually)
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2519
    show "\<forall>\<^sub>F x in at z within S. g x powr of_int n = g x powi n"
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2520
      unfolding eventually_at by (metis \<open>0 < e\<close> complex_powr_of_int dist_commute e_dist)
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2521
  qed (simp add: assms complex_powr_of_int)
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2522
  also have "((\<lambda>z. g z powi n) has_field_derivative E) (at z within S)"
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2523
    unfolding E_def using gderiv assms by (auto intro!: derivative_eq_intros)
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2524
  finally show ?thesis
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  2525
    by (simp add: D_def) 
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2526
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2527
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2528
lemma field_differentiable_powr_of_int:
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2529
  fixes z :: complex
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2530
  assumes "g field_differentiable (at z within S)" and "g z \<noteq> 0"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2531
  shows "(\<lambda>z. g z powr of_int n) field_differentiable (at z within S)"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2532
  using has_field_derivative_powr_of_int assms field_differentiable_def by blast
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2533
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2534
lemma holomorphic_on_powr_of_int [holomorphic_intros]:
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2535
  assumes "f holomorphic_on S" and "\<And>z. z\<in>S \<Longrightarrow> f z \<noteq> 0"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2536
  shows "(\<lambda>z. (f z) powr of_int n) holomorphic_on S"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2537
  using assms field_differentiable_powr_of_int holomorphic_on_def by auto
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2538
65578
e4997c181cce New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents: 65274
diff changeset
  2539
lemma has_field_derivative_powr_right [derivative_intros]:
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2540
    "w \<noteq> 0 \<Longrightarrow> ((\<lambda>z. w powr z) has_field_derivative Ln w * w powr z) (at z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2541
  unfolding powr_def by (intro derivative_eq_intros | simp)+
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2542
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2543
lemma field_differentiable_powr_right [derivative_intros]:
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62393
diff changeset
  2544
  fixes w::complex
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2545
  shows "w \<noteq> 0 \<Longrightarrow> (\<lambda>z. w powr z) field_differentiable (at z)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2546
using field_differentiable_def has_field_derivative_powr_right by blast
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2547
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2548
lemma holomorphic_on_powr_right [holomorphic_intros]:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2549
  assumes "f holomorphic_on S"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2550
  shows "(\<lambda>z. w powr (f z)) holomorphic_on S"
67268
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2551
proof (cases "w = 0")
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2552
  case False
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2553
  with assms show ?thesis
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2554
    unfolding holomorphic_on_def field_differentiable_def
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2555
    by (metis (full_types) DERIV_chain' has_field_derivative_powr_right)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2556
qed simp
67268
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2557
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2558
lemma holomorphic_on_divide_gen [holomorphic_intros]:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2559
  assumes "f holomorphic_on S" "g holomorphic_on S" and "\<And>z z'. \<lbrakk>z \<in> S; z' \<in> S\<rbrakk> \<Longrightarrow> g z = 0 \<longleftrightarrow> g z' = 0"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2560
  shows "(\<lambda>z. f z / g z) holomorphic_on S"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2561
  by (metis (no_types, lifting) assms division_ring_divide_zero holomorphic_on_divide holomorphic_transform)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2562
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2563
lemma norm_powr_real_powr:
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2564
  "w \<in> \<real> \<Longrightarrow> 0 \<le> Re w \<Longrightarrow> cmod (w powr z) = Re w powr Re z"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2565
  by (metis dual_order.order_iff_strict norm_powr_real norm_zero of_real_0 of_real_Re powr_def)
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2566
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2567
lemma tendsto_powr_complex:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2568
  fixes f g :: "_ \<Rightarrow> complex"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2569
  assumes a: "a \<notin> \<real>\<^sub>\<le>\<^sub>0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2570
  assumes f: "(f \<longlongrightarrow> a) F" and g: "(g \<longlongrightarrow> b) F"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2571
  shows   "((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2572
proof -
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2573
  from a have [simp]: "a \<noteq> 0" by auto
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2574
  from f g a have "((\<lambda>z. exp (g z * ln (f z))) \<longlongrightarrow> a powr b) F" (is ?P)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2575
    by (auto intro!: tendsto_intros simp: powr_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2576
  also {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2577
    have "eventually (\<lambda>z. z \<noteq> 0) (nhds a)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2578
      by (intro t1_space_nhds) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2579
    with f have "eventually (\<lambda>z. f z \<noteq> 0) F" using filterlim_iff by blast
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2580
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2581
  hence "?P \<longleftrightarrow> ((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2582
    by (intro tendsto_cong refl) (simp_all add: powr_def mult_ac)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2583
  finally show ?thesis .
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2584
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2585
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2586
lemma tendsto_powr_complex_0:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2587
  fixes f g :: "'a \<Rightarrow> complex"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2588
  assumes f: "(f \<longlongrightarrow> 0) F" and g: "(g \<longlongrightarrow> b) F" and b: "Re b > 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2589
  shows   "((\<lambda>z. f z powr g z) \<longlongrightarrow> 0) F"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2590
proof (rule tendsto_norm_zero_cancel)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2591
  define h where
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2592
    "h = (\<lambda>z. if f z = 0 then 0 else exp (Re (g z) * ln (cmod (f z)) + abs (Im (g z)) * pi))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2593
  {
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2594
    fix z :: 'a assume z: "f z \<noteq> 0" 
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2595
    define c where "c = abs (Im (g z)) * pi"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2596
    from mpi_less_Im_Ln[OF z] Im_Ln_le_pi[OF z]
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2597
      have "abs (Im (Ln (f z))) \<le> pi" by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2598
    from mult_left_mono[OF this, of "abs (Im (g z))"]
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2599
      have "abs (Im (g z) * Im (ln (f z))) \<le> c" by (simp add: abs_mult c_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2600
    hence "-Im (g z) * Im (ln (f z)) \<le> c" by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2601
    hence "norm (f z powr g z) \<le> h z" by (simp add: powr_def field_simps h_def c_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2602
  }
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2603
  hence le: "norm (f z powr g z) \<le> h z" for z
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2604
    by (simp add: h_def) 
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2605
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2606
  have g': "(g \<longlongrightarrow> b) (inf F (principal {z. f z \<noteq> 0}))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2607
    by (rule tendsto_mono[OF _ g]) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2608
  have "((\<lambda>x. norm (f x)) \<longlongrightarrow> 0) (inf F (principal {z. f z \<noteq> 0}))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2609
    by (subst tendsto_norm_zero_iff, rule tendsto_mono[OF _ f]) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2610
  moreover {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2611
    have "filterlim (\<lambda>x. norm (f x)) (principal {0<..}) (principal {z. f z \<noteq> 0})"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2612
      by (auto simp: filterlim_def)
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2613
    hence "filterlim (\<lambda>x. norm (f x)) (principal {0<..}) (inf F (principal {z. f z \<noteq> 0}))"
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2614
      by (rule filterlim_mono) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2615
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2616
  ultimately have norm: "filterlim (\<lambda>x. norm (f x)) (at_right 0) (inf F (principal {z. f z \<noteq> 0}))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2617
    by (simp add: filterlim_inf at_within_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2618
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2619
  have A: "LIM x inf F (principal {z. f z \<noteq> 0}). Re (g x) * -ln (cmod (f x)) :> at_top"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2620
    by (rule filterlim_tendsto_pos_mult_at_top tendsto_intros g' b
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2621
          filterlim_compose[OF filterlim_uminus_at_top_at_bot] filterlim_compose[OF ln_at_0] norm)+
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2622
  have B: "LIM x inf F (principal {z. f z \<noteq> 0}).
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2623
          -\<bar>Im (g x)\<bar> * pi + -(Re (g x) * ln (cmod (f x))) :> at_top"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2624
    by (rule filterlim_tendsto_add_at_top tendsto_intros g')+ (insert A, simp_all)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2625
  have C: "(h \<longlongrightarrow> 0) F" unfolding h_def
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2626
    by (intro filterlim_If tendsto_const filterlim_compose[OF exp_at_bot])
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2627
       (insert B, auto simp: filterlim_uminus_at_bot algebra_simps)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2628
  show "((\<lambda>x. norm (f x powr g x)) \<longlongrightarrow> 0) F"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2629
    by (rule Lim_null_comparison[OF always_eventually C]) (insert le, auto)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2630
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2631
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2632
lemma tendsto_powr_complex' [tendsto_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2633
  fixes f g :: "_ \<Rightarrow> complex"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2634
  assumes "a \<notin> \<real>\<^sub>\<le>\<^sub>0 \<or> (a = 0 \<and> Re b > 0)" and "(f \<longlongrightarrow> a) F" "(g \<longlongrightarrow> b) F"
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2635
  shows   "((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2636
  using assms tendsto_powr_complex tendsto_powr_complex_0 by fastforce
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2637
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2638
lemma tendsto_neg_powr_complex_of_real:
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2639
  assumes "filterlim f at_top F" and "Re s < 0"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2640
  shows   "((\<lambda>x. complex_of_real (f x) powr s) \<longlongrightarrow> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2641
proof -
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2642
  have "((\<lambda>x. norm (complex_of_real (f x) powr s)) \<longlongrightarrow> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2643
  proof (rule Lim_transform_eventually)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2644
    from assms(1) have "eventually (\<lambda>x. f x \<ge> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2645
      by (auto simp: filterlim_at_top)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2646
    thus "eventually (\<lambda>x. f x powr Re s = norm (of_real (f x) powr s)) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2647
      by eventually_elim (simp add: norm_powr_real_powr)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2648
    from assms show "((\<lambda>x. f x powr Re s) \<longlongrightarrow> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2649
      by (intro tendsto_neg_powr)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2650
  qed
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2651
  thus ?thesis by (simp add: tendsto_norm_zero_iff)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2652
qed
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2653
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2654
lemma tendsto_neg_powr_complex_of_nat:
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2655
  assumes "filterlim f at_top F" and "Re s < 0"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2656
  shows   "((\<lambda>x. of_nat (f x) powr s) \<longlongrightarrow> 0) F"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2657
  using tendsto_neg_powr_complex_of_real [of "real o f" F s]
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2658
proof -
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2659
  have "((\<lambda>x. of_real (real (f x)) powr s) \<longlongrightarrow> 0) F" using assms(2)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2660
    by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real]
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2661
              filterlim_compose[OF _ assms(1)] filterlim_real_sequentially filterlim_ident) auto
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2662
  thus ?thesis by simp
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2663
qed
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2664
82338
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2665
lemma continuous_powr_complex [continuous_intros]: 
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2666
  assumes "continuous F f" "continuous F g"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2667
  assumes "Re (f (netlimit F)) \<ge> 0 \<or> Im (f (netlimit F)) \<noteq> 0"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2668
  assumes "f (netlimit F) = 0 \<longrightarrow> Re (g (netlimit F)) > 0"
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2669
  shows   "continuous F (\<lambda>z. f z powr g z :: complex)"
82338
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2670
  using assms
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2671
  unfolding continuous_def
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2672
  by (intro tendsto_powr_complex')
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2673
     (auto simp: complex_nonpos_Reals_iff complex_eq_iff)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2674
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2675
lemma continuous_powr_real [continuous_intros]: 
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2676
  assumes "continuous F f" "continuous F g"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2677
  assumes "f (netlimit F) = 0 \<longrightarrow> g (netlimit F) > 0 \<and> (\<forall>\<^sub>F z in F. 0 \<le> f z)"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2678
  shows   "continuous F (\<lambda>z. f z powr g z :: real)"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2679
  using assms unfolding continuous_def by (intro tendsto_intros) auto
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2680
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2681
lemma isCont_powr_complex [continuous_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2682
  assumes "f z \<notin> \<real>\<^sub>\<le>\<^sub>0" "isCont f z" "isCont g z"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2683
  shows   "isCont (\<lambda>z. f z powr g z :: complex) z"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2684
  using assms unfolding isCont_def by (intro tendsto_powr_complex) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2685
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2686
lemma continuous_on_powr_complex [continuous_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2687
  assumes "A \<subseteq> {z. Re (f z) \<ge> 0 \<or> Im (f z) \<noteq> 0}"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2688
  assumes "\<And>z. z \<in> A \<Longrightarrow> f z = 0 \<Longrightarrow> Re (g z) > 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2689
  assumes "continuous_on A f" "continuous_on A g"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2690
  shows   "continuous_on A (\<lambda>z. f z powr g z)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2691
  unfolding continuous_on_def
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2692
proof
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2693
  fix z assume z: "z \<in> A"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2694
  show "((\<lambda>z. f z powr g z) \<longlongrightarrow> f z powr g z) (at z within A)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2695
  proof (cases "f z = 0")
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2696
    case False
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2697
    from assms(1,2) z have "Re (f z) \<ge> 0 \<or> Im (f z) \<noteq> 0" "f z = 0 \<longrightarrow> Re (g z) > 0" by auto
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2698
    with assms(3,4) z show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2699
      by (intro tendsto_powr_complex')
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2700
         (auto elim!: nonpos_Reals_cases simp: complex_eq_iff continuous_on_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2701
  next
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2702
    case True
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2703
    with assms z show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2704
      by (auto intro!: tendsto_powr_complex_0 simp: continuous_on_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2705
  qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2706
qed
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2707
82338
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2708
lemma analytic_on_powr [analytic_intros]:
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2709
  assumes "f analytic_on X" "g analytic_on X" "\<And>x. x \<in> X \<Longrightarrow> f x \<notin> \<real>\<^sub>\<le>\<^sub>0"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2710
  shows   "(\<lambda>x. f x powr g x) analytic_on X"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2711
proof -
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2712
  from assms(1) obtain X1 where X1: "open X1" "X \<subseteq> X1" "f analytic_on X1"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2713
    unfolding analytic_on_holomorphic by blast
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2714
  from assms(2) obtain X2 where X2: "open X2" "X \<subseteq> X2" "g analytic_on X2"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2715
    unfolding analytic_on_holomorphic by blast
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2716
  have X: "open (X2 \<inter> (X1 \<inter> f -` (-\<real>\<^sub>\<le>\<^sub>0)))"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2717
    by (rule open_Int[OF _ continuous_open_preimage])
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2718
       (use X1 X2 in \<open>auto intro!: holomorphic_on_imp_continuous_on analytic_imp_holomorphic\<close>)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2719
  have X': "X \<subseteq> X2 \<inter> (X1 \<inter> f -` (-\<real>\<^sub>\<le>\<^sub>0))"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2720
    using assms(3) X1(2) X2(2) by blast
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2721
  note [holomorphic_intros] =
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2722
    analytic_imp_holomorphic[OF analytic_on_subset[OF X1(3)]]
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2723
    analytic_imp_holomorphic[OF analytic_on_subset[OF X2(3)]]
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2724
  have "(\<lambda>x. exp (ln (f x) * g x)) holomorphic_on (X2 \<inter> (X1 \<inter> f -` (-\<real>\<^sub>\<le>\<^sub>0)))"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2725
    by (intro holomorphic_intros) auto
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2726
  also have "?this \<longleftrightarrow> (\<lambda>x. f x powr g x) holomorphic_on (X2 \<inter> (X1 \<inter> f -` (-\<real>\<^sub>\<le>\<^sub>0)))"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2727
    by (intro holomorphic_cong) (auto simp: powr_def mult.commute)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2728
  finally show ?thesis
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2729
    using X X' unfolding analytic_on_holomorphic by blast
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2730
qed
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2731
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2732
lemma holomorphic_on_powr [holomorphic_intros]:
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2733
  assumes "f holomorphic_on X" "g holomorphic_on X" "\<And>x. x \<in> X \<Longrightarrow> f x \<notin> \<real>\<^sub>\<le>\<^sub>0"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2734
  shows   "(\<lambda>x. f x powr g x) holomorphic_on X"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2735
proof -
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2736
  have [simp]: "f x \<noteq> 0" if "x \<in> X" for x
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2737
    using assms(3)[OF that] by auto
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2738
  have "(\<lambda>x. exp (ln (f x) * g x)) holomorphic_on X"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2739
    by (auto intro!: holomorphic_intros assms(1,2)) (use assms(3) in auto)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2740
  also have "?this \<longleftrightarrow> ?thesis"
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2741
    by (intro holomorphic_cong) (use assms(3) in \<open>auto simp: powr_def mult_ac\<close>)
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2742
  finally show ?thesis .
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2743
qed
1eb12389c499 New material by Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 80241
diff changeset
  2744
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  2745
subsection\<^marker>\<open>tag unimportant\<close>\<open>Some Limits involving Logarithms\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  2746
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2747
lemma lim_Ln_over_power:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2748
  fixes s::complex
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2749
  assumes "0 < Re s"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2750
    shows "(\<lambda>n. Ln (of_nat n) / of_nat n powr s) \<longlonglongrightarrow> 0"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2751
proof (simp add: lim_sequentially dist_norm, clarify)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  2752
  fix e::real
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2753
  assume e: "0 < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2754
  have "\<exists>xo>0. \<forall>x\<ge>xo. 0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2755
  proof (rule_tac x="2/(e * (Re s)\<^sup>2)" in exI, safe)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2756
    show "0 < 2 / (e * (Re s)\<^sup>2)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2757
      using e assms by (simp add: field_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2758
  next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2759
    fix x::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2760
    assume x: "2 / (e * (Re s)\<^sup>2) \<le> x"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2761
    have "2 / (e * (Re s)\<^sup>2) > 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2762
      using e assms by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2763
    with x have "x > 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2764
      by linarith
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2765
    then have "x * 2 \<le> e * (x\<^sup>2 * (Re s)\<^sup>2)"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2766
      using e assms x by (auto simp: power2_eq_square field_simps)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  2767
    also have "\<dots> < e * (2 + (x * (Re s * 2) + x\<^sup>2 * (Re s)\<^sup>2))"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2768
      using e assms \<open>x > 0\<close>
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2769
      by (auto simp: power2_eq_square field_simps add_pos_pos)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2770
    finally show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2771
      by (auto simp: algebra_simps)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2772
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2773
  then have "\<exists>xo>0. \<forall>x\<ge>xo. x / e < 1 + (Re s * x) + (1/2) * (Re s * x)^2"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2774
    using e  by (simp add: field_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2775
  then have "\<exists>xo>0. \<forall>x\<ge>xo. x / e < exp (Re s * x)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2776
    using assms
69529
4ab9657b3257 capitalize proper names in lemma names
nipkow
parents: 69508
diff changeset
  2777
    by (force intro: less_le_trans [OF _ exp_lower_Taylor_quadratic])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2778
  then obtain xo where "xo > 0" and xo: "\<And>x. x \<ge> xo \<Longrightarrow> x < e * exp (Re s * x)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2779
    using e  by (auto simp: field_simps)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2780
  have "norm (Ln (of_nat n) / of_nat n powr s) < e" if "n \<ge> nat \<lceil>exp xo\<rceil>" for n
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2781
  proof -
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2782
    have "ln (real n) \<ge> xo"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2783
      using that exp_gt_zero ln_ge_iff [of n] nat_ceiling_le_eq by fastforce
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2784
    then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2785
      using e xo [of "ln n"]  by (auto simp: norm_divide norm_powr_real field_split_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2786
  qed
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2787
  then show "\<exists>no. \<forall>n\<ge>no. norm (Ln (of_nat n) / of_nat n powr s) < e"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2788
    by blast
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2789
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2790
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2791
lemma lim_Ln_over_n: "((\<lambda>n. Ln(of_nat n) / of_nat n) \<longlongrightarrow> 0) sequentially"
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65585
diff changeset
  2792
  using lim_Ln_over_power [of 1] by simp
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65585
diff changeset
  2793
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2794
lemma lim_ln_over_power:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2795
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2796
  assumes "0 < s"
77273
f82317de6f28 A bit more tidying and some new material
paulson <lp15@cam.ac.uk>
parents: 77230
diff changeset
  2797
  shows "(\<lambda>n. ln (real n) / real n powr s) \<longlonglongrightarrow> 0"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2798
proof -
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2799
  have "(\<lambda>n. ln (Suc n) / (Suc n) powr s) \<longlonglongrightarrow> 0"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2800
    using lim_Ln_over_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2801
    by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2802
  then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2803
    using filterlim_sequentially_Suc[of "\<lambda>n::nat. ln n / n powr s"] by auto
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2804
qed
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2805
70724
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2806
lemma lim_ln_over_n [tendsto_intros]: "((\<lambda>n. ln(real_of_nat n) / of_nat n) \<longlongrightarrow> 0) sequentially"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2807
  using lim_ln_over_power [of 1] by auto
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2808
70724
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2809
lemma lim_log_over_n [tendsto_intros]:
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2810
  "(\<lambda>n. log k n/n) \<longlonglongrightarrow> 0"
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2811
proof -
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2812
  have *: "log k n/n = (1/ln k) * (ln n / n)" for n
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2813
    unfolding log_def by auto
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2814
  have "(\<lambda>n. (1/ln k) * (ln n / n)) \<longlonglongrightarrow> (1/ln k) * 0"
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2815
    by (intro tendsto_intros)
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2816
  then show ?thesis
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2817
    unfolding * by auto
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2818
qed
65371451fde8 A few more simple results
paulson <lp15@cam.ac.uk>
parents: 70367
diff changeset
  2819
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2820
lemma lim_1_over_complex_power:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2821
  assumes "0 < Re s"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2822
  shows "(\<lambda>n. 1 / of_nat n powr s) \<longlonglongrightarrow> 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2823
proof (rule Lim_null_comparison)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2824
  have "\<forall>n>0. 3 \<le> n \<longrightarrow> 1 \<le> ln (real_of_nat n)"
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2825
    using ln_272_gt_1
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2826
    by (force intro: order_trans [of _ "ln (272/100)"])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2827
  then show "\<forall>\<^sub>F x in sequentially. cmod (1 / of_nat x powr s) \<le> cmod (Ln (of_nat x) / of_nat x powr s)"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  2828
    by (auto simp: norm_divide field_split_simps eventually_sequentially)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2829
  show "(\<lambda>n. cmod (Ln (of_nat n) / of_nat n powr s)) \<longlonglongrightarrow> 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2830
    using lim_Ln_over_power [OF assms] by (metis tendsto_norm_zero_iff)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2831
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2832
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2833
lemma lim_1_over_real_power:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2834
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2835
  assumes "0 < s"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2836
  shows "((\<lambda>n. 1 / (of_nat n powr s)) \<longlongrightarrow> 0) sequentially"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2837
  using lim_1_over_complex_power [of "of_real s", THEN filterlim_sequentially_Suc [THEN iffD2]] assms
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2838
  apply (subst filterlim_sequentially_Suc [symmetric])
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2839
  by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2840
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  2841
lemma lim_1_over_Ln: "(\<lambda>n. 1 / Ln (complex_of_nat n)) \<longlonglongrightarrow> 0"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  2842
proof (clarsimp simp add: lim_sequentially dist_norm norm_divide field_split_simps)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2843
  fix r::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2844
  assume "0 < r"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2845
  have ir: "inverse (exp (inverse r)) > 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2846
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2847
  obtain n where n: "1 < of_nat n * inverse (exp (inverse r))"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2848
    using ex_less_of_nat_mult [of _ 1, OF ir]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2849
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2850
  then have "exp (inverse r) < of_nat n"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  2851
    by (simp add: field_split_simps)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2852
  then have "ln (exp (inverse r)) < ln (of_nat n)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2853
    by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2854
  with \<open>0 < r\<close> have "1 < r * ln (real_of_nat n)"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2855
    by (simp add: field_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2856
  moreover have "n > 0" using n
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2857
    using neq0_conv by fastforce
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2858
  ultimately show "\<exists>no. \<forall>k. Ln (of_nat k) \<noteq> 0 \<longrightarrow> no \<le> k \<longrightarrow> 1 < r * cmod (Ln (of_nat k))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2859
    using n \<open>0 < r\<close>
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  2860
    by (rule_tac x=n in exI) (force simp: field_split_simps intro: less_le_trans)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2861
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2862
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  2863
lemma lim_1_over_ln: "(\<lambda>n. 1 / ln (real n)) \<longlonglongrightarrow> 0"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  2864
  using lim_1_over_Ln [THEN filterlim_sequentially_Suc [THEN iffD2]]
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2865
  apply (subst filterlim_sequentially_Suc [symmetric])
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2866
  by (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2867
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2868
lemma lim_ln1_over_ln: "(\<lambda>n. ln(Suc n) / ln n) \<longlonglongrightarrow> 1"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2869
proof (rule Lim_transform_eventually)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2870
  have "(\<lambda>n. ln(1 + 1/n) / ln n) \<longlonglongrightarrow> 0"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2871
  proof (rule Lim_transform_bound)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2872
    show "(inverse o real) \<longlonglongrightarrow> 0"
70367
81b65ddac59f fixed renaming issues
paulson <lp15@cam.ac.uk>
parents: 70365
diff changeset
  2873
      by (metis comp_def lim_inverse_n lim_explicit)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2874
    show "\<forall>\<^sub>F n in sequentially. norm (ln (1 + 1 / n) / ln n) \<le> norm ((inverse \<circ> real) n)"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2875
    proof
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2876
      fix n::nat
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2877
      assume n: "3 \<le> n"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2878
      then have "ln 3 \<le> ln n" and ln0: "0 \<le> ln n"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2879
        by auto
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2880
      with ln3_gt_1 have "1/ ln n \<le> 1"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  2881
        by (simp add: field_split_simps)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2882
      moreover have "ln (1 + 1 / real n) \<le> 1/n"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2883
        by (simp add: ln_add_one_self_le_self)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2884
      ultimately have "ln (1 + 1 / real n) * (1 / ln n) \<le> (1/n) * 1"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2885
        by (intro mult_mono) (use n in auto)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2886
      then show "norm (ln (1 + 1 / n) / ln n) \<le> norm ((inverse \<circ> real) n)"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2887
        by (simp add: field_simps ln0)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2888
      qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2889
  qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2890
  then show "(\<lambda>n. 1 + ln(1 + 1/n) / ln n) \<longlonglongrightarrow> 1"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2891
    by (metis (full_types) add.right_neutral tendsto_add_const_iff)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2892
  show "\<forall>\<^sub>F k in sequentially. 1 + ln (1 + 1 / k) / ln k = ln(Suc k) / ln k"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  2893
    by (simp add: field_split_simps ln_div eventually_sequentiallyI [of 2])
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2894
qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2895
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2896
lemma lim_ln_over_ln1: "(\<lambda>n. ln n / ln(Suc n)) \<longlonglongrightarrow> 1"
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  2897
  using tendsto_inverse [OF lim_ln1_over_ln] by force
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2898
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2899
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  2900
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relation between Square Root and exp/ln, hence its derivative\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2901
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2902
lemma csqrt_exp_Ln:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2903
  assumes "z \<noteq> 0"
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  2904
    shows "csqrt z = exp(Ln z / 2)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2905
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2906
  have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))"
64240
eabf80376aab more standardized names
haftmann
parents: 63918
diff changeset
  2907
    by (metis exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  2908
  also have "\<dots> = z"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2909
    using assms exp_Ln by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2910
  finally have "csqrt z = csqrt ((exp (Ln z / 2))\<^sup>2)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2911
    by simp
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  2912
  also have "\<dots> = exp (Ln z / 2)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  2913
    apply (rule csqrt_square)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2914
    using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2915
    by (fastforce simp: Re_exp Im_exp)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2916
  finally show ?thesis using assms csqrt_square
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2917
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2918
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2919
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2920
lemma csqrt_in_nonpos_Reals_iff [simp]: "csqrt z \<in> \<real>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> z = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2921
proof
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2922
  assume "csqrt z \<in> \<real>\<^sub>\<le>\<^sub>0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2923
  hence "csqrt z = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2924
    unfolding complex_eq_iff using csqrt_principal[of z]
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2925
    by (auto simp: complex_nonpos_Reals_iff sgn_if simp del: csqrt.sel split: if_splits)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2926
  thus "z = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2927
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2928
qed auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2929
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2930
lemma Im_csqrt_eq_0_iff: "Im (csqrt z) = 0 \<longleftrightarrow> z \<in> \<real>\<^sub>\<ge>\<^sub>0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2931
proof
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2932
  assume *: "Im (csqrt z) = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2933
  define x where "x = Re (csqrt z)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2934
  have "z = csqrt z ^ 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2935
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2936
  also have "csqrt z = of_real x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2937
    by (simp add: complex_eq_iff x_def * del: csqrt.sel)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2938
  also have "\<dots> ^ 2 = of_real (x ^ 2)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2939
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2940
  also have "\<dots> \<in> \<real>\<^sub>\<ge>\<^sub>0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2941
    unfolding nonneg_Reals_of_real_iff by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2942
  finally show "z \<in> \<real>\<^sub>\<ge>\<^sub>0" .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2943
qed (auto elim!: nonneg_Reals_cases)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  2944
77221
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  2945
lemma csqrt_conv_powr: "csqrt z = z powr (1/2)"
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  2946
  by (auto simp: csqrt_exp_Ln powr_def)
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  2947
77140
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2948
lemma csqrt_mult:
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2949
  assumes "Arg z + Arg w \<in> {-pi<..pi}"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2950
  shows   "csqrt (z * w) = csqrt z * csqrt w"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2951
proof (cases "z = 0 \<or> w = 0")
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2952
  case False
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2953
  have "csqrt (z * w) = exp ((ln (z * w)) / 2)"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2954
    using False by (intro csqrt_exp_Ln) auto
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2955
  also have "\<dots> = exp ((Ln z + Ln w) / 2)"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2956
    using False assms by (subst Ln_times_simple) (auto simp: Arg_eq_Im_Ln)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2957
  also have "(Ln z + Ln w) / 2 = Ln z / 2 + Ln w / 2"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2958
    by (simp add: add_divide_distrib)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2959
  also have "exp \<dots> = csqrt z * csqrt w"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2960
    using False by (simp add: exp_add csqrt_exp_Ln)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2961
  finally show ?thesis .
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2962
qed auto
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2963
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2964
lemma Arg_csqrt [simp]: "Arg (csqrt z) = Arg z / 2"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2965
proof (cases "z = 0")
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2966
  case False
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2967
  have "Im (Ln z) \<in> {-pi<..pi}"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2968
    by (simp add: False Im_Ln_le_pi mpi_less_Im_Ln)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2969
  also have "\<dots> \<subseteq> {-2*pi<..2*pi}"
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2970
    by auto
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2971
  finally show ?thesis
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2972
    using False by (auto simp: csqrt_exp_Ln Arg_exp Arg_eq_Im_Ln)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2973
qed (auto simp: Arg_zero)
9a60c1759543 Lots more new material thanks to Manuel Eberl
paulson <lp15@cam.ac.uk>
parents: 77103
diff changeset
  2974
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2975
lemma csqrt_inverse:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2976
  "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> csqrt (inverse z) = inverse (csqrt z)"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2977
  by (metis Ln_inverse csqrt_eq_0 csqrt_exp_Ln divide_minus_left exp_minus 
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2978
      inverse_nonzero_iff_nonzero)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2979
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2980
lemma cnj_csqrt: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> cnj(csqrt z) = csqrt(cnj z)"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  2981
  by (metis cnj_Ln complex_cnj_divide complex_cnj_numeral complex_cnj_zero_iff csqrt_eq_0 csqrt_exp_Ln exp_cnj)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2982
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2983
lemma has_field_derivative_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2984
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2985
    shows "(csqrt has_field_derivative inverse(2 * csqrt z)) (at z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2986
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2987
  have z: "z \<noteq> 0"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2988
    using assms by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2989
  then have *: "inverse z = inverse (2*z) * 2"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  2990
    by (simp add: field_split_simps)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2991
  have [simp]: "exp (Ln z / 2) * inverse z = inverse (csqrt z)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2992
    by (simp add: z field_simps csqrt_exp_Ln [symmetric]) (metis power2_csqrt power2_eq_square)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2993
  have "Im z = 0 \<Longrightarrow> 0 < Re z"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2994
    using assms complex_nonpos_Reals_iff not_less by blast
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2995
  with z have "((\<lambda>z. exp (Ln z / 2)) has_field_derivative inverse (2 * csqrt z)) (at z)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2996
    by (force intro: derivative_eq_intros * simp add: assms)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2997
  then show ?thesis
71029
934e0044e94b Moved or deleted some out of place material, also eliminating obsolete naming conventions
paulson <lp15@cam.ac.uk>
parents: 71001
diff changeset
  2998
  proof (rule has_field_derivative_transform_within)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2999
    show "\<And>x. dist x z < cmod z \<Longrightarrow> exp (Ln x / 2) = csqrt x"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  3000
      by (metis csqrt_exp_Ln dist_0_norm less_irrefl)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  3001
  qed (use z in auto)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3002
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3003
82459
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3004
lemma has_field_derivative_csqrt' [derivative_intros]:
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3005
  assumes "(f has_field_derivative f') (at x within A)" "f x \<notin> \<real>\<^sub>\<le>\<^sub>0"
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3006
  shows   "((\<lambda>x. csqrt (f x)) has_field_derivative (f' / (2 * csqrt (f x)))) (at x within A)"
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3007
proof -
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3008
  have "((csqrt \<circ> f) has_field_derivative (inverse (2 * csqrt (f x)) * f')) (at x within A)"
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3009
    using has_field_derivative_csqrt assms(1) by (rule DERIV_chain) fact
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3010
  thus ?thesis
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3011
    by (simp add: o_def field_simps)
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3012
qed
a1de627d417a More of Manuel's material
paulson <lp15@cam.ac.uk>
parents: 82338
diff changeset
  3013
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3014
lemma field_differentiable_at_csqrt:
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3015
    "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> csqrt field_differentiable at z"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3016
  using field_differentiable_def has_field_derivative_csqrt by blast
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3017
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3018
lemma field_differentiable_within_csqrt:
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3019
    "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> csqrt field_differentiable (at z within s)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3020
  using field_differentiable_at_csqrt field_differentiable_within_subset by blast
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3021
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3022
lemma continuous_at_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3023
    "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z) csqrt"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3024
  by (simp add: field_differentiable_within_csqrt field_differentiable_imp_continuous_at)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3025
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3026
corollary\<^marker>\<open>tag unimportant\<close> isCont_csqrt' [simp]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3027
   "\<lbrakk>isCont f z; f z \<notin> \<real>\<^sub>\<le>\<^sub>0\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. csqrt (f x)) z"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  3028
  by (blast intro: isCont_o2 [OF _ continuous_at_csqrt])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  3029
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3030
lemma continuous_within_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3031
    "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within s) csqrt"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3032
  by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_csqrt)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3033
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3034
lemma continuous_on_csqrt [continuous_intros]:
77221
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3035
    "continuous_on (-\<real>\<^sub>\<le>\<^sub>0) csqrt"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3036
  by (simp add: continuous_at_imp_continuous_on continuous_within_csqrt)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3037
77221
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3038
lemma holomorphic_on_csqrt [holomorphic_intros]: "csqrt holomorphic_on -\<real>\<^sub>\<le>\<^sub>0"
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3039
  by (simp add: field_differentiable_within_csqrt holomorphic_on_def)
77221
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3040
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3041
lemma holomorphic_on_csqrt' [holomorphic_intros]:
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3042
  "f holomorphic_on A \<Longrightarrow> (\<And>z. z \<in> A \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> (\<lambda>z. csqrt (f z)) holomorphic_on A"
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3043
  using holomorphic_on_compose_gen[OF _ holomorphic_on_csqrt, of f A] by (auto simp: o_def)
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3044
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3045
lemma analytic_on_csqrt [analytic_intros]: "csqrt analytic_on -\<real>\<^sub>\<le>\<^sub>0"
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3046
  using holomorphic_on_csqrt by (subst analytic_on_open) auto
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3047
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3048
lemma analytic_on_csqrt' [analytic_intros]:
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3049
  "f analytic_on A \<Longrightarrow> (\<And>z. z \<in> A \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> (\<lambda>z. csqrt (f z)) analytic_on A"
0cdb384bf56a More new theorems from the number theory development
paulson <lp15@cam.ac.uk>
parents: 77200
diff changeset
  3050
  using analytic_on_compose_gen[OF _ analytic_on_csqrt, of f A] by (auto simp: o_def)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3051
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3052
lemma continuous_within_closed_nontrivial:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3053
    "closed s \<Longrightarrow> a \<notin> s ==> continuous (at a within s) f"
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3054
  using Compl_iff continuous_within_topological open_Compl by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3055
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3056
lemma continuous_within_csqrt_posreal:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3057
    "continuous (at z within (\<real> \<inter> {w. 0 \<le> Re(w)})) csqrt"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3058
proof (cases "z \<in> \<real>\<^sub>\<le>\<^sub>0")
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3059
  case True
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3060
  then have [simp]: "Im z = 0" and 0: "Re z < 0 \<or> z = 0"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3061
    using complex_nonpos_Reals_iff complex_eq_iff by force+
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3062
  show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3063
    using 0
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3064
  proof
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3065
    assume "Re z < 0"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3066
    then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3067
      by (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge])
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3068
  next
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3069
    assume "z = 0"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3070
    moreover
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3071
    have "\<And>e. 0 < e
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3072
         \<Longrightarrow> \<forall>x'\<in>\<real> \<inter> {w. 0 \<le> Re w}. cmod x' < e^2 \<longrightarrow> cmod (csqrt x') < e"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3073
      by (auto simp: Reals_def real_less_lsqrt)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3074
    ultimately show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3075
      using zero_less_power by (fastforce simp: continuous_within_eps_delta)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3076
  qed
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3077
qed (blast intro: continuous_within_csqrt)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3078
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3079
subsection\<open>Complex arctangent\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3080
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3081
text\<open>The branch cut gives standard bounds in the real case.\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3082
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3083
definition\<^marker>\<open>tag important\<close> Arctan :: "complex \<Rightarrow> complex" where
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3084
    "Arctan \<equiv> \<lambda>z. (\<i>/2) * Ln((1 - \<i>*z) / (1 + \<i>*z))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3085
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3086
lemma Arctan_def_moebius: "Arctan z = \<i>/2 * Ln (moebius (-\<i>) 1 \<i> 1 z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3087
  by (simp add: Arctan_def moebius_def add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3088
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3089
lemma Ln_conv_Arctan:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3090
  assumes "z \<noteq> -1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3091
  shows   "Ln z = -2*\<i> * Arctan (moebius 1 (- 1) (- \<i>) (- \<i>) z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3092
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3093
  have "Arctan (moebius 1 (- 1) (- \<i>) (- \<i>) z) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3094
             \<i>/2 * Ln (moebius (- \<i>) 1 \<i> 1 (moebius 1 (- 1) (- \<i>) (- \<i>) z))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3095
    by (simp add: Arctan_def_moebius)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3096
  also from assms have "\<i> * z \<noteq> \<i> * (-1)" by (subst mult_left_cancel) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3097
  hence "\<i> * z - -\<i> \<noteq> 0" by (simp add: eq_neg_iff_add_eq_0)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3098
  from moebius_inverse'[OF _ this, of 1 1]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3099
    have "moebius (- \<i>) 1 \<i> 1 (moebius 1 (- 1) (- \<i>) (- \<i>) z) = z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3100
  finally show ?thesis by (simp add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3101
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3102
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3103
lemma Arctan_0 [simp]: "Arctan 0 = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3104
  by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3105
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3106
lemma Im_complex_div_lemma: "Im((1 - \<i>*z) / (1 + \<i>*z)) = 0 \<longleftrightarrow> Re z = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3107
  by (auto simp: Im_complex_div_eq_0 algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3108
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3109
lemma Re_complex_div_lemma: "0 < Re((1 - \<i>*z) / (1 + \<i>*z)) \<longleftrightarrow> norm z < 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3110
  by (simp add: Re_complex_div_gt_0 algebra_simps cmod_def power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3111
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3112
lemma tan_Arctan:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3113
  assumes "z\<^sup>2 \<noteq> -1"
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3114
  shows [simp]: "tan(Arctan z) = z"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3115
proof -
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3116
  obtain "1 + \<i>*z \<noteq> 0" "1 - \<i>*z \<noteq> 0"
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3117
    by (metis add_diff_cancel_left' assms diff_0 i_times_eq_iff mult_cancel_left2 power2_i power2_minus right_minus_eq)
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3118
  then show ?thesis
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3119
    by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps 
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3120
        flip: csqrt_exp_Ln power2_eq_square)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3121
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3122
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3123
lemma Arctan_tan [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3124
  assumes "\<bar>Re z\<bar> < pi/2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3125
    shows "Arctan(tan z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3126
proof -
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3127
  have "Ln ((1 - \<i> * tan z) / (1 + \<i> * tan z)) = 2 * z / \<i>"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3128
  proof (rule Ln_unique)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3129
    have ge_pi2: "\<And>n::int. \<bar>of_int (2*n + 1) * pi/2\<bar> \<ge> pi/2"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3130
      by (case_tac n rule: int_cases) (auto simp: abs_mult)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3131
    have "exp (\<i>*z)*exp (\<i>*z) = -1 \<longleftrightarrow> exp (2*\<i>*z) = -1"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3132
      by (metis distrib_right exp_add mult_2)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3133
    also have "\<dots> \<longleftrightarrow> exp (2*\<i>*z) = exp (\<i>*pi)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3134
      using cis_conv_exp cis_pi by auto
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3135
    also have "\<dots> \<longleftrightarrow> exp (2*\<i>*z - \<i>*pi) = 1"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3136
      by (metis (no_types) diff_add_cancel diff_minus_eq_add exp_add exp_minus_inverse mult.commute)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3137
    also have "\<dots> \<longleftrightarrow> Re(\<i>*2*z - \<i>*pi) = 0 \<and> (\<exists>n::int. Im(\<i>*2*z - \<i>*pi) = of_int (2 * n) * pi)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3138
      by (simp add: exp_eq_1)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3139
    also have "\<dots> \<longleftrightarrow> Im z = 0 \<and> (\<exists>n::int. 2 * Re z = of_int (2*n + 1) * pi)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3140
      by (simp add: algebra_simps)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3141
    also have "\<dots> \<longleftrightarrow> False"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3142
      using assms ge_pi2
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3143
      by (metis eq_divide_eq linorder_not_less mult.commute zero_neq_numeral)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3144
    finally have "exp (\<i>*z)*exp (\<i>*z) + 1 \<noteq> 0"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3145
      by (auto simp: add.commute minus_unique)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3146
    then show "exp (2 * z / \<i>) = (1 - \<i> * tan z) / (1 + \<i> * tan z)"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3147
      apply (simp add: tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3148
      by (simp add: algebra_simps flip: power2_eq_square exp_double)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3149
  qed (use assms in auto)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3150
  then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3151
    by (auto simp: Arctan_def)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3152
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3153
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3154
lemma
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3155
  assumes "Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3156
  shows Re_Arctan_bounds: "\<bar>Re(Arctan z)\<bar> < pi/2"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3157
    and has_field_derivative_Arctan: "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3158
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3159
  have nz0: "1 + \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3160
    using assms
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  3161
    by (metis abs_one add_diff_cancel_left' complex_i_mult_minus diff_0 i_squared imaginary_unit.simps
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3162
                less_asym neg_equal_iff_equal)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3163
  have "z \<noteq> -\<i>" using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3164
    by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3165
  then have zz: "1 + z * z \<noteq> 0"
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  3166
    by (metis abs_one assms i_squared imaginary_unit.simps less_irrefl minus_unique square_eq_iff)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3167
  have nz1: "1 - \<i>*z \<noteq> 0"
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  3168
    using assms by (force simp add: i_times_eq_iff)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3169
  have nz2: "inverse (1 + \<i>*z) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3170
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3171
    by (metis Im_complex_div_lemma Re_complex_div_lemma cmod_eq_Im divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3172
              less_irrefl mult_zero_right zero_complex.simps(1) zero_complex.simps(2))
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3173
  have nzi: "((1 - \<i>*z) * inverse (1 + \<i>*z)) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3174
    using nz1 nz2 by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3175
  have "Im ((1 - \<i>*z) / (1 + \<i>*z)) = 0 \<Longrightarrow> 0 < Re ((1 - \<i>*z) / (1 + \<i>*z))"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3176
    by (simp add: Im_complex_div_lemma Re_complex_div_lemma assms cmod_eq_Im)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3177
  then have *: "((1 - \<i>*z) / (1 + \<i>*z)) \<notin> \<real>\<^sub>\<le>\<^sub>0"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3178
    by (auto simp add: complex_nonpos_Reals_iff)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3179
  show "\<bar>Re(Arctan z)\<bar> < pi/2"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3180
    unfolding Arctan_def divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3181
    using mpi_less_Im_Ln [OF nzi]
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3182
    by (auto simp: abs_if intro!: Im_Ln_less_pi * [unfolded divide_complex_def])
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3183
  show "(Arctan has_field_derivative inverse(1 + z\<^sup>2)) (at z)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3184
    unfolding Arctan_def scaleR_conv_of_real
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3185
    apply (intro derivative_eq_intros | simp add: nz0 *)+
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3186
    using nz1 zz
71633
07bec530f02e cleaned proofs
nipkow
parents: 71184
diff changeset
  3187
    apply (simp add: field_split_simps power2_eq_square)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3188
    apply algebra
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3189
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3190
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3191
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3192
lemma field_differentiable_at_Arctan: "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan field_differentiable at z"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3193
  using has_field_derivative_Arctan
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3194
  by (auto simp: field_differentiable_def)
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3195
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3196
lemma field_differentiable_within_Arctan:
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3197
    "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan field_differentiable (at z within s)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3198
  using field_differentiable_at_Arctan field_differentiable_at_within by blast
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3199
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3200
declare has_field_derivative_Arctan [derivative_intros]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3201
declare has_field_derivative_Arctan [THEN DERIV_chain2, derivative_intros]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3202
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3203
lemma continuous_at_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3204
    "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> continuous (at z) Arctan"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3205
  by (simp add: field_differentiable_imp_continuous_at field_differentiable_within_Arctan)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3206
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3207
lemma continuous_within_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3208
    "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> continuous (at z within s) Arctan"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3209
  using continuous_at_Arctan continuous_at_imp_continuous_within by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3210
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3211
lemma continuous_on_Arctan [continuous_intros]:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3212
    "(\<And>z. z \<in> s \<Longrightarrow> Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> continuous_on s Arctan"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3213
  by (auto simp: continuous_at_imp_continuous_on continuous_within_Arctan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3214
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3215
lemma holomorphic_on_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3216
    "(\<And>z. z \<in> s \<Longrightarrow> Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3217
  by (simp add: field_differentiable_within_Arctan holomorphic_on_def)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3218
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3219
theorem Arctan_series:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3220
  assumes z: "norm (z :: complex) < 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3221
  defines "g \<equiv> \<lambda>n. if odd n then -\<i>*\<i>^n / n else 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3222
  defines "h \<equiv> \<lambda>z n. (-1)^n / of_nat (2*n+1) * (z::complex)^(2*n+1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3223
  shows   "(\<lambda>n. g n * z^n) sums Arctan z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3224
  and     "h z sums Arctan z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3225
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  3226
  define G where [abs_def]: "G z = (\<Sum>n. g n * z^n)" for z
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3227
  have summable: "summable (\<lambda>n. g n * u^n)" if "norm u < 1" for u
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3228
  proof (cases "u = 0")
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3229
    case False
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3230
    have "(\<lambda>n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\<lambda>n. ereal (inverse (norm u)^2) *
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3231
              ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3232
    proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3233
      fix n
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3234
      have "ereal (norm (h u n) / norm (h u (Suc n))) =
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3235
             ereal (inverse (norm u)^2) * ereal (((2*Suc n+1) / (Suc n)) /
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3236
                 ((2*Suc n-1) / (Suc n)))"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  3237
      by (simp add: h_def norm_mult norm_power norm_divide field_split_simps
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3238
                    power2_eq_square eval_nat_numeral del: of_nat_add of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3239
      also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  3240
        by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all?
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3241
      also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc n))"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  3242
        by (auto simp: field_split_simps simp del: of_nat_Suc) simp_all?
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3243
      finally show "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) *
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3244
              ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n))))" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3245
    qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3246
    also have "\<dots> \<longlonglongrightarrow> ereal (inverse (norm u)^2) * ereal ((2 + 0) / (2 - 0))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3247
      by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3248
    finally have "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) = inverse (norm u)^2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3249
      by (intro lim_imp_Liminf) simp_all
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3250
    moreover from power_strict_mono[OF that, of 2] False have "inverse (norm u)^2 > 1"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  3251
      by (simp add: field_split_simps)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3252
    ultimately have A: "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3253
    from False have "summable (h u)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3254
      by (intro summable_norm_cancel[OF ratio_test_convergence[OF _ A]])
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3255
         (auto simp: h_def norm_divide norm_mult norm_power simp del: of_nat_Suc
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3256
               intro!: mult_pos_pos divide_pos_pos always_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3257
    thus "summable (\<lambda>n. g n * u^n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3258
      by (subst summable_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66252
diff changeset
  3259
         (auto simp: power_mult strict_mono_def g_def h_def elim!: oddE)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3260
  qed (simp add: h_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3261
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3262
  have "\<exists>c. \<forall>u\<in>ball 0 1. Arctan u - G u = c"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3263
  proof (rule has_field_derivative_zero_constant)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3264
    fix u :: complex assume "u \<in> ball 0 1"
71633
07bec530f02e cleaned proofs
nipkow
parents: 71184
diff changeset
  3265
    hence u: "norm u < 1" by (simp)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  3266
    define K where "K = (norm u + 1) / 2"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3267
    from u and abs_Im_le_cmod[of u] have Im_u: "\<bar>Im u\<bar> < 1" by linarith
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3268
    from u have K: "0 \<le> K" "norm u < K" "K < 1" by (simp_all add: K_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3269
    hence "(G has_field_derivative (\<Sum>n. diffs g n * u ^ n)) (at u)" unfolding G_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3270
      by (intro termdiffs_strong[of _ "of_real K"] summable) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3271
    also have "(\<lambda>n. diffs g n * u^n) = (\<lambda>n. if even n then (\<i>*u)^n else 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3272
      by (intro ext) (simp_all del: of_nat_Suc add: g_def diffs_def power_mult_distrib)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3273
    also have "suminf \<dots> = (\<Sum>n. (-(u^2))^n)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3274
      by (subst suminf_mono_reindex[of "\<lambda>n. 2*n", symmetric])
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66252
diff changeset
  3275
         (auto elim!: evenE simp: strict_mono_def power_mult power_mult_distrib)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3276
    also from u have "norm u^2 < 1^2" by (intro power_strict_mono) simp_all
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3277
    hence "(\<Sum>n. (-(u^2))^n) = inverse (1 + u^2)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3278
      by (subst suminf_geometric) (simp_all add: norm_power inverse_eq_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3279
    finally have "(G has_field_derivative inverse (1 + u\<^sup>2)) (at u)" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3280
    from DERIV_diff[OF has_field_derivative_Arctan this] Im_u u
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3281
      show "((\<lambda>u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3282
      by (simp_all add: at_within_open[OF _ open_ball])
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3283
  qed simp_all
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3284
  then obtain c where c: "\<And>u. norm u < 1 \<Longrightarrow> Arctan u - G u = c" by auto
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3285
  from this[of 0] have "c = 0" by (simp add: G_def g_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3286
  with c z have "Arctan z = G z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3287
  with summable[OF z] show "(\<lambda>n. g n * z^n) sums Arctan z" unfolding G_def by (simp add: sums_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3288
  thus "h z sums Arctan z" by (subst (asm) sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66252
diff changeset
  3289
                              (auto elim!: oddE simp: strict_mono_def power_mult g_def h_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3290
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3291
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3292
text \<open>A quickly-converging series for the logarithm, based on the arctangent.\<close>
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3293
theorem ln_series_quadratic:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3294
  assumes x: "x > (0::real)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3295
  shows "(\<lambda>n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3296
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  3297
  define y :: complex where "y = of_real ((x-1)/(x+1))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3298
  from x have x': "complex_of_real x \<noteq> of_real (-1)"  by (subst of_real_eq_iff) auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3299
  from x have "\<bar>x - 1\<bar> < \<bar>x + 1\<bar>" by linarith
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3300
  hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3301
    by (simp add: norm_divide del: of_real_add of_real_diff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3302
  hence "norm (\<i> * y) < 1" unfolding y_def by (subst norm_mult) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3303
  hence "(\<lambda>n. (-2*\<i>) * ((-1)^n / of_nat (2*n+1) * (\<i>*y)^(2*n+1))) sums ((-2*\<i>) * Arctan (\<i>*y))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3304
    by (intro Arctan_series sums_mult) simp_all
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3305
  also have "(\<lambda>n. (-2*\<i>) * ((-1)^n / of_nat (2*n+1) * (\<i>*y)^(2*n+1))) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3306
                 (\<lambda>n. (-2*\<i>) * ((-1)^n * (\<i>*y*(-y\<^sup>2)^n)/of_nat (2*n+1)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3307
    by (intro ext) (simp_all add: power_mult power_mult_distrib)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3308
  also have "\<dots> = (\<lambda>n. 2*y* ((-1) * (-y\<^sup>2))^n/of_nat (2*n+1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3309
    by (intro ext, subst power_mult_distrib) (simp add: algebra_simps power_mult)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3310
  also have "\<dots> = (\<lambda>n. 2*y^(2*n+1) / of_nat (2*n+1))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3311
    by (subst power_add, subst power_mult) (simp add: mult_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3312
  also have "\<dots> = (\<lambda>n. of_real (2*((x-1)/(x+1))^(2*n+1) / of_nat (2*n+1)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3313
    by (intro ext) (simp add: y_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3314
  also have "\<i> * y = (of_real x - 1) / (-\<i> * (of_real x + 1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3315
    by (subst divide_divide_eq_left [symmetric]) (simp add: y_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3316
  also have "\<dots> = moebius 1 (-1) (-\<i>) (-\<i>) (of_real x)" by (simp add: moebius_def algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3317
  also from x' have "-2*\<i>*Arctan \<dots> = Ln (of_real x)" by (intro Ln_conv_Arctan [symmetric]) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3318
  also from x have "\<dots> = ln x" by (rule Ln_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3319
  finally show ?thesis by (subst (asm) sums_of_real_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3320
qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3321
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3322
subsection\<^marker>\<open>tag unimportant\<close> \<open>Real arctangent\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3323
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3324
lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3325
proof -
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3326
  have ne: "1 + x\<^sup>2 \<noteq> 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3327
    by (metis power_one sum_power2_eq_zero_iff zero_neq_one)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3328
  have ne1: "1 + \<i> * complex_of_real x \<noteq> 0"
75494
eded3fe9e600 Five slightly useful lemmas
paulson <lp15@cam.ac.uk>
parents: 74513
diff changeset
  3329
    using Complex_eq complex_eq_cancel_iff2 by fastforce
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3330
  have "Re (Ln ((1 - \<i> * x) * inverse (1 + \<i> * x))) = 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3331
    apply (rule norm_exp_imaginary)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3332
    using ne
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3333
    apply (simp add: ne1 cmod_def)
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  3334
    apply (auto simp: field_split_simps)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3335
    apply algebra
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3336
    done
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3337
  then show ?thesis
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3338
    unfolding Arctan_def divide_complex_def by (simp add: complex_eq_iff)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3339
qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3340
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3341
lemma arctan_eq_Re_Arctan: "arctan x = Re (Arctan (of_real x))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3342
proof (rule arctan_unique)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3343
  have "(1 - \<i> * x) / (1 + \<i> * x) \<notin> \<real>\<^sub>\<le>\<^sub>0"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3344
    by (auto simp: Im_complex_div_lemma complex_nonpos_Reals_iff)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3345
  then show "- (pi / 2) < Re (Arctan (complex_of_real x))"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3346
    by (simp add: Arctan_def Im_Ln_less_pi)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3347
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3348
  have *: " (1 - \<i>*x) / (1 + \<i>*x) \<noteq> 0"
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  3349
    by (simp add: field_split_simps) ( simp add: complex_eq_iff)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3350
  show "Re (Arctan (complex_of_real x)) < pi / 2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3351
    using mpi_less_Im_Ln [OF *]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3352
    by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3353
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3354
  have "tan (Re (Arctan (of_real x))) = Re (tan (Arctan (of_real x)))"
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3355
    by (metis Im_Arctan_of_real Re_complex_of_real complex_is_Real_iff of_real_Re tan_of_real)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3356
  also have "\<dots> = x"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3357
  proof -
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3358
    have "(complex_of_real x)\<^sup>2 \<noteq> - 1"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3359
      by (smt (verit, best) Im_complex_of_real imaginary_unit.sel(2) of_real_minus power2_eq_iff power2_i)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3360
    then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3361
      by simp
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3362
  qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3363
  finally show "tan (Re (Arctan (complex_of_real x))) = x" .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3364
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3365
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3366
lemma Arctan_of_real: "Arctan (of_real x) = of_real (arctan x)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3367
  unfolding arctan_eq_Re_Arctan divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3368
  by (simp add: complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3369
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3370
lemma Arctan_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Arctan z \<in> \<real>"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3371
  by (metis Reals_cases Reals_of_real Arctan_of_real)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3372
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3373
declare arctan_one [simp]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3374
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3375
lemma arctan_less_pi4_pos: "x < 1 \<Longrightarrow> arctan x < pi/4"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3376
  by (metis arctan_less_iff arctan_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3377
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3378
lemma arctan_less_pi4_neg: "-1 < x \<Longrightarrow> -(pi/4) < arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3379
  by (metis arctan_less_iff arctan_minus arctan_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3380
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3381
lemma arctan_less_pi4: "\<bar>x\<bar> < 1 \<Longrightarrow> \<bar>arctan x\<bar> < pi/4"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3382
  by (metis abs_less_iff arctan_less_pi4_pos arctan_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3383
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3384
lemma arctan_le_pi4: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> \<bar>arctan x\<bar> \<le> pi/4"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3385
  by (metis abs_le_iff arctan_le_iff arctan_minus arctan_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3386
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3387
lemma abs_arctan: "\<bar>arctan x\<bar> = arctan \<bar>x\<bar>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3388
  by (simp add: abs_if arctan_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3389
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3390
lemma arctan_add_raw:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3391
  assumes "\<bar>arctan x + arctan y\<bar> < pi/2"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3392
    shows "arctan x + arctan y = arctan((x + y) / (1 - x * y))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3393
proof (rule arctan_unique [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3394
  show 12: "- (pi / 2) < arctan x + arctan y" "arctan x + arctan y < pi / 2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3395
    using assms by linarith+
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3396
  show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)"
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3397
    using cos_gt_zero_pi [OF 12] by (simp add: arctan tan_add)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3398
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3399
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3400
lemma arctan_inverse:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3401
  "0 < x \<Longrightarrow>arctan(inverse x) = pi/2 - arctan x"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3402
  by (smt (verit, del_insts) arctan arctan_unique tan_cot zero_less_arctan_iff)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3403
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3404
lemma arctan_add_small:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3405
  assumes "\<bar>x * y\<bar> < 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3406
    shows "(arctan x + arctan y = arctan((x + y) / (1 - x * y)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3407
proof (cases "x = 0 \<or> y = 0")
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3408
  case False
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3409
  with assms have "\<bar>x\<bar> < inverse \<bar>y\<bar>"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3410
    by (simp add: field_split_simps abs_mult)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3411
  with False have "\<bar>arctan x\<bar> < pi / 2 - \<bar>arctan y\<bar>" using assms
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3412
    by (auto simp add: abs_arctan arctan_inverse [symmetric] arctan_less_iff)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3413
  then show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3414
    by (intro arctan_add_raw) linarith
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3415
qed auto
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3416
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3417
lemma abs_arctan_le:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3418
  fixes x::real shows "\<bar>arctan x\<bar> \<le> \<bar>x\<bar>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3419
proof -
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3420
  have 1: "\<And>x. x \<in> \<real> \<Longrightarrow> cmod (inverse (1 + x\<^sup>2)) \<le> 1"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3421
    by (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3422
  have "cmod (Arctan w - Arctan z) \<le> 1 * cmod (w-z)" if "w \<in> \<real>" "z \<in> \<real>" for w z
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3423
    apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1])
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3424
       apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan])
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3425
    using 1 that by (auto simp: Reals_def)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3426
  then have "cmod (Arctan (of_real x) - Arctan 0) \<le> 1 * cmod (of_real x - 0)"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3427
    using Reals_0 Reals_of_real by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3428
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3429
    by (simp add: Arctan_of_real)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3430
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3431
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3432
lemma arctan_le_self: "0 \<le> x \<Longrightarrow> arctan x \<le> x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3433
  by (metis abs_arctan_le abs_of_nonneg zero_le_arctan_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3434
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3435
lemma abs_tan_ge: "\<bar>x\<bar> < pi/2 \<Longrightarrow> \<bar>x\<bar> \<le> \<bar>tan x\<bar>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3436
  by (metis abs_arctan_le abs_less_iff arctan_tan minus_less_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3437
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3438
lemma arctan_bounds:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3439
  assumes "0 \<le> x" "x < 1"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3440
  shows arctan_lower_bound:
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3441
    "(\<Sum>k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \<le> arctan x" (is "(\<Sum>k<_. _ * ?a k) \<le> _")
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3442
    and arctan_upper_bound:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3443
    "arctan x \<le> (\<Sum>k<2 * n + 1. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1)))"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3444
proof -
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3445
  have tendsto_zero: "?a \<longlonglongrightarrow> 0"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3446
  proof (rule tendsto_eq_rhs)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3447
    show "(\<lambda>k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \<longlonglongrightarrow> 0 * 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3448
      using assms
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3449
      by (intro tendsto_mult real_tendsto_divide_at_top)
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3450
        (auto simp: filterlim_sequentially_iff_filterlim_real
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3451
          intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3452
          tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3453
  qed simp
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3454
  have nonneg: "0 \<le> ?a n" for n
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3455
    by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3456
  have le: "?a (Suc n) \<le> ?a n" for n
70817
dd675800469d dedicated fact collections for algebraic simplification rules potentially splitting goals
haftmann
parents: 70724
diff changeset
  3457
    by (rule mult_mono[OF _ power_decreasing]) (auto simp: field_split_simps assms less_imp_le)
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3458
  from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3459
    summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3460
    assms
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3461
  show "(\<Sum>k<2*n. (- 1)^ k * ?a k) \<le> arctan x" "arctan x \<le> (\<Sum>k<2 * n + 1. (- 1)^ k * ?a k)"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3462
    by (auto simp: arctan_series)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3463
qed
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3464
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3465
subsection\<^marker>\<open>tag unimportant\<close> \<open>Bounds on pi using real arctangent\<close>
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3466
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3467
lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3468
  using machin by simp
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3469
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3470
lemma pi_approx: "3.141592653588 \<le> pi" "pi \<le> 3.1415926535899"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3471
  unfolding pi_machin
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3472
  using arctan_bounds[of "1/5"   4]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3473
        arctan_bounds[of "1/239" 4]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3474
  by (simp_all add: eval_nat_numeral)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  3475
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3476
lemma pi_gt3: "pi > 3"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  3477
  using pi_approx by simp
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3478
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3479
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3480
subsection\<open>Inverse Sine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3481
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3482
definition\<^marker>\<open>tag important\<close> Arcsin :: "complex \<Rightarrow> complex" where
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3483
   "Arcsin \<equiv> \<lambda>z. -\<i> * Ln(\<i> * z + csqrt(1 - z\<^sup>2))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3484
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3485
lemma Arcsin_body_lemma: "\<i> * z + csqrt(1 - z\<^sup>2) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3486
  using power2_csqrt [of "1 - z\<^sup>2"]
77275
386b1b33785c New material due to Eberl on Formal Laurent Series
paulson <lp15@cam.ac.uk>
parents: 77273
diff changeset
  3487
  by (metis add.inverse_unique diff_0 diff_add_cancel mult.left_commute mult_minus1_right power2_i power2_minus power_mult_distrib zero_neq_one)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3488
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3489
lemma Arcsin_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Re(\<i> * z + csqrt(1 - z\<^sup>2))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3490
  using Complex.cmod_power2 [of z, symmetric]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3491
  by (simp add: real_less_rsqrt algebra_simps Re_power2 cmod_square_less_1_plus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3492
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3493
lemma Re_Arcsin: "Re(Arcsin z) = Im (Ln (\<i> * z + csqrt(1 - z\<^sup>2)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3494
  by (simp add: Arcsin_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3495
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3496
lemma Im_Arcsin: "Im(Arcsin z) = - ln (cmod (\<i> * z + csqrt (1 - z\<^sup>2)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3497
  by (simp add: Arcsin_def Arcsin_body_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3498
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3499
lemma one_minus_z2_notin_nonpos_Reals:
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3500
  assumes "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3501
  shows "1 - z\<^sup>2 \<notin> \<real>\<^sub>\<le>\<^sub>0"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3502
proof (cases "Im z = 0")
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3503
  case True
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3504
  with assms show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3505
    by (simp add: complex_nonpos_Reals_iff flip: abs_square_less_1)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3506
next
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3507
  case False
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3508
  have "\<not> (Im z)\<^sup>2 \<le> - 1"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3509
    using False power2_less_eq_zero_iff by fastforce
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3510
  with False show ?thesis
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3511
    by (auto simp add: complex_nonpos_Reals_iff Re_power2 Im_power2)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3512
qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3513
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3514
lemma isCont_Arcsin_lemma:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3515
  assumes le0: "Re (\<i> * z + csqrt (1 - z\<^sup>2)) \<le> 0" and "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3516
    shows False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3517
proof (cases "Im z = 0")
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3518
  case True
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3519
  then show ?thesis
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3520
    using assms by (fastforce simp: cmod_def abs_square_less_1 [symmetric])
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3521
next
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3522
  case False
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3523
  have leim: "(cmod (1 - z\<^sup>2) + (1 - Re (z\<^sup>2))) / 2 \<le> (Im z)\<^sup>2"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3524
    using le0 sqrt_le_D by fastforce
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3525
  have neq: "(cmod z)\<^sup>2 \<noteq> 1 + cmod (1 - z\<^sup>2)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3526
  proof (clarsimp simp add: cmod_def)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3527
    assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 = 1 + sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3528
    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)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3529
      by simp
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3530
    then show False using False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3531
      by (simp add: power2_eq_square algebra_simps)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3532
  qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3533
  moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3534
    using leim cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1]
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3535
    by (simp add: norm_power Re_power2 norm_minus_commute [of 1])
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3536
  ultimately show False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3537
    by (simp add: Re_power2 Im_power2 cmod_power2)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3538
qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3539
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3540
lemma isCont_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3541
  assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3542
    shows "isCont Arcsin z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3543
proof -
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3544
  have 1: "\<i> * z + csqrt (1 - z\<^sup>2) \<notin> \<real>\<^sub>\<le>\<^sub>0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3545
    by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3546
  have 2: "1 - z\<^sup>2 \<notin> \<real>\<^sub>\<le>\<^sub>0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3547
    by (simp add: one_minus_z2_notin_nonpos_Reals assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3548
  show ?thesis
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3549
    using assms unfolding Arcsin_def by (intro isCont_Ln' isCont_csqrt' continuous_intros 1 2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3550
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3551
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3552
lemma isCont_Arcsin' [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3553
  shows "isCont f z \<Longrightarrow> (Im (f z) = 0 \<Longrightarrow> \<bar>Re (f z)\<bar> < 1) \<Longrightarrow> isCont (\<lambda>x. Arcsin (f x)) z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3554
  by (blast intro: isCont_o2 [OF _ isCont_Arcsin])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3555
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3556
lemma sin_Arcsin [simp]: "sin(Arcsin z) = z"
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60020
diff changeset
  3557
proof -
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3558
  have "\<i>*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \<longleftrightarrow> (\<i>*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67371
diff changeset
  3559
    by (simp add: algebra_simps)  \<comment> \<open>Cancelling a factor of 2\<close>
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3560
  moreover have "\<dots> \<longleftrightarrow> (\<i>*z) + csqrt (1 - z\<^sup>2) = 0"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3561
    by (metis Arcsin_body_lemma distrib_right no_zero_divisors zero_neq_numeral)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3562
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3563
    apply (simp add: sin_exp_eq Arcsin_def Arcsin_body_lemma exp_minus divide_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3564
    apply (simp add: algebra_simps)
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3565
    apply (simp add: right_diff_distrib flip: power2_eq_square)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3566
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3567
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3568
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3569
lemma Re_eq_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3570
    "\<bar>Re z\<bar> = pi/2 \<Longrightarrow> Im z = 0 \<Longrightarrow>
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3571
      Re ((exp (\<i>*z) + inverse (exp (\<i>*z))) / 2) = 0 \<and> 0 \<le> Im ((exp (\<i>*z) + inverse (exp (\<i>*z))) / 2)"
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  3572
  apply (simp add: cos_i_times [symmetric] Re_cos Im_cos abs_if del: eq_divide_eq_numeral1)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3573
  by (metis cos_minus cos_pi_half)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3574
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3575
lemma Re_less_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3576
  assumes "\<bar>Re z\<bar> < pi / 2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3577
    shows "0 < Re ((exp (\<i>*z) + inverse (exp (\<i>*z))) / 2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3578
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3579
  have "0 < cos (Re z)" using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3580
    using cos_gt_zero_pi by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3581
  then show ?thesis
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  3582
    by (simp add: cos_i_times [symmetric] Re_cos Im_cos add_pos_pos)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3583
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3584
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3585
lemma Arcsin_sin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3586
    assumes "\<bar>Re z\<bar> < pi/2 \<or> (\<bar>Re z\<bar> = pi/2 \<and> Im z = 0)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3587
      shows "Arcsin(sin z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3588
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3589
  have "Arcsin(sin z) = - (\<i> * Ln (csqrt (1 - (\<i> * (exp (\<i>*z) - inverse (exp (\<i>*z))))\<^sup>2 / 4) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  3590
    by (simp add: sin_exp_eq Arcsin_def exp_minus power_divide)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3591
  also have "\<dots> = - (\<i> * Ln (csqrt (((exp (\<i>*z) + inverse (exp (\<i>*z)))/2)\<^sup>2) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3592
    by (simp add: field_simps power2_eq_square)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3593
  also have "\<dots> = - (\<i> * Ln (((exp (\<i>*z) + inverse (exp (\<i>*z)))/2) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3594
    apply (subst csqrt_square)
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3595
    using assms Re_eq_pihalf_lemma Re_less_pihalf_lemma by auto
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3596
  also have "\<dots> =  - (\<i> * Ln (exp (\<i>*z)))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3597
    by (simp add: field_simps power2_eq_square)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3598
  also have "\<dots> = z"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3599
    using assms by (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3600
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3601
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3602
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3603
lemma Arcsin_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3604
    "\<lbrakk>sin z = w; \<bar>Re z\<bar> < pi/2 \<or> (\<bar>Re z\<bar> = pi/2 \<and> Im z = 0)\<rbrakk> \<Longrightarrow> Arcsin w = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3605
  by (metis Arcsin_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3606
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3607
lemma Arcsin_0 [simp]: "Arcsin 0 = 0"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3608
  by (simp add: Arcsin_unique)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3609
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3610
lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3611
  using Arcsin_unique sin_of_real_pi_half by fastforce
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3612
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3613
lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3614
  by (simp add: Arcsin_unique)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3615
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3616
lemma has_field_derivative_Arcsin:
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3617
  assumes "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3618
    shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  3619
proof -
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3620
  have "(sin (Arcsin z))\<^sup>2 \<noteq> 1"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3621
    using assms one_minus_z2_notin_nonpos_Reals by force
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3622
  then have "cos (Arcsin z) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3623
    by (metis diff_0_right power_zero_numeral sin_squared_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3624
  then show ?thesis
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3625
    by (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) (auto intro: isCont_Arcsin assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3626
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3627
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3628
declare has_field_derivative_Arcsin [derivative_intros]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3629
declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3630
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3631
lemma field_differentiable_at_Arcsin:
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3632
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin field_differentiable at z"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3633
  using field_differentiable_def has_field_derivative_Arcsin by blast
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3634
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3635
lemma field_differentiable_within_Arcsin:
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3636
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin field_differentiable (at z within s)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3637
  using field_differentiable_at_Arcsin field_differentiable_within_subset by blast
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3638
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3639
lemma continuous_within_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3640
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous (at z within s) Arcsin"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3641
  using continuous_at_imp_continuous_within isCont_Arcsin by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3642
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3643
lemma continuous_on_Arcsin [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3644
    "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous_on s Arcsin"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3645
  by (simp add: continuous_at_imp_continuous_on)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3646
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3647
lemma holomorphic_on_Arcsin: "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3648
  by (simp add: field_differentiable_within_Arcsin holomorphic_on_def)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3649
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3650
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3651
subsection\<open>Inverse Cosine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3652
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3653
definition\<^marker>\<open>tag important\<close> Arccos :: "complex \<Rightarrow> complex" where
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3654
   "Arccos \<equiv> \<lambda>z. -\<i> * Ln(z + \<i> * csqrt(1 - z\<^sup>2))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3655
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3656
lemma Arccos_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Im(z + \<i> * csqrt(1 - z\<^sup>2))"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3657
  using Arcsin_range_lemma [of "-z"]  by simp
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3658
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3659
lemma Arccos_body_lemma: "z + \<i> * csqrt(1 - z\<^sup>2) \<noteq> 0"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3660
  by (metis Arcsin_body_lemma complex_i_mult_minus diff_0 diff_eq_eq power2_minus)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3661
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3662
lemma Re_Arccos: "Re(Arccos z) = Im (Ln (z + \<i> * csqrt(1 - z\<^sup>2)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3663
  by (simp add: Arccos_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3664
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3665
lemma Im_Arccos: "Im(Arccos z) = - ln (cmod (z + \<i> * csqrt (1 - z\<^sup>2)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3666
  by (simp add: Arccos_def Arccos_body_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3667
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3668
text\<open>A very tricky argument to find!\<close>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3669
lemma isCont_Arccos_lemma:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3670
  assumes eq0: "Im (z + \<i> * csqrt (1 - z\<^sup>2)) = 0" and "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3671
    shows False
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3672
proof (cases "Im z = 0")
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3673
  case True
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3674
  then show ?thesis
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3675
    using assms by (fastforce simp add: cmod_def abs_square_less_1 [symmetric])
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3676
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3677
  case False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3678
  have Imz: "Im z = - sqrt ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3679
    using eq0 abs_Re_le_cmod [of "1-z\<^sup>2"]
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3680
    by (simp add: Re_power2 algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3681
  have "(cmod z)\<^sup>2 - 1 \<noteq> cmod (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3682
  proof (clarsimp simp add: cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3683
    assume "(Re z)\<^sup>2 + (Im z)\<^sup>2 - 1 = sqrt ((1 - Re (z\<^sup>2))\<^sup>2 + (Im (z\<^sup>2))\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3684
    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)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3685
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3686
    then show False using False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3687
      by (simp add: power2_eq_square algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3688
  qed
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3689
  moreover have "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3690
    using abs_Re_le_cmod [of "1-z\<^sup>2"] by (subst Imz) (simp add: Re_power2)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3691
  ultimately show False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3692
    by (simp add: cmod_power2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3693
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3694
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3695
lemma isCont_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3696
  assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3697
    shows "isCont Arccos z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3698
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3699
  have "z + \<i> * csqrt (1 - z\<^sup>2) \<notin> \<real>\<^sub>\<le>\<^sub>0"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3700
    by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3701
  with assms show ?thesis
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3702
    unfolding Arccos_def
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3703
    by (simp_all add: one_minus_z2_notin_nonpos_Reals assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3704
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3705
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3706
lemma isCont_Arccos' [simp]:
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3707
  "isCont f z \<Longrightarrow> (Im (f z) = 0 \<Longrightarrow> \<bar>Re (f z)\<bar> < 1) \<Longrightarrow> isCont (\<lambda>x. Arccos (f x)) z"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3708
  by (blast intro: isCont_o2 [OF _ isCont_Arccos])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3709
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3710
lemma cos_Arccos [simp]: "cos(Arccos z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3711
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3712
  have "z*2 + \<i> * (2 * csqrt (1 - z\<^sup>2)) = 0 \<longleftrightarrow> z*2 + \<i> * csqrt (1 - z\<^sup>2)*2 = 0"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67371
diff changeset
  3713
    by (simp add: algebra_simps)  \<comment> \<open>Cancelling a factor of 2\<close>
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3714
  moreover have "\<dots> \<longleftrightarrow> z + \<i> * csqrt (1 - z\<^sup>2) = 0"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3715
    by (metis distrib_right mult_eq_0_iff zero_neq_numeral)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3716
  ultimately show ?thesis
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3717
    by (simp add: cos_exp_eq Arccos_def Arccos_body_lemma exp_minus field_simps flip: power2_eq_square)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3718
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3719
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3720
lemma Arccos_cos:
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3721
    assumes "0 < Re z \<and> Re z < pi \<or>
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3722
             Re z = 0 \<and> 0 \<le> Im z \<or>
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3723
             Re z = pi \<and> Im z \<le> 0"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3724
      shows "Arccos(cos z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3725
proof -
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3726
  have *: "((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z))) = sin z"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3727
    by (simp add: sin_exp_eq exp_minus field_simps power2_eq_square)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3728
  have "1 - (exp (\<i> * z) + inverse (exp (\<i> * z)))\<^sup>2 / 4 = ((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z)))\<^sup>2"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3729
    by (simp add: field_simps power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3730
  then have "Arccos(cos z) = - (\<i> * Ln ((exp (\<i> * z) + inverse (exp (\<i> * z))) / 2 +
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3731
                           \<i> * csqrt (((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z)))\<^sup>2)))"
61694
6571c78c9667 Removed some legacy theorems; minor adjustments to simplification rules; new material on homotopic paths
paulson <lp15@cam.ac.uk>
parents: 61649
diff changeset
  3732
    by (simp add: cos_exp_eq Arccos_def exp_minus power_divide)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3733
  also have "\<dots> = - (\<i> * Ln ((exp (\<i> * z) + inverse (exp (\<i> * z))) / 2 +
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3734
                              \<i> * ((\<i> - (exp (\<i> * z))\<^sup>2 * \<i>) / (2 * exp (\<i> * z)))))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3735
    apply (subst csqrt_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3736
    using assms Re_sin_pos [of z] Im_sin_nonneg [of z] Im_sin_nonneg2 [of z]
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3737
    by (auto simp: * Re_sin Im_sin)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3738
  also have "\<dots> =  - (\<i> * Ln (exp (\<i>*z)))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3739
    by (simp add: field_simps power2_eq_square)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3740
  also have "\<dots> = z"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3741
    using assms
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3742
    by (subst Complex_Transcendental.Ln_exp, auto)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3743
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3744
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3745
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3746
lemma Arccos_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3747
    "\<lbrakk>cos z = w;
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3748
      0 < Re z \<and> Re z < pi \<or>
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3749
      Re z = 0 \<and> 0 \<le> Im z \<or>
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3750
      Re z = pi \<and> Im z \<le> 0\<rbrakk> \<Longrightarrow> Arccos w = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3751
  using Arccos_cos by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3752
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3753
lemma Arccos_0 [simp]: "Arccos 0 = pi/2"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3754
  by (rule Arccos_unique) auto
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3755
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3756
lemma Arccos_1 [simp]: "Arccos 1 = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3757
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3758
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3759
lemma Arccos_minus1: "Arccos(-1) = pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3760
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3761
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3762
lemma has_field_derivative_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3763
  assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3764
    shows "(Arccos has_field_derivative - inverse(sin(Arccos z))) (at z)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3765
proof -
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3766
  have "x\<^sup>2 \<noteq> -1" for x::real
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3767
    by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x__]^2)))))")
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3768
  with assms have "(cos (Arccos z))\<^sup>2 \<noteq> 1"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3769
    by (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3770
  then have "- sin (Arccos z) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3771
    by (metis cos_squared_eq diff_0_right mult_zero_left neg_0_equal_iff_equal power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3772
  then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3773
    by (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]])
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3774
       (auto intro: isCont_Arccos assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3775
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3776
    by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3777
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3778
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3779
declare has_field_derivative_Arcsin [derivative_intros]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3780
declare has_field_derivative_Arcsin [THEN DERIV_chain2, derivative_intros]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3781
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3782
lemma field_differentiable_at_Arccos:
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3783
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos field_differentiable at z"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3784
  using field_differentiable_def has_field_derivative_Arccos by blast
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3785
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3786
lemma field_differentiable_within_Arccos:
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3787
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos field_differentiable (at z within s)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3788
  using field_differentiable_at_Arccos field_differentiable_within_subset by blast
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3789
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3790
lemma continuous_within_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3791
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous (at z within s) Arccos"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3792
  using continuous_at_imp_continuous_within isCont_Arccos by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3793
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3794
lemma continuous_on_Arccos [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3795
    "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> continuous_on s Arccos"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3796
  by (simp add: continuous_at_imp_continuous_on)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3797
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3798
lemma holomorphic_on_Arccos: "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3799
  by (simp add: field_differentiable_within_Arccos holomorphic_on_def)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3800
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3801
subsection\<^marker>\<open>tag unimportant\<close>\<open>Upper and Lower Bounds for Inverse Sine and Cosine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3802
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3803
lemma Arcsin_bounds: "\<bar>Re z\<bar> < 1 \<Longrightarrow> \<bar>Re(Arcsin z)\<bar> < pi/2"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3804
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3805
  by (blast intro: Re_Ln_pos_lt_imp Arcsin_range_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3806
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3807
lemma Arccos_bounds: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Re(Arccos z) \<and> Re(Arccos z) < pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3808
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3809
  by (blast intro!: Im_Ln_pos_lt_imp Arccos_range_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3810
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3811
lemma Re_Arccos_bounds: "-pi < Re(Arccos z) \<and> Re(Arccos z) \<le> pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3812
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3813
  by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arccos_body_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3814
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3815
lemma Re_Arccos_bound: "\<bar>Re(Arccos z)\<bar> \<le> pi"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
  3816
  by (meson Re_Arccos_bounds abs_le_iff less_eq_real_def minus_less_iff)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3817
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3818
lemma Im_Arccos_bound: "\<bar>Im (Arccos w)\<bar> \<le> cmod w"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3819
proof -
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3820
  have "(Im (Arccos w))\<^sup>2 \<le> (cmod (cos (Arccos w)))\<^sup>2 - (cos (Re (Arccos w)))\<^sup>2"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3821
    using norm_cos_squared [of "Arccos w"] real_le_abs_sinh [of "Im (Arccos w)"]
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3822
    by (simp only: abs_le_square_iff) (simp add: field_split_simps)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3823
  also have "\<dots> \<le> (cmod w)\<^sup>2"
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3824
    by (auto simp: cmod_power2)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3825
  finally show ?thesis
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  3826
    using abs_le_square_iff by force
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3827
qed
65064
a4abec71279a Renamed ii to imaginary_unit in order to free up ii as a variable name. Also replaced some legacy def commands
paulson <lp15@cam.ac.uk>
parents: 65036
diff changeset
  3828
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3829
lemma Re_Arcsin_bounds: "-pi < Re(Arcsin z) & Re(Arcsin z) \<le> pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3830
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3831
  by (blast intro!: mpi_less_Im_Ln Im_Ln_le_pi Arcsin_body_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3832
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3833
lemma Re_Arcsin_bound: "\<bar>Re(Arcsin z)\<bar> \<le> pi"
61649
268d88ec9087 Tweaks for "real": Removal of [iff] status for some lemmas, adding [simp] for others. Plus fixes.
paulson <lp15@cam.ac.uk>
parents: 61610
diff changeset
  3834
  by (meson Re_Arcsin_bounds abs_le_iff less_eq_real_def minus_less_iff)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  3835
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3836
lemma norm_Arccos_bounded:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3837
  fixes w :: complex
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3838
  shows "norm (Arccos w) \<le> pi + norm w"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3839
proof -
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3840
  have Re: "(Re (Arccos w))\<^sup>2 \<le> pi\<^sup>2" "(Im (Arccos w))\<^sup>2 \<le> (cmod w)\<^sup>2"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3841
    using Re_Arccos_bound [of w] Im_Arccos_bound [of w] abs_le_square_iff by force+
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3842
  have "Arccos w \<bullet> Arccos w \<le> pi\<^sup>2 + (cmod w)\<^sup>2"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3843
    using Re by (simp add: dot_square_norm cmod_power2 [of "Arccos w"])
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3844
  then have "cmod (Arccos w) \<le> pi + cmod (cos (Arccos w))"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3845
    by (smt (verit) Im_Arccos_bound Re_Arccos_bound cmod_le cos_Arccos)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3846
  then show "cmod (Arccos w) \<le> pi + cmod w"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3847
    by auto
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3848
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3849
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3850
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3851
subsection\<^marker>\<open>tag unimportant\<close>\<open>Interrelations between Arcsin and Arccos\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3852
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3853
lemma cos_Arcsin_nonzero: "z\<^sup>2 \<noteq> 1 \<Longrightarrow>cos(Arcsin z) \<noteq> 0"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3854
  by (metis diff_0_right power_zero_numeral sin_Arcsin sin_squared_eq)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3855
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3856
lemma sin_Arccos_nonzero: "z\<^sup>2 \<noteq> 1 \<Longrightarrow>sin(Arccos z) \<noteq> 0"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3857
  by (metis add.right_neutral cos_Arccos power2_eq_square power_zero_numeral sin_cos_squared_add3)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3858
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3859
lemma cos_sin_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3860
  assumes "0 < cos(Re z)  \<or>  cos(Re z) = 0 \<and> Im z * sin(Re z) \<le> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3861
    shows "cos z = csqrt(1 - (sin z)\<^sup>2)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3862
proof (rule csqrt_unique [THEN sym])
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3863
  show "(cos z)\<^sup>2 = 1 - (sin z)\<^sup>2"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3864
    by (simp add: cos_squared_eq)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3865
qed (use assms in \<open>auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff\<close>)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3866
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3867
lemma sin_cos_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3868
  assumes "0 < sin(Re z)  \<or>  sin(Re z) = 0 \<and> 0 \<le> Im z * cos(Re z)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3869
    shows "sin z = csqrt(1 - (cos z)\<^sup>2)"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3870
proof (rule csqrt_unique [THEN sym])
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3871
  show "(sin z)\<^sup>2 = 1 - (cos z)\<^sup>2"
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3872
    by (simp add: sin_squared_eq)
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  3873
qed (use assms in \<open>auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff\<close>)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3874
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3875
lemma Arcsin_Arccos_csqrt_pos:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3876
    "(0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z) \<Longrightarrow> Arcsin z = Arccos(csqrt(1 - z\<^sup>2))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3877
  by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3878
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3879
lemma Arccos_Arcsin_csqrt_pos:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3880
    "(0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z) \<Longrightarrow> Arccos z = Arcsin(csqrt(1 - z\<^sup>2))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3881
  by (simp add: Arcsin_def Arccos_def Complex.csqrt_square add.commute)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3882
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3883
lemma sin_Arccos:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3884
    "0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z \<Longrightarrow> sin(Arccos z) = csqrt(1 - z\<^sup>2)"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3885
  by (simp add: Arccos_Arcsin_csqrt_pos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3886
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3887
lemma cos_Arcsin:
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3888
    "0 < Re z \<or> Re z = 0 \<and> 0 \<le> Im z \<Longrightarrow> cos(Arcsin z) = csqrt(1 - z\<^sup>2)"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3889
  by (simp add: Arcsin_Arccos_csqrt_pos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3890
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3891
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3892
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationship with Arcsin on the Real Numbers\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3893
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3894
lemma of_real_arcsin: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arcsin x) = Arcsin(of_real x)"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3895
  by (smt (verit, best) Arcsin_sin Im_complex_of_real Re_complex_of_real arcsin sin_of_real)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3896
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3897
lemma Im_Arcsin_of_real: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> Im (Arcsin (of_real x)) = 0"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3898
  by (metis Im_complex_of_real of_real_arcsin)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3899
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3900
corollary\<^marker>\<open>tag unimportant\<close> Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3901
  by (metis Im_Arcsin_of_real Re_complex_of_real Reals_cases complex_is_Real_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3902
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3903
lemma arcsin_eq_Re_Arcsin: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> arcsin x = Re (Arcsin (of_real x))"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3904
  by (metis Re_complex_of_real of_real_arcsin)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3905
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3906
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3907
subsection\<^marker>\<open>tag unimportant\<close>\<open>Relationship with Arccos on the Real Numbers\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3908
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3909
lemma of_real_arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arccos x) = Arccos(of_real x)"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3910
  by (smt (verit, del_insts) Arccos_unique Im_complex_of_real Re_complex_of_real arccos_lbound 
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3911
      arccos_ubound cos_arccos_abs cos_of_real)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3912
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3913
lemma Im_Arccos_of_real: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> Im (Arccos (of_real x)) = 0"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3914
  by (metis Im_complex_of_real of_real_arccos)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3915
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3916
corollary\<^marker>\<open>tag unimportant\<close> Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>"
76819
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3917
  by (metis Im_Arccos_of_real complex_is_Real_iff of_real_Re)
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3918
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3919
lemma arccos_eq_Re_Arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> arccos x = Re (Arccos (of_real x))"
fc4ad2a2b6b1 reorganisation and simplification of theorems about transcendental functions
paulson <lp15@cam.ac.uk>
parents: 76724
diff changeset
  3920
  by (metis Re_complex_of_real of_real_arccos)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3921
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69986
diff changeset
  3922
subsection\<^marker>\<open>tag unimportant\<close>\<open>Continuity results for arcsin and arccos\<close>
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3923
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3924
lemma continuous_on_Arcsin_real [continuous_intros]:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3925
    "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} Arcsin"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3926
proof -
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3927
  have "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (arcsin (Re x))) =
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3928
        continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (Re (Arcsin (of_real (Re x)))))"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3929
    by (rule continuous_on_cong [OF refl]) (simp add: arcsin_eq_Re_Arcsin)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3930
  also have "\<dots> = ?thesis"
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3931
    by (rule continuous_on_cong [OF refl]) simp
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3932
  finally show ?thesis
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3933
    using continuous_on_arcsin [OF continuous_on_Re [OF continuous_on_id], of "{w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}"]
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3934
          continuous_on_of_real
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3935
    by fastforce
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3936
qed
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3937
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3938
lemma continuous_within_Arcsin_real:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3939
    "continuous (at z within {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}) Arcsin"
77324
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  3940
  using closed_real_abs_le continuous_on_Arcsin_real continuous_on_eq_continuous_within 
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  3941
        continuous_within_closed_nontrivial by blast
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3942
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3943
lemma continuous_on_Arccos_real:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3944
    "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} Arccos"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3945
proof -
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3946
  have "continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (arccos (Re x))) =
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3947
        continuous_on {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1} (\<lambda>x. complex_of_real (Re (Arccos (of_real (Re x)))))"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3948
    by (rule continuous_on_cong [OF refl]) (simp add: arccos_eq_Re_Arccos)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3949
  also have "\<dots> = ?thesis"
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3950
    by (rule continuous_on_cong [OF refl]) simp
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3951
  finally show ?thesis
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3952
    using continuous_on_arccos [OF continuous_on_Re [OF continuous_on_id], of "{w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}"]
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3953
          continuous_on_of_real
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3954
    by fastforce
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3955
qed
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3956
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3957
lemma continuous_within_Arccos_real:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3958
    "continuous (at z within {w \<in> \<real>. \<bar>Re w\<bar> \<le> 1}) Arccos"
77324
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  3959
  using closed_real_abs_le continuous_on_Arccos_real continuous_on_eq_continuous_within 
66c7ec736c36 Simplified some more proofs
paulson <lp15@cam.ac.uk>
parents: 77279
diff changeset
  3960
        continuous_within_closed_nontrivial by blast
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3961
67578
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3962
lemma sinh_ln_complex: "x \<noteq> 0 \<Longrightarrow> sinh (ln x :: complex) = (x - inverse x) / 2"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3963
  by (simp add: sinh_def exp_minus scaleR_conv_of_real exp_of_real)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3964
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3965
lemma cosh_ln_complex: "x \<noteq> 0 \<Longrightarrow> cosh (ln x :: complex) = (x + inverse x) / 2"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3966
  by (simp add: cosh_def exp_minus scaleR_conv_of_real)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3967
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3968
lemma tanh_ln_complex: "x \<noteq> 0 \<Longrightarrow> tanh (ln x :: complex) = (x ^ 2 - 1) / (x ^ 2 + 1)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3969
  by (simp add: tanh_def sinh_ln_complex cosh_ln_complex divide_simps power2_eq_square)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3970
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3971
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3972
subsection\<open>Roots of unity\<close>
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3973
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3974
theorem complex_root_unity:
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3975
  fixes j::nat
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3976
  assumes "n \<noteq> 0"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3977
    shows "exp(2 * of_real pi * \<i> * of_nat j / of_nat n)^n = 1"
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3978
  by (metis assms bot_nat_0.not_eq_extremum exp_divide_power_eq exp_of_nat2_mult exp_two_pi_i power_one)
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3979
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3980
lemma complex_root_unity_eq:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3981
  fixes j::nat and k::nat
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3982
  assumes "1 \<le> n"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3983
    shows "(exp(2 * of_real pi * \<i> * of_nat j / of_nat n) = exp(2 * of_real pi * \<i> * of_nat k / of_nat n)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3984
           \<longleftrightarrow> j mod n = k mod n)"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3985
proof -
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3986
    have "(\<exists>z::int. \<i> * (of_nat j * (of_real pi * 2)) =
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3987
               \<i> * (of_nat k * (of_real pi * 2)) + \<i> * (of_int z * (of_nat n * (of_real pi * 2)))) \<longleftrightarrow>
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3988
          (\<exists>z::int. of_nat j * (\<i> * (of_real pi * 2)) =
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3989
              (of_nat k + of_nat n * of_int z) * (\<i> * (of_real pi * 2)))"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3990
      by (simp add: algebra_simps)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3991
    also have "\<dots> \<longleftrightarrow> (\<exists>z::int. of_nat j = of_nat k + of_nat n * (of_int z :: complex))"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3992
      by simp
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3993
    also have "\<dots> \<longleftrightarrow> (\<exists>z::int. of_nat j = of_nat k + of_nat n * z)"
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 72301
diff changeset
  3994
      by (metis (mono_tags, opaque_lifting) of_int_add of_int_eq_iff of_int_mult of_int_of_nat_eq)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3995
    also have "\<dots> \<longleftrightarrow> int j mod int n = int k mod int n"
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 64508
diff changeset
  3996
      by (auto simp: mod_eq_dvd_iff dvd_def algebra_simps)
76137
175e6d47e3af tidied a few ugly proofs
paulson <lp15@cam.ac.uk>
parents: 75494
diff changeset
  3997
    also have "\<dots> \<longleftrightarrow> j mod n = k mod n"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3998
      by (metis of_nat_eq_iff zmod_int)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3999
    finally have "(\<exists>z. \<i> * (of_nat j * (of_real pi * 2)) =
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4000
             \<i> * (of_nat k * (of_real pi * 2)) + \<i> * (of_int z * (of_nat n * (of_real pi * 2)))) \<longleftrightarrow> j mod n = k mod n" .
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4001
   note * = this
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4002
  show ?thesis
78475
a5f6d2fc1b1f More cosmetic changes
paulson <lp15@cam.ac.uk>
parents: 77324
diff changeset
  4003
    using assms by (simp add: exp_eq field_split_simps *)
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4004
qed
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4005
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4006
corollary bij_betw_roots_unity:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4007
    "bij_betw (\<lambda>j. exp(2 * of_real pi * \<i> * of_nat j / of_nat n))
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4008
              {..<n}  {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j. j < n}"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4009
  by (auto simp: bij_betw_def inj_on_def complex_root_unity_eq)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4010
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4011
lemma complex_root_unity_eq_1:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4012
  fixes j::nat and k::nat
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4013
  assumes "1 \<le> n"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4014
    shows "exp(2 * of_real pi * \<i> * of_nat j / of_nat n) = 1 \<longleftrightarrow> n dvd j"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4015
proof -
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4016
  have "1 = exp(2 * of_real pi * \<i> * (of_nat n / of_nat n))"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4017
    using assms by simp
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4018
  then have "exp(2 * of_real pi * \<i> * (of_nat j / of_nat n)) = 1 \<longleftrightarrow> j mod n = n mod n"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4019
     using complex_root_unity_eq [of n j n] assms
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4020
     by simp
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4021
  then show ?thesis
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4022
    by auto
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4023
qed
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4024
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4025
lemma finite_complex_roots_unity_explicit:
77230
2d26af072990 Some basis results about trigonometric functions
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  4026
  "finite {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j::nat. j < n}"
2d26af072990 Some basis results about trigonometric functions
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  4027
  by simp
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4028
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4029
lemma card_complex_roots_unity_explicit:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4030
     "card {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j::nat. j < n} = n"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4031
  by (simp add:  Finite_Set.bij_betw_same_card [OF bij_betw_roots_unity, symmetric])
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4032
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4033
lemma complex_roots_unity:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4034
  assumes "1 \<le> n"
72301
c5d1a01d2d6c de-applying and tidying
paulson <lp15@cam.ac.uk>
parents: 71633
diff changeset
  4035
    shows "{z::complex. z^n = 1} = {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j. j < n}"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4036
  apply (rule Finite_Set.card_seteq [symmetric])
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4037
  using assms
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4038
  apply (auto simp: card_complex_roots_unity_explicit finite_roots_unity complex_root_unity card_roots_unity)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4039
  done
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4040
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4041
lemma card_complex_roots_unity: "1 \<le> n \<Longrightarrow> card {z::complex. z^n = 1} = n"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4042
  by (simp add: card_complex_roots_unity_explicit complex_roots_unity)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4043
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4044
lemma complex_not_root_unity:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4045
    "1 \<le> n \<Longrightarrow> \<exists>u::complex. norm u = 1 \<and> u^n \<noteq> 1"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4046
  apply (rule_tac x="exp (of_real pi * \<i> * of_real (1 / n))" in exI)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4047
  apply (auto simp: Re_complex_div_eq_0 exp_of_nat_mult [symmetric] mult_ac exp_Euler)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4048
  done
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4049
82518
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4050
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4051
subsection \<open>Normalisation of angles\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4052
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4053
text \<open>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4054
  The following operation normalises an angle $\varphi$, i.e.\ returns the unique
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4055
  $\psi$ in the range $(-\pi, \pi]$ such that
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4056
  $\varphi\equiv\psi\hskip.5em(\text{mod}\ 2\pi)$.
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4057
  This is the same convention used by the \<^const>\<open>Arg\<close> function.
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4058
\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4059
definition normalize_angle :: "real \<Rightarrow> real" where
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4060
  "normalize_angle x = x - \<lceil>x / (2 * pi) - 1 / 2\<rceil> * (2 * pi)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4061
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4062
lemma normalize_angle_id [simp]:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4063
  assumes "x \<in> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4064
  shows   "normalize_angle x = x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4065
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4066
  have "-1 < x / (2 * pi) - 1 / 2" "x / (2 * pi) - 1 / 2 \<le> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4067
    using assms pi_gt3 by (auto simp: field_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4068
  hence "ceiling (x / (2 * pi) - 1 / 2) = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4069
    by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4070
  thus ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4071
    by (simp add: normalize_angle_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4072
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4073
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4074
lemma normalize_angle_normalized: "normalize_angle x \<in> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4075
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4076
  have "-1 < x / (2 * pi) - 1 / 2 - ceiling (x / (2 * pi) - 1 / 2)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4077
    by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4078
  moreover have "x / (2 * pi) - 1 / 2 - ceiling (x / (2 * pi) - 1 / 2) \<le> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4079
    by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4080
  ultimately show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4081
    using pi_gt3 by (auto simp: field_simps normalize_angle_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4082
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4083
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4084
lemma cis_normalize_angle [simp]: "cis (normalize_angle x) = cis x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4085
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4086
  have "cis (normalize_angle x) = cis x / cis (real_of_int \<lceil>x / (2 * pi) - 1 / 2\<rceil> * (2 * pi))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4087
    by (simp add: normalize_angle_def flip: cis_divide)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4088
  also have "real_of_int \<lceil>x / (2 * pi) - 1 / 2\<rceil> * (2 * pi) =
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4089
               2 * pi * real_of_int \<lceil>x / (2 * pi) - 1 / 2\<rceil>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4090
    by (simp add: algebra_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4091
  also have "cis \<dots> = 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4092
    by (rule cis_multiple_2pi) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4093
  finally show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4094
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4095
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4096
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4097
lemma rcis_normalize_angle [simp]: "rcis r (normalize_angle x) = rcis r x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4098
  by (simp add: rcis_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4099
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4100
lemma normalize_angle_lbound [intro]: "normalize_angle x > -pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4101
  and normalize_angle_ubound [intro]: "normalize_angle x \<le> pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4102
  using normalize_angle_normalized[of x] by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4103
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4104
lemma normalize_angle_idem [simp]: "normalize_angle (normalize_angle x) = normalize_angle x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4105
  by (rule normalize_angle_id) (use normalize_angle_normalized[of x] in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4106
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4107
lemma Arg_rcis': "r > 0 \<Longrightarrow> Arg (rcis r \<phi>) = normalize_angle \<phi>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4108
  by (rule Arg_unique'[of r]) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4109
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4110
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4111
subsection \<open>Convexity of circular sectors in the complex plane\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4112
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4113
text \<open>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4114
  In this section we will show that if we have two non-zero points $w$ and $z$ in the complex plane
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4115
  whose arguments differ by less than $\pi$, then the argument of any point on the line connecting
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4116
  $w$ and $z$ lies between the arguments of $w$ and $z$. Moreover, the norm of any such point is
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4117
  no greater than the norms of $w$ and $z$.
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4118
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4119
  Geometrically, this means that if we have a sector around the origin with a central angle
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4120
  less than $\pi$ (minus the origin itself) then that region is convex.
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4121
\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4122
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4123
lemma Arg_closed_segment_aux1:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4124
  assumes "x \<noteq> 0" "y \<noteq> 0" "Re x > 0" "Re x = Re y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4125
  assumes "z \<in> closed_segment x y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4126
  shows   "Arg z \<in> closed_segment (Arg x) (Arg y)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4127
  using assms
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4128
proof (induction "Arg x" "Arg y" arbitrary: x y rule: linorder_wlog)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4129
  case (le x y)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4130
  from le have "Re z = Re x" "Im z \<in> closed_segment (Im x) (Im y)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4131
    by (auto simp: closed_segment_same_Re)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4132
  then obtain t where t: "t \<in> {0..1}" "Im z = linepath (Im x) (Im y) t"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4133
    by (metis image_iff linepath_image_01)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4134
  have *: "Im x \<le> Im y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4135
    using le by (auto simp: arg_conv_arctan arctan_le_iff field_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4136
  have "Im x / Re x \<le> linepath (Im x) (Im y) t / Re x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4137
    using le t * by (intro divide_right_mono linepath_real_ge_left) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4138
  hence "Arg x \<le> Arg z"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4139
    using t le \<open>Re z = Re x\<close> by (auto simp: arg_conv_arctan arctan_le_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4140
  moreover have "Im y / Re x \<ge> linepath (Im x) (Im y) t / Re x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4141
    using le t * by (intro divide_right_mono linepath_real_le_right) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4142
  hence "Arg y \<ge> Arg z"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4143
    using t le \<open>Re z = Re x\<close> by (auto simp: arg_conv_arctan arctan_le_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4144
  ultimately show ?case
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4145
    using le by (auto simp: closed_segment_same_Re closed_segment_eq_real_ivl)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4146
next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4147
  case (sym x y)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4148
  have "Arg z \<in> closed_segment (Arg y) (Arg x)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4149
    by (rule sym(1))
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4150
       (use sym(2-) in \<open>simp_all add: dist_commute closed_segment_commute\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4151
  thus ?case
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4152
    by (simp add: closed_segment_commute)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4153
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4154
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4155
lemma Arg_closed_segment_aux1':
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4156
  fixes r1 r2 \<phi>1 \<phi>2 :: real
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4157
  defines "x \<equiv> rcis r1 \<phi>1" and "y \<equiv> rcis r2 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4158
  assumes "z \<in> closed_segment x y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4159
  assumes "r1 > 0" "r2 > 0" "Re x = Re y" "Re x \<ge> 0" "Re x = 0 \<longrightarrow> Im x * Im y > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4160
  assumes "dist \<phi>1 \<phi>2 < pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4161
  obtains r \<phi> where "r \<in> {0<..max r1 r2}" "\<phi> \<in> closed_segment \<phi>1 \<phi>2" "z = rcis r \<phi>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4162
proof (cases "Re x = 0")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4163
  case True
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4164
  have [simp]: "cos \<phi>1 = 0" "cos \<phi>2 = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4165
    using assms True by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4166
  have "sin \<phi>1 = 1 \<and> sin \<phi>2 = 1 \<or> sin \<phi>1 = -1 \<and> sin \<phi>2 = -1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4167
    using sin_cos_squared_add[of \<phi>1] sin_cos_squared_add[of \<phi>2] assms
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4168
    by (auto simp: zero_less_mult_iff power2_eq_1_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4169
  thus ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4170
  proof (elim disjE conjE)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4171
    assume [simp]: "sin \<phi>1 = 1" "sin \<phi>2 = 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4172
    have xy_eq: "x = of_real r1 * \<i>" "y = of_real r2 * \<i>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4173
      by (auto simp: complex_eq_iff x_def y_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4174
    hence z: "Re z = 0" "Im z \<in> closed_segment r1 r2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4175
      using \<open>z \<in> closed_segment x y \<close> by (auto simp: xy_eq closed_segment_same_Re)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4176
    have "closed_segment r1 r2 \<subseteq> {0<..max r1 r2}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4177
      by (rule closed_segment_subset) (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4178
    with z have "Im z \<in> {0<..max r1 r2}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4179
      by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4180
    show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4181
      by (rule that[of "Im z" \<phi>1])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4182
         (use z \<open>Im z \<in> {0<..max r1 r2}\<close> in \<open>auto simp: complex_eq_iff\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4183
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4184
    assume [simp]: "sin \<phi>1 = -1" "sin \<phi>2 = -1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4185
    have xy_eq: "x = -of_real r1 * \<i>" "y = -of_real r2 * \<i>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4186
      by (auto simp: complex_eq_iff x_def y_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4187
    hence z: "Re z = 0" "Im z \<in> closed_segment (-r1) (-r2)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4188
      using \<open>z \<in> closed_segment x y \<close> by (auto simp: xy_eq closed_segment_same_Re)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4189
    have "closed_segment (-r1) (-r2) \<subseteq> {-max r1 r2..<0}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4190
      by (rule closed_segment_subset) (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4191
    with z have "-Im z \<in> {0<..max r1 r2}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4192
      by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4193
    show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4194
      by (rule that[of "-Im z" \<phi>1])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4195
         (use z \<open>-Im z \<in> {0<..max r1 r2}\<close> in \<open>auto simp: complex_eq_iff\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4196
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4197
next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4198
  case False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4199
  hence Re_pos: "Re x > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4200
    using \<open>Re x \<ge> 0\<close> by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4201
  define n :: int where "n = \<lceil>\<phi>1 / (2 * pi) - 1 / 2\<rceil>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4202
  define n' :: int where "n' = \<lceil>\<phi>2 / (2 * pi) - 1 / 2\<rceil>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4203
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4204
  have "Re z = Re x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4205
    using assms by (auto simp: closed_segment_same_Re)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4206
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4207
  have Arg_z: "Arg z \<in> closed_segment (Arg x) (Arg y)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4208
    by (rule Arg_closed_segment_aux1) (use assms Re_pos in \<open>simp_all add: dist_norm\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4209
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4210
  have "z \<in> closed_segment x y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4211
    by fact
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4212
  also have "\<dots> \<subseteq> cball 0 (max r1 r2)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4213
    using _ _ convex_cball by (rule closed_segment_subset) (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4214
  finally have "norm z \<le> max r1 r2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4215
    by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4216
  moreover have "z \<noteq> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4217
    by (intro notI) (use \<open>Re x > 0\<close> \<open>Re z = Re x\<close> in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4218
  ultimately have norm_z: "norm z \<in> {0<..max r1 r2}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4219
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4220
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4221
  have Arg_x: "Arg x = \<phi>1 - 2 * pi * of_int n"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4222
    using assms by (simp add: x_def Arg_rcis' normalize_angle_def n_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4223
  have Arg_y: "Arg y = \<phi>2 - 2 * pi * of_int n'"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4224
    using assms by (simp add: x_def Arg_rcis' normalize_angle_def n'_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4225
  have Arg_bounds: "\<bar>Arg x\<bar> \<le> pi/2" "\<bar>Arg y\<bar> \<le> pi/2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4226
    by (subst Arg_Re_nonneg; use assms in simp)+
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4227
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4228
  have "pi * of_int (2 * \<bar>n - n'\<bar> - 1) = 2 * pi * of_int (\<bar>n - n'\<bar>) - pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4229
    by (simp add: algebra_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4230
  also have "\<dots> = \<bar>2 * pi * of_int (n - n')\<bar> - pi / 2 - pi / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4231
    by (simp add: abs_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4232
  also have "\<dots> \<le> \<bar>2 * pi * of_int (n - n') + Arg x - Arg y\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4233
    using Arg_bounds pi_gt_zero by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4234
  also have "\<dots> \<le> dist \<phi>1 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4235
    using Arg_x Arg_y unfolding dist_norm real_norm_def by (simp add: algebra_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4236
  also have "\<dots> < pi * 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4237
    using assms by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4238
  finally have "2 * \<bar>n - n'\<bar> - 1 < 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4239
    by (subst (asm) mult_less_cancel_left_pos) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4240
  hence [simp]: "n' = n"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4241
    by presburger    
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4242
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4243
  show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4244
    using norm_z
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4245
  proof (rule that[of "norm z" "Arg z + 2 * pi * of_int n"])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4246
    have "2 * pi * of_int n + Arg z \<in> ((+) (2 * pi * of_int n)) ` closed_segment (Arg x) (Arg y)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4247
      using Arg_z by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4248
    also have "\<dots> = closed_segment (2 * pi * real_of_int n + Arg x) (2 * pi * real_of_int n + Arg y)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4249
      by (rule closed_segment_translation [symmetric])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4250
    also have "2 * pi * real_of_int n + Arg x = \<phi>1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4251
      by (simp add: Arg_x)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4252
    also have "2 * pi * real_of_int n + Arg y = \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4253
      by (simp add: Arg_y)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4254
    finally show "Arg z + 2 * pi * real_of_int n \<in> closed_segment \<phi>1 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4255
      by (simp add: add_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4256
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4257
    have "z = rcis (norm z) (Arg z)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4258
      by (simp add: rcis_cmod_Arg)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4259
    also have "\<dots> = rcis (cmod z) (Arg z + 2 * pi * real_of_int n)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4260
      by (simp add: rcis_def flip: cis_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4261
    finally show "z = \<dots>" .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4262
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4263
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4264
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4265
lemma Arg_closed_segment':
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4266
  fixes r1 r2 \<phi>1 \<phi>2 :: real
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4267
  defines "x \<equiv> rcis r1 \<phi>1" and "y \<equiv> rcis r2 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4268
  assumes "r1 > 0" "r2 > 0" "dist \<phi>1 \<phi>2 < pi" "z \<in> closed_segment x y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4269
  obtains r \<phi> where "r \<in> {0<..max r1 r2}" "\<phi> \<in> closed_segment \<phi>1 \<phi>2" "z = rcis r \<phi>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4270
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4271
  define u_aux :: real where
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4272
    "u_aux = (if Im x = Im y then pi/2 else arctan (Re (x-y) / Im (x-y)))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4273
  define u :: real where
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4274
    "u = (if Re (x * cis u_aux) < 0 then if u_aux \<le> 0 then u_aux + pi else u_aux - pi else u_aux)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4275
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4276
  have "u_aux \<in> {-pi/2<..pi/2}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4277
    using arctan_lbound[of "Re (x-y) / Im (x-y)"] arctan_ubound[of "Re (x-y) / Im (x-y)"]
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4278
    by (auto simp: u_aux_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4279
  have u_bounds: "u \<in> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4280
    using \<open>u_aux \<in> _\<close> by (auto simp: u_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4281
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4282
  have u_aux: "(Re x - Re y) * cos u_aux = (Im x - Im y) * sin u_aux"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4283
  proof (cases "Im x = Im y")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4284
    case False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4285
    hence "tan u_aux = (Re x - Re y) / (Im x - Im y)" and "cos u_aux \<noteq> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4286
      by (auto simp: u_aux_def tan_arctan)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4287
    thus ?thesis using False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4288
      by (simp add: tan_def divide_simps mult_ac split: if_splits)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4289
  qed (auto simp: u_aux_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4290
  hence "Re (x * cis u_aux) = Re (y * cis u_aux)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4291
    by (auto simp: algebra_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4292
  hence same_Re: "Re (x * cis u) = Re (y * cis u)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4293
    by (auto simp: u_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4294
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4295
  have Re_nonneg: "Re (x * cis u) \<ge> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4296
    by (auto simp: u_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4297
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4298
  have "linear (\<lambda>w. w * cis u)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4299
    by (intro linearI) (auto simp: algebra_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4300
  hence "closed_segment (x * cis u) (y * cis u) = (\<lambda>w. w * cis u) ` closed_segment x y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4301
    by (intro closed_segment_linear_image)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4302
  hence z'_in: "z * cis u \<in> closed_segment (x * cis u) (y * cis u)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4303
    using assms by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4304
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4305
  obtain r \<phi> where r\<phi>:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4306
    "r \<in> {0<..max r1 r2}" "\<phi> \<in> closed_segment (\<phi>1 + u) (\<phi>2 + u)" "z * cis u = rcis r \<phi>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4307
  proof (rule Arg_closed_segment_aux1'[of "z * cis u" r1 "\<phi>1 + u" r2 "\<phi>2 + u"])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4308
    show "z * cis u \<in> closed_segment (rcis r1 (\<phi>1 + u)) (rcis r2 (\<phi>2 + u))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4309
      using z'_in by (simp add: x_def y_def rcis_def mult.assoc flip: cis_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4310
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4311
    show "r1 > 0" "r2 > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4312
      by fact+
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4313
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4314
    show "Re (rcis r1 (\<phi>1 + u)) = Re (rcis r2 (\<phi>2 + u))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4315
      using same_Re by (simp add: x_def y_def cos_add field_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4316
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4317
    show "Re (rcis r1 (\<phi>1 + u)) \<ge> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4318
      using \<open>r1 > 0\<close> Re_nonneg by (auto intro!: mult_nonneg_nonneg simp: cos_add x_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4319
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4320
    show "dist (\<phi>1 + u) (\<phi>2 + u) < pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4321
      using assms by (simp add: dist_norm)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4322
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4323
    show "Re (rcis r1 (\<phi>1 + u)) = 0 \<longrightarrow> 0 < Im (rcis r1 (\<phi>1 + u)) * Im (rcis r2 (\<phi>2 + u))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4324
    proof
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4325
      assume *: "Re (rcis r1 (\<phi>1 + u)) = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4326
      hence "cos (\<phi>1 + u) = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4327
        using assms by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4328
      then obtain n1 where "\<phi>1 + u = real_of_int n1 * pi + pi / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4329
        by (subst (asm) cos_zero_iff_int2) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4330
      hence n1: "\<phi>1 = real_of_int n1 * pi + pi / 2 - u"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4331
        by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4332
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4333
      have "Re (rcis r1 (\<phi>1 + u)) = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4334
        by fact
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4335
      also have "rcis r1 (\<phi>1 + u) = x * cis u"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4336
        by (simp add: x_def rcis_def cis_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4337
      also have "Re (x * cis u) = Re (y * cis u)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4338
        by (fact same_Re)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4339
      also have "y * cis u = rcis r2 (\<phi>2 + u)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4340
        by (simp add: y_def rcis_def cis_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4341
      finally have "cos (\<phi>2 + u) = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4342
        using assms by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4343
      then obtain n2 where "\<phi>2 + u = real_of_int n2 * pi + pi / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4344
        by (subst (asm) cos_zero_iff_int2) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4345
      hence n2: "\<phi>2 = real_of_int n2 * pi + pi / 2 - u"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4346
        by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4347
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4348
      have "pi * real_of_int \<bar>n2 - n1\<bar> = \<bar>real_of_int (n2 - n1) * pi\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4349
        by (simp add: abs_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4350
      also have "\<dots> = dist \<phi>1 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4351
        by (simp add: n1 n2 dist_norm ring_distribs)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4352
      also have "\<dots> < pi * 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4353
        using \<open>dist \<phi>1 \<phi>2 < pi\<close> by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4354
      finally have "real_of_int \<bar>n2 - n1\<bar> < 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4355
        by (subst (asm) mult_less_cancel_left_pos) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4356
      hence "n1 = n2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4357
        by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4358
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4359
      have "Im (rcis r1 (\<phi>1 + u)) * Im (rcis r2 (\<phi>2 + u)) = r1 * r2 * cos (real_of_int n2 * pi) ^ 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4360
        by (simp add: n1 n2 sin_add \<open>n1 = n2\<close> power2_eq_square)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4361
      also have "cos (real_of_int n2 * pi) ^ 2 = (cos (2 * (real_of_int n2 * pi)) + 1) / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4362
        by (subst cos_double_cos) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4363
      also have "2 * (real_of_int n2 * pi) = 2 * pi * real_of_int n2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4364
        by (simp add: mult_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4365
      also have "(cos \<dots> + 1) / 2 = 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4366
        by (subst cos_int_2pin) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4367
      also have "r1 * r2 * 1 > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4368
        using assms by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4369
      finally show "Im (rcis r1 (\<phi>1 + u)) * Im (rcis r2 (\<phi>2 + u)) > 0" .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4370
    qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4371
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4372
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4373
  show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4374
  proof (rule that[of r "\<phi> - u"])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4375
    show "r \<in> {0<..max r1 r2}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4376
      by fact
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4377
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4378
    have "u + (\<phi> - u) \<in> closed_segment (\<phi>1 + u) (\<phi>2 + u)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4379
      using r\<phi> by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4380
    also have "\<dots> = (+) u ` closed_segment \<phi>1 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4381
      by (subst (1 2) add.commute, rule closed_segment_translation)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4382
    finally show "\<phi> - u \<in> closed_segment \<phi>1 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4383
      by (subst (asm) inj_image_mem_iff) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4384
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4385
    show "z = rcis r (\<phi> - u)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4386
      using r\<phi> by (simp add: rcis_def field_simps flip: cis_divide)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4387
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4388
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4389
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4390
lemma Arg_closed_segment:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4391
  assumes "x \<noteq> 0" "y \<noteq> 0" "dist (Arg x) (Arg y) < pi" "z \<in> closed_segment x y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4392
  shows   "Arg z \<in> closed_segment (Arg x) (Arg y)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4393
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4394
  define r1 \<phi>1 where "r1 = norm x" and "\<phi>1 = Arg x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4395
  define r2 \<phi>2 where "r2 = norm y" and "\<phi>2 = Arg y"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4396
  note defs = r1_def r2_def \<phi>1_def \<phi>2_def
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4397
  obtain r \<phi> where *: "r \<in> {0<..max r1 r2}" "\<phi> \<in> closed_segment \<phi>1 \<phi>2" "z = rcis r \<phi>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4398
    by (rule Arg_closed_segment'[of r1 r2 \<phi>1 \<phi>2 z])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4399
       (use assms in \<open>auto simp: defs rcis_cmod_Arg\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4400
  have "Arg z = \<phi>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4401
  proof (rule Arg_unique')
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4402
    show "z = rcis r \<phi>" "r > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4403
      using * by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4404
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4405
    have "\<phi> \<in> closed_segment \<phi>1 \<phi>2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4406
      by (fact *)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4407
    also have "\<dots> \<subseteq> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4408
      by (rule closed_segment_subset)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4409
         (use assms * Arg_bounded[of x] Arg_bounded[of y] in \<open>auto simp: defs\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4410
    finally show "\<phi> \<in> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4411
      by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4412
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4413
  with * show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4414
    by (simp add: defs)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4415
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4416
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4417
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4418
subsection \<open>Complex cones\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4419
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4420
text \<open>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4421
  We introduce the following notation to describe the set of all complex numbers of the form
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4422
  $c e^{ix}$ where $c\geq 0$ and $x\in [a, b]$.
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4423
\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4424
definition complex_cone :: "real \<Rightarrow> real \<Rightarrow> complex set" where
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4425
  "complex_cone a b = (\<lambda>(r,a). rcis r a) ` ({0..} \<times> closed_segment a b)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4426
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4427
lemma in_complex_cone_iff: "z \<in> complex_cone a b \<longleftrightarrow> (\<exists>x\<in>closed_segment a b. z = rcis (norm z) x)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4428
  by (auto simp: complex_cone_def image_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4429
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4430
lemma zero_in_complex_cone [simp, intro]: "0 \<in> complex_cone a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4431
  by (auto simp: in_complex_cone_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4432
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4433
lemma in_complex_cone_iff_Arg: 
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4434
  assumes "a \<in> {-pi<..pi}" "b \<in> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4435
  shows   "z \<in> complex_cone a b \<longleftrightarrow> z = 0 \<or> Arg z \<in> closed_segment a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4436
proof
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4437
  assume "z \<in> complex_cone a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4438
  then obtain r x where *: "x \<in> closed_segment a b" "z = rcis r x" "r \<ge> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4439
    by (auto simp: complex_cone_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4440
  have "closed_segment a b \<subseteq> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4441
    by (rule closed_segment_subset) (use assms in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4442
  with * have **: "x \<in> {-pi<..pi}"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4443
    by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4444
  show "z = 0 \<or> Arg z \<in> closed_segment a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4445
  proof (cases "z = 0")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4446
    case False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4447
    with * have "r \<noteq> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4448
      by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4449
    with * have [simp]: "r > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4450
      by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4451
    show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4452
      by (use * ** in \<open>auto simp: Arg_rcis\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4453
  qed auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4454
next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4455
  assume "z = 0 \<or> Arg z \<in> closed_segment a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4456
  thus "z \<in> complex_cone a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4457
  proof
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4458
    assume *: "Arg z \<in> closed_segment a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4459
    have "z = rcis (norm z) (Arg z)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4460
      by (simp_all add: rcis_cmod_Arg)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4461
    thus ?thesis using *
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4462
      unfolding in_complex_cone_iff by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4463
  qed auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4464
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4465
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4466
lemma complex_cone_rotate: "complex_cone (c + a) (c + b) = (*) (cis c) ` complex_cone a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4467
proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4468
  have *: "(*) (cis c) ` complex_cone a b \<subseteq> complex_cone (c + a) (c + b)" for c a b
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4469
    by (auto simp: closed_segment_translation in_complex_cone_iff norm_mult rcis_def simp flip: cis_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4470
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4471
  have "complex_cone (c + a) (c + b) = (*) (cis c) ` (*) (cis (-c)) ` complex_cone (c + a) (c + b)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4472
    by (simp add: image_image field_simps flip: cis_inverse)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4473
  also have "\<dots> \<subseteq> (*) (cis c) ` complex_cone ((-c) + (c + a)) ((-c) + (c + b))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4474
    by (intro image_mono *)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4475
  also have "\<dots> = (*) (cis c) ` complex_cone a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4476
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4477
  finally show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4478
    using *[of c a b] by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4479
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4480
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4481
lemma complex_cone_subset: 
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4482
  "a \<in> closed_segment a' b' \<Longrightarrow> b \<in> closed_segment a' b' \<Longrightarrow> complex_cone a b \<subseteq> complex_cone a' b'"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4483
  unfolding complex_cone_def 
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4484
  by (intro image_mono Sigma_mono order.refl, unfold subset_closed_segment) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4485
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4486
lemma complex_cone_commute: "complex_cone a b = complex_cone b a"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4487
  by (simp add: complex_cone_def closed_segment_commute)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4488
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4489
lemma complex_cone_eq_UNIV:
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4490
  assumes "dist a b \<ge> 2*pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4491
  shows   "complex_cone a b = UNIV"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4492
  using assms
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4493
proof (induction a b rule: linorder_wlog)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4494
  case (le a b)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4495
  have "bij ((*) (cis (a+pi)))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4496
    by (rule bij_betwI[of _ _ _ "(*) (cis (-a-pi))"])
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4497
       (auto simp: field_simps simp flip: cis_inverse cis_divide cis_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4498
  hence "UNIV = (*) (cis (a+pi)) ` UNIV"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4499
    unfolding bij_betw_def by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4500
  also have "UNIV \<subseteq> complex_cone (-pi) pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4501
  proof safe
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4502
    fix z :: complex
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4503
    have "z = rcis (norm z) (Arg z)" "norm z \<ge> 0" "Arg z \<in> closed_segment (-pi) pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4504
      using Arg_bounded[of z] by (auto simp: closed_segment_eq_real_ivl rcis_cmod_Arg)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4505
    thus "z \<in> complex_cone (-pi) pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4506
      unfolding in_complex_cone_iff by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4507
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4508
  also have "(*) (cis (a + pi)) ` complex_cone (- pi) pi = complex_cone a (a + 2 * pi)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4509
    using complex_cone_rotate[of "a+pi" "-pi" pi] by (simp add: add_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4510
  also have "\<dots> \<subseteq> complex_cone a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4511
    by (rule complex_cone_subset) (use le in \<open>auto simp: closed_segment_eq_real_ivl1 dist_norm\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4512
  finally show ?case by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4513
qed (simp_all add: complex_cone_commute dist_commute)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4514
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4515
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4516
text \<open>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4517
  A surprisingly tough argument: cones in the complex plane are closed.
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4518
\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4519
lemma closed_complex_cone [continuous_intros, intro]: "closed (complex_cone a b)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4520
proof (induction a b rule: linorder_wlog)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4521
  case (sym a b)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4522
  thus ?case
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4523
    by (simp add: complex_cone_commute)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4524
next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4525
  case (le a b)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4526
  show ?case
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4527
  proof (cases "b - a < 2 * pi")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4528
    case False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4529
    hence "complex_cone a b = UNIV"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4530
      by (intro complex_cone_eq_UNIV) (auto simp: dist_norm)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4531
    thus ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4532
      by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4533
  next
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4534
    case True
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4535
    define c where "c = (b - a) / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4536
    define d where "d = (b + a) / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4537
    have ab_eq: "a = d - c" "b = c + d"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4538
      by (simp_all add: c_def d_def field_simps)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4539
    have "c \<ge> 0" "c < pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4540
      using True le by (simp_all add: c_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4541
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4542
    define e where "e = (if c \<le> pi / 2 then 1 else sin c)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4543
    have "e > 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4544
    proof (cases "c \<le> pi / 2")
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4545
      case False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4546
      have "0 < pi / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4547
        by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4548
      also have "pi / 2 < c"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4549
        using False by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4550
      finally have "c > 0" .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4551
      moreover have "c < pi"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4552
        using True by (simp add: c_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4553
      ultimately show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4554
        using False by (auto simp: e_def intro!: sin_gt_zero)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4555
    qed (auto simp: e_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4556
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4557
    define A where "A = -ball 0 1 - {z. Re z < 0} \<inter> ({z. Im z < e} \<inter> {z. Im z > -e})"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4558
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4559
    have "closed (A \<inter> (Arg -` {-c..c}))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4560
    proof (intro continuous_closed_preimage)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4561
      show "closed A" unfolding A_def
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4562
        by (intro closed_Diff closed_Compl open_Int open_halfspace_Re_lt
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4563
                  open_halfspace_Im_lt open_halfspace_Im_gt open_ball)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4564
      show "continuous_on A Arg"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4565
        unfolding A_def using \<open>e > 0\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4566
        by (intro continuous_intros) (auto elim!: nonpos_Reals_cases)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4567
    qed auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4568
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4569
    also have "A \<inter> (Arg -` {-c..c}) =
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4570
                 (Arg -` {-c..c} - {z. Re z < 0} \<inter> ({z. Im z < e} \<inter> {z. Im z > -e})) - ball 0 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4571
      by (auto simp: A_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4572
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4573
    also have "\<dots> = Arg -` {-c..c} - ball 0 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4574
    proof (intro equalityI subsetI)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4575
      fix z assume z: "z \<in> Arg -` {-c..c} - ball 0 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4576
      define r where "r = norm z"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4577
      define x where "x = Arg z"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4578
      have "\<bar>x\<bar> \<le> c"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4579
        using z by (auto simp: x_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4580
      also note \<open>c < pi\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4581
      finally have "\<bar>x\<bar> < pi" .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4582
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4583
      have False if *: "Re z < 0" "Im z < e" "Im z > -e"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4584
      proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4585
        have "r \<ge> 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4586
          using z by (auto simp: r_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4587
        have z_eq: "z = rcis r x"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4588
          by (simp add: r_def x_def rcis_cmod_Arg)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4589
        from * and \<open>r \<ge> 1\<close> have "cos x < 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4590
          by (simp add: z_eq mult_less_0_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4591
        with \<open>\<bar>x\<bar> < pi\<close> have "\<bar>x\<bar> > pi / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4592
          using cos_ge_zero[of x] by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4593
        hence "c > pi / 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4594
          using \<open>\<bar>x\<bar> \<le> c\<close> by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4595
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4596
        have "sin c \<le> sin \<bar>x\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4597
        proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4598
          have "sin (pi - c) \<le> sin (pi - \<bar>x\<bar>)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4599
            by (rule sin_monotone_2pi_le)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4600
               (use \<open>\<bar>x\<bar> \<le> c\<close> \<open>\<bar>x\<bar> < pi\<close> \<open>\<bar>x\<bar> > pi / 2\<close> \<open>c < pi\<close> in \<open>auto simp: field_simps\<close>)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4601
          thus ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4602
            by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4603
        qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4604
        also have "sin \<bar>x\<bar> \<le> 1 * \<bar>sin x\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4605
          by (auto simp: abs_if)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4606
        also have "1 * \<bar>sin x\<bar> \<le> r * \<bar>sin x\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4607
          by (rule mult_right_mono) (use \<open>r \<ge> 1\<close> in auto)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4608
        also have "r * \<bar>sin x\<bar> = \<bar>Im z\<bar>"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4609
          using \<open>r \<ge> 1\<close> by (simp add: z_eq abs_mult)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4610
        also have "\<bar>Im z\<bar> < e"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4611
          using * by linarith
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4612
        finally show False
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4613
          using \<open>c > pi / 2\<close> by (auto simp: e_def split: if_splits)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4614
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4615
      qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4616
      thus "z \<in> Arg -` {-c..c} - {z. Re z < 0} \<inter> ({z. Im z < e} \<inter> {z. Im z > -e}) - ball 0 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4617
        using z by blast
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4618
    qed auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4619
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4620
    also have "Arg -` {-c..c} - ball 0 1 = complex_cone (-c) c - ball 0 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4621
      using \<open>c < pi\<close> \<open>c \<ge> 0\<close>
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4622
      by (auto simp: in_complex_cone_iff_Arg closed_segment_eq_real_ivl1)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4623
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4624
    finally have "closed (complex_cone (-c) c - ball 0 1)" .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4625
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4626
    moreover have "closed (complex_cone (-c) c \<inter> cball 0 1)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4627
    proof -
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4628
      have "compact ((\<lambda>(r,x). rcis r x) ` ({0..1} \<times> closed_segment (-c) c))"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4629
        by (intro compact_continuous_image)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4630
           (auto intro!: continuous_intros compact_Times simp: case_prod_unfold)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4631
      also have "((\<lambda>(r,x). rcis r x) ` ({0..1} \<times> closed_segment (-c) c)) = complex_cone (-c) c \<inter> cball 0 1"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4632
        by (auto simp: in_complex_cone_iff image_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4633
      finally show ?thesis
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4634
        by (rule compact_imp_closed)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4635
    qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4636
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4637
    ultimately have "closed (complex_cone (-c) c - ball 0 1 \<union> complex_cone (-c) c \<inter> cball 0 1)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4638
      by (intro closed_Un)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4639
    also have "\<dots> = complex_cone (-c) c"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4640
      by auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4641
    finally have "closed (complex_cone (-c) c)" .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4642
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4643
    hence "closed ((*) (cis d) ` complex_cone (-c) c)"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4644
      by (intro closed_injective_linear_image) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4645
    also have "\<dots> = complex_cone a b"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4646
      using complex_cone_rotate[of d "-c" c] by (simp add: ab_eq add_ac)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4647
    finally show ?thesis .
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4648
  qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4649
qed
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4650
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4651
lemma norm_eq_Re_iff: "norm z = Re z \<longleftrightarrow> z \<in> \<real>\<^sub>\<ge>\<^sub>0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4652
proof
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4653
  assume "norm z = Re z"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4654
  hence "norm z ^ 2 = Re z ^ 2"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4655
    by simp
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4656
  hence "Im z = 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4657
    by (auto simp: cmod_def)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4658
  moreover have "Re z \<ge> 0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4659
    by (subst \<open>norm z = Re z\<close> [symmetric]) auto
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4660
  ultimately show "z \<in> \<real>\<^sub>\<ge>\<^sub>0"
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4661
    by (auto simp: complex_nonneg_Reals_iff)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4662
qed (auto elim!: nonneg_Reals_cases)
da14e77a48b2 lots of lemmas for HOL, HOL-{Complex_}Analysis, HOL-Number_Theory
Manuel Eberl <manuel@pruvisto.org>
parents: 82459
diff changeset
  4663
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  4664
end