src/HOL/Analysis/Complex_Transcendental.thy
author Manuel Eberl <eberlm@in.tum.de>
Mon, 21 Aug 2017 20:49:15 +0200
changeset 66480 4b8d1df8933b
parent 66466 aec5d9c88d69
child 66611 c375b64a6c24
permissions -rw-r--r--
HOL-Analysis: Convergent FPS and infinite sums
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
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
     7
  Complex_Analysis_Basics
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
     8
  Summation_Tests
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66447
diff changeset
     9
   "HOL-Library.Periodic_Fun"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    10
begin
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
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 *)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    13
definition "moebius a b c d = (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    14
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    15
lemma 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)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    36
  have "r * x / \<bar>r\<bar> < sqrt (x*x + y*y)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    37
    apply (rule real_less_rsqrt)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    38
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    39
    apply (simp add: Complex power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    40
    using not_real_square_gt_zero by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    41
  then show ?thesis using assms Complex
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    42
    apply (auto simp: cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    43
    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
    44
    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
    45
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    46
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    47
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    48
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
    49
  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
    50
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    51
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    52
lemma cmod_square_less_1_plus:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    53
  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
    54
    shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    55
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    56
  apply (cases "Im z = 0 \<or> Re z = 0")
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    57
  using abs_square_less_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    58
    apply (force simp add: Re_power2 Im_power2 cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    59
  using cmod_diff_real_less [of "1 - z\<^sup>2" "1"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    60
  apply (simp add: norm_power Im_power2)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    61
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    62
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    63
subsection\<open>The Exponential Function is Differentiable and Continuous\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    65
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
    66
  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
    67
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
lemma continuous_within_exp:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
  fixes z::"'a::{real_normed_field,banach}"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
  shows "continuous (at z within s) exp"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
by (simp add: continuous_at_imp_continuous_within)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    72
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
    73
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
    74
  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
    75
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    76
lemma holomorphic_on_exp' [holomorphic_intros]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    77
  "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
    78
  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
    79
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    80
subsection\<open>Euler and de Moivre formulas.\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    81
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    82
text\<open>The sine series times @{term i}\<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
    83
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
    84
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
    85
  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
    86
    using sin_converges sums_mult by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
  then show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
    by (simp add: scaleR_conv_of_real field_simps)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
    91
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
    92
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
    93
  have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n)
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
    94
        = (\<lambda>n. (\<i> * z) ^ n /\<^sub>R (fact n))"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
  proof
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
    fix n
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
    97
    show "(cos_coeff n + \<i> * sin_coeff n) * z^n = (\<i> * z) ^ n /\<^sub>R (fact n)"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
      by (auto simp: cos_coeff_def sin_coeff_def scaleR_conv_of_real field_simps elim!: evenE oddE)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
  qed
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   100
  also have "... sums (exp (\<i> * z))"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
    by (rule exp_converges)
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   102
  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
   103
  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
   104
    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
   105
    by (simp add: field_simps scaleR_conv_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
  ultimately show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
    using sums_unique2 by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   108
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   110
corollary exp_minus_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
   111
  using exp_Euler [of "-z"]
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  by simp
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   114
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
   115
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   117
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
   118
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   120
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
   121
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   123
subsection\<open>Relationships between real and complex trig functions\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
lemma real_sin_eq [simp]:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  fixes x::real
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   127
  shows "Re(sin(of_real x)) = sin x"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
  by (simp add: sin_of_real)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   129
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
lemma real_cos_eq [simp]:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
  fixes x::real
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
  shows "Re(cos(of_real x)) = cos x"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
  by (simp add: cos_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   135
lemma DeMoivre: "(cos z + \<i> * sin z) ^ n = cos(n * z) + \<i> * sin(n * z)"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   136
  apply (simp add: exp_Euler [symmetric])
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
  by (metis exp_of_nat_mult mult.left_commute)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   139
lemma exp_cnj:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
  fixes z::complex
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  shows "cnj (exp z) = exp (cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
proof -
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
  have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) = (\<lambda>n. (cnj z)^n /\<^sub>R (fact n))"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
    by auto
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
  also have "... sums (exp (cnj z))"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
    by (rule exp_converges)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
  finally have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) sums (exp (cnj z))" .
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
  moreover have "(\<lambda>n. cnj (z ^ n /\<^sub>R (fact n))) sums (cnj (exp z))"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   149
    by (metis exp_converges sums_cnj)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
  ultimately show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
    using sums_unique2
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   152
    by blast
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
lemma cnj_sin: "cnj(sin z) = sin(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
  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
   157
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
lemma cnj_cos: "cnj(cos z) = cos(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
  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
   160
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   161
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
   162
  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
   163
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   164
lemma field_differentiable_within_sin: "sin 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
   165
  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
   166
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   167
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
   168
  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
   169
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   170
lemma field_differentiable_within_cos: "cos 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
   171
  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
   172
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
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
   174
  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
   175
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   176
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
   177
  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
   178
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   179
subsection\<open>Get a nice real/imaginary separation in Euler's formula.\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   181
lemma Euler: "exp(z) = of_real(exp(Re z)) *
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   182
              (of_real(cos(Im z)) + \<i> * of_real(sin(Im z)))"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   183
by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real Complex_eq)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   184
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  by (simp add: sin_exp_eq field_simps Re_divide Im_exp)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   187
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   188
lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   189
  by (simp add: sin_exp_eq field_simps Im_divide Re_exp)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   190
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   191
lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   192
  by (simp add: cos_exp_eq field_simps Re_divide Re_exp)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   193
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   194
lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   195
  by (simp add: cos_exp_eq field_simps Im_divide Im_exp)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   196
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   197
lemma Re_sin_pos: "0 < Re z \<Longrightarrow> Re z < pi \<Longrightarrow> Re (sin z) > 0"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   198
  by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   199
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   200
lemma Im_sin_nonneg: "Re z = 0 \<Longrightarrow> 0 \<le> Im z \<Longrightarrow> 0 \<le> Im (sin z)"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   201
  by (simp add: Re_sin Im_sin algebra_simps)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   202
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   203
lemma Im_sin_nonneg2: "Re z = pi \<Longrightarrow> Im z \<le> 0 \<Longrightarrow> 0 \<le> Im (sin z)"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   204
  by (simp add: Re_sin Im_sin algebra_simps)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   205
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   206
subsection\<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
   207
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   208
lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   209
  by (simp add: Complex_eq exp_add exp_Euler exp_of_real sin_of_real cos_of_real)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   210
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   211
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
   212
                 (is "?lhs = ?rhs")
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   213
proof 
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   214
  assume "exp z = 1"
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   215
  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
   216
    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
   217
  with \<open>?lhs\<close> show ?rhs
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   218
    by (metis Re_exp complex_Re_of_int cos_one_2pi_int exp_zero mult.commute mult_numeral_1 numeral_One of_int_mult of_int_numeral)
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   219
next
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   220
  assume ?rhs then show ?lhs
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   221
    using Im_exp Re_exp complex_Re_Im_cancel_iff by force
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   222
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   223
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   224
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
   225
                (is "?lhs = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   226
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   227
  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
   228
    by (simp add: exp_diff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   229
  also have "... \<longleftrightarrow> (Re w = Re z \<and> (\<exists>n::int. Im w - Im z = of_int (2 * n) * pi))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   230
    by (simp add: exp_eq_1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   231
  also have "... \<longleftrightarrow> ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   232
    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
   233
  finally show ?thesis .
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   234
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   235
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   236
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
   237
  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
   238
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   239
lemma exp_integer_2pi:
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   240
  assumes "n \<in> \<int>"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   241
  shows "exp((2 * n * pi) * \<i>) = 1"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   242
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   243
  have "exp((2 * n * pi) * \<i>) = exp 0"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   244
    using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   245
    by (simp only: Ints_def exp_eq) auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   246
  also have "... = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   247
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   248
  finally show ?thesis .
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   249
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   250
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   251
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
   252
  by (simp add: exp_eq)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   253
66466
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   254
lemma exp_integer_2pi_plus1:
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   255
  assumes "n \<in> \<int>"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   256
  shows "exp(((2 * n + 1) * pi) * \<i>) = - 1"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   257
proof -
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   258
  from assms obtain n' where [simp]: "n = of_int n'"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   259
    by (auto simp: Ints_def)
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   260
  have "exp(((2 * n + 1) * pi) * \<i>) = exp (pi * \<i>)"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   261
    using assms by (subst exp_eq) (auto intro!: exI[of _ n'] simp: algebra_simps)
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   262
  also have "... = - 1"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   263
    by simp
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   264
  finally show ?thesis .
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   265
qed
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   266
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   267
lemma inj_on_exp_pi:
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   268
  fixes z::complex shows "inj_on exp (ball z pi)"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   269
proof (clarsimp simp: inj_on_def exp_eq)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   270
  fix y n
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   271
  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
   272
         "dist z y < pi"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   273
  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
   274
    using dist_commute_lessI dist_triangle_less_add by blast
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   275
  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
   276
    by (simp add: dist_norm)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   277
  then show "n = 0"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   278
    by (auto simp: norm_mult)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   279
qed
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   280
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   281
lemma sin_cos_eq_iff: "sin y = sin x \<and> cos y = cos x \<longleftrightarrow> (\<exists>n::int. y = x + 2 * n * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   282
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   283
  { assume "sin y = sin x" "cos y = cos x"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   284
    then have "cos (y-x) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   285
      using cos_add [of y "-x"] by simp
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
   286
    then have "\<exists>n::int. y-x = n * 2 * pi"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   287
      using cos_one_2pi_int by blast }
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   288
  then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   289
  apply (auto simp: sin_add cos_add)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   290
  apply (metis add.commute diff_add_cancel mult.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   291
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   292
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   293
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   294
lemma exp_i_ne_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   295
  assumes "0 < x" "x < 2*pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   296
  shows "exp(\<i> * of_real x) \<noteq> 1"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   297
proof
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   298
  assume "exp (\<i> * of_real x) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   299
  then have "exp (\<i> * of_real x) = exp 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   300
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   301
  then obtain n where "\<i> * of_real x = (of_int (2 * n) * pi) * \<i>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   302
    by (simp only: Ints_def exp_eq) auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   303
  then have  "of_real x = (of_int (2 * n) * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   304
    by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   305
  then have  "x = (of_int (2 * n) * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   306
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   307
  then show False using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   308
    by (cases n) (auto simp: zero_less_mult_iff mult_less_0_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   309
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   310
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   311
lemma sin_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   312
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   313
  shows "sin z = 0 \<longleftrightarrow> (\<exists>n::int. z = of_real(n * pi))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   314
  by (simp add: sin_exp_eq exp_eq of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   315
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   316
lemma cos_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   317
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   318
  shows "cos z = 0 \<longleftrightarrow> (\<exists>n::int. z = of_real(n * pi) + of_real pi/2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   319
  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
   320
  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
   321
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   322
lemma cos_eq_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   323
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   324
  shows "cos z = 1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   325
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   326
  have "cos z = cos (2*(z/2))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   327
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   328
  also have "... = 1 - 2 * sin (z/2) ^ 2"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   329
    by (simp only: cos_double_sin)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   330
  finally have [simp]: "cos z = 1 \<longleftrightarrow> sin (z/2) = 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   331
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   332
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   333
    by (auto simp: sin_eq_0 of_real_numeral)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   334
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   335
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   336
lemma csin_eq_1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   337
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   338
  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
   339
  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
   340
  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
   341
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   342
lemma csin_eq_minus1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   343
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   344
  shows "sin z = -1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi) + 3/2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   345
        (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   346
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   347
  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
   348
    by (simp add: equation_minus_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   349
  also have "...  \<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
   350
    by (simp only: csin_eq_1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   351
  also have "...  \<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
   352
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   353
    by (metis (no_types)  is_num_normalize(8) minus_minus of_real_def real_vector.scale_minus_left uminus_add_conv_diff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   354
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   355
    apply (auto simp: of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   356
    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
   357
    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
   358
    apply (simp_all add: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   359
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   360
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   361
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   362
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   363
lemma ccos_eq_minus1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   364
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   365
  shows "cos z = -1 \<longleftrightarrow> (\<exists>n::int. z = of_real(2 * n * pi) + pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   366
  using csin_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
   367
  apply (simp add: sin_diff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   368
  apply (simp add: algebra_simps of_real_numeral equation_minus_iff)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   369
  done
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   370
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   371
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
   372
                (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   373
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   374
  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
   375
    by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   376
  also have "...  \<longleftrightarrow> (\<exists>n::int. complex_of_real x = 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
   377
    by (simp only: csin_eq_1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   378
  also have "...  \<longleftrightarrow> (\<exists>n::int. x = 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
   379
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   380
    apply (auto simp: algebra_simps of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   381
    apply (rule injD [OF inj_of_real [where 'a = complex]])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   382
    apply (auto simp: of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   383
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   384
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   385
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   386
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   387
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   388
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   389
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
   390
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   391
  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
   392
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   393
  also have "...  \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   394
    by (simp only: csin_eq_minus1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   395
  also have "...  \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + 3/2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   396
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   397
    apply (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   398
    apply (rule injD [OF inj_of_real [where 'a = complex]], auto)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   399
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   400
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   401
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   402
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   403
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   404
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   405
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
   406
                      (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   407
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   408
  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
   409
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   410
  also have "...  \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   411
    by (simp only: ccos_eq_minus1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   412
  also have "...  \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   413
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   414
    apply (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   415
    apply (rule injD [OF inj_of_real [where 'a = complex]], auto)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   416
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   417
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   418
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   419
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   420
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   421
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
   422
lemma dist_exp_i_1: "norm(exp(\<i> * of_real t) - 1) = 2 * \<bar>sin(t / 2)\<bar>"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   423
  apply (simp add: exp_Euler cmod_def power2_diff sin_of_real cos_of_real algebra_simps)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   424
  using cos_double_sin [of "t/2"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   425
  apply (simp add: real_sqrt_mult)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   426
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   427
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   428
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   429
lemma complex_sin_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   430
  fixes w :: complex
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   431
  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
   432
        (is "?lhs = ?rhs")
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   433
proof
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   434
  assume ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   435
  then have "sin w - sin z = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   436
    by (auto simp: algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   437
  then have "sin ((w - z) / 2)*cos ((w + z) / 2) = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   438
    by (auto simp: sin_diff_sin)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   439
  then consider "sin ((w - z) / 2) = 0" | "cos ((w + z) / 2) = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   440
    using mult_eq_0_iff by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   441
  then show ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   442
  proof cases
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   443
    case 1
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   444
    then show ?thesis
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   445
      apply (auto simp: sin_eq_0 algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   446
      by (metis Ints_of_int of_real_of_int_eq)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   447
  next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   448
    case 2
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   449
    then show ?thesis
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   450
      apply (auto simp: cos_eq_0 algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   451
      by (metis Ints_of_int of_real_of_int_eq)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   452
  qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   453
next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   454
  assume ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   455
  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
   456
                               w = -z + of_real ((2* of_int n + 1)*pi)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   457
    using Ints_cases by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   458
  then show ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   459
    using Periodic_Fun.sin.plus_of_int [of z n]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   460
    apply (auto simp: algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   461
    by (metis (no_types, hide_lams) add_diff_cancel_left add_diff_cancel_left' add_minus_cancel
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   462
              mult.commute sin.plus_of_int sin_minus sin_plus_pi)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   463
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   464
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   465
lemma complex_cos_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   466
  fixes w :: complex
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   467
  shows "cos w = cos z \<longleftrightarrow>
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   468
         (\<exists>n \<in> \<int>. w = z + of_real(2*n*pi) \<or> w = -z + of_real(2*n*pi))"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   469
        (is "?lhs = ?rhs")
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   470
proof
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   471
  assume ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   472
  then have "cos w - cos z = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   473
    by (auto simp: algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   474
  then have "sin ((w + z) / 2) * sin ((z - w) / 2) = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   475
    by (auto simp: cos_diff_cos)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   476
  then consider "sin ((w + z) / 2) = 0" | "sin ((z - w) / 2) = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   477
    using mult_eq_0_iff by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   478
  then show ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   479
  proof cases
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   480
    case 1
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   481
    then show ?thesis
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   482
      apply (auto simp: sin_eq_0 algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   483
      by (metis Ints_of_int of_real_of_int_eq)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   484
  next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   485
    case 2
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   486
    then show ?thesis
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   487
      apply (auto simp: sin_eq_0 algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   488
      by (metis Ints_of_int add_minus_cancel distrib_right mult_of_int_commute mult_zero_right of_int_0 of_int_add of_real_of_int_eq)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   489
  qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   490
next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   491
  assume ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   492
  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
   493
                               w = -z + of_real(2*n*pi)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   494
    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
   495
  then show ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   496
    using Periodic_Fun.cos.plus_of_int [of z n]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   497
    apply (auto simp: algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   498
    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
   499
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   500
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   501
lemma sin_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   502
   "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
   503
  using complex_sin_eq [of x y]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   504
  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
   505
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   506
lemma cos_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   507
   "cos x = cos y \<longleftrightarrow> (\<exists>n \<in> \<int>. x = y + 2*n*pi \<or> x = -y + 2*n*pi)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   508
  using complex_cos_eq [of x y]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   509
  by (simp only: cos_of_real Re_complex_of_real of_real_add [symmetric] of_real_minus [symmetric] of_real_mult [symmetric] of_real_eq_iff)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   510
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   511
lemma sinh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   512
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   513
  shows "(exp z - inverse (exp z)) / 2 = -\<i> * sin(\<i> * z)"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   514
  by (simp add: sin_exp_eq divide_simps exp_minus)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   515
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
   516
lemma sin_i_times:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   517
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   518
  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
   519
  using sinh_complex by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   520
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   521
lemma sinh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   522
  fixes x :: real
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   523
  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
   524
  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
   525
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   526
lemma cosh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   527
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   528
  shows "(exp z + inverse (exp z)) / 2 = cos(\<i> * z)"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   529
  by (simp add: cos_exp_eq divide_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
   530
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   531
lemma cosh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   532
  fixes x :: real
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   533
  shows "of_real((exp x + inverse (exp x)) / 2) = cos(\<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
   534
  by (simp add: cos_exp_eq divide_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
   535
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
   536
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
   537
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   538
lemma norm_cos_squared:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   539
    "norm(cos z) ^ 2 = cos(Re z) ^ 2 + (exp(Im z) - inverse(exp(Im z))) ^ 2 / 4"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   540
  apply (cases z)
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   541
  apply (simp add: cos_add cmod_power2 cos_of_real sin_of_real Complex_eq)
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
   542
  apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   543
  apply (simp only: left_diff_distrib [symmetric] power_mult_distrib)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   544
  apply (simp add: sin_squared_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   545
  apply (simp add: power2_eq_square algebra_simps divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   546
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   547
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   548
lemma norm_sin_squared:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   549
    "norm(sin z) ^ 2 = (exp(2 * Im z) + inverse(exp(2 * Im z)) - 2 * cos(2 * Re z)) / 4"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   550
  apply (cases z)
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   551
  apply (simp add: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double Complex_eq)
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
   552
  apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   553
  apply (simp only: left_diff_distrib [symmetric] power_mult_distrib)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   554
  apply (simp add: cos_squared_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   555
  apply (simp add: power2_eq_square algebra_simps divide_simps)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   556
  done
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   557
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   558
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
   559
  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
   560
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   561
lemma norm_cos_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   562
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   563
  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
   564
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   565
  have "Im z \<le> cmod z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   566
    using abs_Im_le_cmod abs_le_D1 by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   567
  with exp_uminus_Im show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   568
    apply (simp add: cos_exp_eq norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   569
    apply (rule order_trans [OF norm_triangle_ineq], simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   570
    apply (metis add_mono exp_le_cancel_iff mult_2_right)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   571
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   572
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   573
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   574
lemma norm_cos_plus1_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   575
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   576
  shows "norm(1 + cos z) \<le> 2 * exp(norm z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   577
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   578
  have mono: "\<And>u w z::real. (1 \<le> w | 1 \<le> z) \<Longrightarrow> (w \<le> u & z \<le> u) \<Longrightarrow> 2 + w + z \<le> 4 * u"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   579
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   580
  have *: "Im z \<le> cmod z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   581
    using abs_Im_le_cmod abs_le_D1 by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   582
  have triangle3: "\<And>x y z. norm(x + y + z) \<le> norm(x) + norm(y) + norm(z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   583
    by (simp add: norm_add_rule_thm)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   584
  have "norm(1 + cos z) = cmod (1 + (exp (\<i> * z) + exp (- (\<i> * z))) / 2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   585
    by (simp add: cos_exp_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   586
  also have "... = cmod ((2 + exp (\<i> * z) + exp (- (\<i> * z))) / 2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   587
    by (simp add: field_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   588
  also have "... = cmod (2 + exp (\<i> * z) + exp (- (\<i> * z))) / 2"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   589
    by (simp add: norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   590
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   591
    apply (rule ssubst, simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   592
    apply (rule order_trans [OF triangle3], simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   593
    using exp_uminus_Im *
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   594
    apply (auto intro: mono)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   595
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   596
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   597
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   598
subsection\<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
   599
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   600
declare power_Suc [simp del]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   601
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   602
lemma Taylor_exp_field:
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   603
  fixes z::"'a::{banach,real_normed_field}"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   604
  shows "norm (exp z - (\<Sum>i\<le>n. z ^ i / fact i)) \<le> exp (norm z) * (norm z ^ Suc n) / fact n"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   605
proof (rule field_taylor[of _ n "\<lambda>k. exp" "exp (norm z)" 0 z, simplified])
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   606
  show "convex (closed_segment 0 z)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   607
    by (rule convex_closed_segment [of 0 z])
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   608
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   609
  fix k x
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   610
  assume "x \<in> closed_segment 0 z" "k \<le> n"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   611
  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
   612
    using DERIV_exp DERIV_subset by blast
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   613
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   614
  fix x
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   615
  assume x: "x \<in> closed_segment 0 z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   616
  have "norm (exp x) \<le> exp (norm x)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   617
    by (rule norm_exp)
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   618
  also have "norm x \<le> norm z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   619
    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
   620
  finally show "norm (exp x) \<le> exp (norm z)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   621
    by simp
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   622
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   623
  show "0 \<in> closed_segment 0 z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   624
    by (auto simp: closed_segment_def)
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   625
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   626
  show "z \<in> closed_segment 0 z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   627
    apply (simp add: closed_segment_def scaleR_conv_of_real)
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   628
    using of_real_1 zero_le_one by blast
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   629
qed
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   630
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   631
lemma Taylor_exp:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   632
  "norm(exp z - (\<Sum>k\<le>n. z ^ k / (fact k))) \<le> exp\<bar>Re 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
   633
proof (rule complex_taylor [of _ n "\<lambda>k. exp" "exp\<bar>Re z\<bar>" 0 z, simplified])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   634
  show "convex (closed_segment 0 z)"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   635
    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
   636
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   637
  fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   638
  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
   639
  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
   640
    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
   641
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   642
  fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   643
  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
   644
  then show "Re x \<le> \<bar>Re z\<bar>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   645
    apply (auto simp: closed_segment_def scaleR_conv_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   646
    by (meson abs_ge_self abs_ge_zero linear mult_left_le_one_le mult_nonneg_nonpos order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   647
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   648
  show "0 \<in> closed_segment 0 z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   649
    by (auto simp: closed_segment_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   650
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   651
  show "z \<in> closed_segment 0 z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   652
    apply (simp add: closed_segment_def scaleR_conv_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   653
    using of_real_1 zero_le_one by blast
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   654
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   655
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   656
lemma
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   657
  assumes "0 \<le> u" "u \<le> 1"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   658
  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
   659
    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
   660
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   661
  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
   662
    by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   663
  show "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   664
    apply (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   665
    apply (rule order_trans [OF norm_triangle_ineq4])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   666
    apply (rule mono)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   667
    apply (auto simp: abs_if mult_left_le_one_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   668
    apply (meson mult_nonneg_nonneg neg_le_0_iff_le not_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   669
    apply (meson less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   670
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   671
  show "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   672
    apply (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   673
    apply (rule order_trans [OF norm_triangle_ineq])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   674
    apply (rule mono)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   675
    apply (auto simp: abs_if mult_left_le_one_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   676
    apply (meson mult_nonneg_nonneg neg_le_0_iff_le not_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   677
    apply (meson less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   678
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   679
qed
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   680
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   681
lemma Taylor_sin:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   682
  "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
   683
   \<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
   684
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   685
  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
   686
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   687
  have *: "cmod (sin z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   688
                 (\<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
   689
           \<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact 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
   690
  proof (rule complex_taylor [of "closed_segment 0 z" n
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
   691
                                 "\<lambda>k x. (-1)^(k div 2) * (if even k then sin x else cos x)"
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
   692
                                 "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
   693
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   694
    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
   695
            (- 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
   696
            (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
   697
      apply (auto simp: power_Suc)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   698
      apply (intro derivative_eq_intros | simp)+
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   699
      done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   700
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   701
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   702
    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
   703
    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
   704
      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
   705
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   706
  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
   707
            = (-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
   708
    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
   709
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   710
    apply (rule order_trans [OF _ *])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   711
    apply (simp add: **)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   712
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   713
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   714
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   715
lemma Taylor_cos:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   716
  "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
   717
   \<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
   718
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   719
  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
   720
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   721
  have *: "cmod (cos z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   722
                 (\<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
   723
           \<le> exp \<bar>Im z\<bar> * cmod z ^ Suc n / (fact n)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   724
  proof (rule complex_taylor [of "closed_segment 0 z" n "\<lambda>k x. (-1)^(Suc k div 2) * (if even k then cos x else sin x)" "exp\<bar>Im z\<bar>" 0 z,
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   725
simplified])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   726
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   727
    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
   728
    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
   729
            (- 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
   730
             (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
   731
      apply (auto simp: power_Suc)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   732
      apply (intro derivative_eq_intros | simp)+
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   733
      done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   734
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   735
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   736
    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
   737
    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
   738
      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
   739
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   740
  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
   741
            = (-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
   742
    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
   743
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   744
    apply (rule order_trans [OF _ *])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   745
    apply (simp add: **)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   746
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   747
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   748
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
   749
declare power_Suc [simp]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   750
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   751
text\<open>32-bit Approximation to e\<close>
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   752
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
   753
  using Taylor_exp [of 1 14] exp_le
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   754
  apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   755
  apply (simp only: pos_le_divide_eq [symmetric], linarith)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   756
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   757
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   758
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
   759
  using e_approx_32
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
   760
  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
   761
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   762
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
   763
  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
   764
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   765
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
   766
lemma ln3_gt_1: "ln 3 > (1::real)"
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   767
  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
   768
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   769
subsection\<open>The argument of a complex number\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   770
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   771
definition Arg :: "complex \<Rightarrow> real" where
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   772
 "Arg z \<equiv> if z = 0 then 0
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   773
           else THE t. 0 \<le> t \<and> t < 2*pi \<and>
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   774
                    z = of_real(norm z) * exp(\<i> * of_real t)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   775
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   776
lemma Arg_0 [simp]: "Arg(0) = 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   777
  by (simp add: Arg_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   778
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   779
lemma Arg_unique_lemma:
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   780
  assumes z:  "z = of_real(norm z) * exp(\<i> * of_real t)"
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   781
      and z': "z = of_real(norm z) * exp(\<i> * of_real t')"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   782
      and 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
   783
      and 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
   784
      and nz: "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   785
  shows "t' = t"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   786
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   787
  have [dest]: "\<And>x y z::real. x\<ge>0 \<Longrightarrow> x+y < z \<Longrightarrow> y<z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   788
    by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   789
  have "of_real (cmod z) * exp (\<i> * of_real t') = of_real (cmod z) * exp (\<i> * of_real t)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   790
    by (metis z z')
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   791
  then have "exp (\<i> * of_real t') = exp (\<i> * of_real t)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   792
    by (metis nz mult_left_cancel mult_zero_left z)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   793
  then have "sin t' = sin t \<and> cos t' = cos t"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   794
    apply (simp add: exp_Euler sin_of_real cos_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   795
    by (metis Complex_eq complex.sel)
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
   796
  then obtain n::int where n: "t' = t + 2 * n * pi"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   797
    by (auto simp: sin_cos_eq_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   798
  then have "n=0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   799
    apply (rule_tac z=n in int_cases)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   800
    using t t'
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   801
    apply (auto simp: mult_less_0_iff algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   802
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   803
  then show "t' = t"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   804
      by (simp add: n)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   805
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   806
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   807
lemma Arg: "0 \<le> Arg z & Arg z < 2*pi & z = of_real(norm z) * exp(\<i> * of_real(Arg z))"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   808
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   809
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   810
    by (simp add: Arg_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   811
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   812
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   813
  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
   814
             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
   815
    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
   816
    by blast
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   817
  have z: "z = of_real(norm z) * exp(\<i> * of_real t)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   818
    apply (rule complex_eqI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   819
    using t False ReIm
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   820
    apply (auto simp: exp_Euler sin_of_real cos_of_real divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   821
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   822
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   823
    apply (simp add: Arg_def False)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   824
    apply (rule theI [where a=t])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   825
    using t z False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   826
    apply (auto intro: Arg_unique_lemma)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   827
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   828
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   829
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   830
corollary
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   831
  shows Arg_ge_0: "0 \<le> Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   832
    and Arg_lt_2pi: "Arg z < 2*pi"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   833
    and Arg_eq: "z = of_real(norm z) * exp(\<i> * of_real(Arg z))"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   834
  using Arg by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   835
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
   836
lemma complex_norm_eq_1_exp: "norm z = 1 \<longleftrightarrow> exp(\<i> * of_real (Arg z)) = z"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
   837
  by (metis Arg_eq cis_conv_exp mult.left_neutral norm_cis of_real_1)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   838
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   839
lemma Arg_unique: "\<lbrakk>of_real r * exp(\<i> * of_real a) = z; 0 < r; 0 \<le> a; a < 2*pi\<rbrakk> \<Longrightarrow> Arg z = a"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   840
  apply (rule Arg_unique_lemma [OF _ Arg_eq])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   841
  using Arg [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   842
  apply (auto simp: norm_mult)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   843
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   844
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   845
lemma Arg_minus: "z \<noteq> 0 \<Longrightarrow> Arg (-z) = (if Arg z < pi then Arg z + pi else Arg z - pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   846
  apply (rule Arg_unique [of "norm z"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   847
  apply (rule complex_eqI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   848
  using Arg_ge_0 [of z] Arg_eq [of z] Arg_lt_2pi [of z] Arg_eq [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   849
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   850
  apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   851
  apply (metis Re_rcis Im_rcis rcis_def)+
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   852
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   853
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   854
lemma Arg_times_of_real [simp]: "0 < r \<Longrightarrow> Arg (of_real r * z) = Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   855
  apply (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   856
  apply (rule Arg_unique [of "r * norm z"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   857
  using Arg
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   858
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   859
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   860
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   861
lemma Arg_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg (z * of_real r) = Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   862
  by (metis Arg_times_of_real mult.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   863
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   864
lemma Arg_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg (z / of_real r) = Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   865
  by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   866
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   867
lemma Arg_le_pi: "Arg z \<le> pi \<longleftrightarrow> 0 \<le> Im z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   868
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   869
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   870
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   871
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   872
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   873
  have "0 \<le> Im z \<longleftrightarrow> 0 \<le> Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   874
    by (metis Arg_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   875
  also have "... = (0 \<le> Im (exp (\<i> * complex_of_real (Arg z))))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   876
    using False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   877
    by (simp add: zero_le_mult_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   878
  also have "... \<longleftrightarrow> Arg z \<le> pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   879
    by (simp add: Im_exp) (metis Arg_ge_0 Arg_lt_2pi sin_lt_zero sin_ge_zero not_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   880
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   881
    by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   882
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   883
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   884
lemma Arg_lt_pi: "0 < Arg z \<and> Arg z < pi \<longleftrightarrow> 0 < Im z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   885
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   886
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   887
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   888
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   889
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   890
  have "0 < Im z \<longleftrightarrow> 0 < Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   891
    by (metis Arg_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   892
  also have "... = (0 < Im (exp (\<i> * complex_of_real (Arg z))))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   893
    using False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   894
    by (simp add: zero_less_mult_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   895
  also have "... \<longleftrightarrow> 0 < Arg z \<and> Arg z < pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   896
    using Arg_ge_0  Arg_lt_2pi sin_le_zero sin_gt_zero
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   897
    apply (auto simp: Im_exp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   898
    using le_less apply fastforce
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   899
    using not_le by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   900
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   901
    by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   902
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   903
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   904
lemma Arg_eq_0: "Arg 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
   905
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   906
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   907
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   908
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   909
  case False
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   910
  have "z \<in> \<real> \<and> 0 \<le> Re z \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   911
    by (metis Arg_eq)
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   912
  also have "... \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg z)))"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   913
    using False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   914
    by (simp add: zero_le_mult_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   915
  also have "... \<longleftrightarrow> Arg z = 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   916
    apply (auto simp: Re_exp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   917
    apply (metis Arg_lt_pi Arg_ge_0 Arg_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   918
    using Arg_eq [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   919
    apply (auto simp: Reals_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   920
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   921
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   922
    by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   923
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   924
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
   925
corollary Arg_gt_0:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   926
  assumes "z \<in> \<real> \<Longrightarrow> Re z < 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   927
    shows "Arg z > 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   928
  using Arg_eq_0 Arg_ge_0 assms dual_order.strict_iff_order by fastforce
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   929
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   930
lemma Arg_of_real: "Arg(of_real x) = 0 \<longleftrightarrow> 0 \<le> x"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   931
  by (simp add: Arg_eq_0)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   932
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   933
lemma Arg_eq_pi: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   934
  apply  (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   935
  using Arg_eq_0 [of "-z"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   936
  apply (auto simp: complex_is_Real_iff Arg_minus)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   937
  apply (simp add: complex_Re_Im_cancel_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   938
  apply (metis Arg_minus pi_gt_zero add.left_neutral minus_minus minus_zero)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   939
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   940
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   941
lemma Arg_eq_0_pi: "Arg z = 0 \<or> Arg z = pi \<longleftrightarrow> z \<in> \<real>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   942
  using Arg_eq_0 Arg_eq_pi not_le by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   943
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   944
lemma Arg_inverse: "Arg(inverse z) = (if z \<in> \<real> \<and> 0 \<le> Re z then Arg z else 2*pi - Arg z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   945
  apply (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   946
  apply (rule Arg_unique [of "inverse (norm z)"])
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
   947
  using Arg_ge_0 [of z] Arg_lt_2pi [of z] Arg_eq [of z] Arg_eq_0 [of z] exp_two_pi_i
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   948
  apply (auto simp: of_real_numeral algebra_simps exp_diff divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   949
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   950
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   951
lemma Arg_eq_iff:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   952
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   953
     shows "Arg w = Arg z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   954
  using assms Arg_eq [of z] Arg_eq [of w]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   955
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   956
  apply (rule_tac x="norm w / norm z" in exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   957
  apply (simp add: divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   958
  by (metis mult.commute mult.left_commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   959
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   960
lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \<longleftrightarrow> Arg z = 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   961
  using complex_is_Real_iff
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   962
  apply (simp add: Arg_eq_0)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   963
  apply (auto simp: divide_simps not_sum_power2_lt_zero)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   964
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   965
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   966
lemma Arg_divide:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   967
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg w \<le> Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   968
    shows "Arg(z / w) = Arg z - Arg w"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   969
  apply (rule Arg_unique [of "norm(z / w)"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   970
  using assms Arg_eq [of z] Arg_eq [of w] Arg_ge_0 [of w] Arg_lt_2pi [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   971
  apply (auto simp: exp_diff norm_divide algebra_simps divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   972
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   973
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   974
lemma Arg_le_div_sum:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   975
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg w \<le> Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   976
    shows "Arg z = Arg w + Arg(z / w)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   977
  by (simp add: Arg_divide assms)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   978
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   979
lemma Arg_le_div_sum_eq:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   980
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   981
    shows "Arg w \<le> Arg z \<longleftrightarrow> Arg z = Arg w + Arg(z / w)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   982
  using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   983
  by (auto simp: Arg_ge_0 intro: Arg_le_div_sum)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   984
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   985
lemma Arg_diff:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   986
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   987
    shows "Arg w - Arg z = (if Arg z \<le> Arg w then Arg(w / z) else Arg(w/z) - 2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   988
  using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   989
  apply (auto simp: Arg_ge_0 Arg_divide not_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   990
  using Arg_divide [of w z] Arg_inverse [of "w/z"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   991
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   992
  by (metis Arg_eq_0 less_irrefl minus_diff_eq right_minus_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   993
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   994
lemma Arg_add:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   995
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   996
    shows "Arg w + Arg z = (if Arg w + Arg z < 2*pi then Arg(w * z) else Arg(w * z) + 2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   997
  using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   998
  using Arg_diff [of "w*z" z] Arg_le_div_sum_eq [of z "w*z"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   999
  apply (auto simp: Arg_ge_0 Arg_divide not_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1000
  apply (metis Arg_lt_2pi add.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1001
  apply (metis (no_types) Arg add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1002
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1003
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1004
lemma Arg_times:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1005
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1006
    shows "Arg (w * z) = (if Arg w + Arg z < 2*pi then Arg w + Arg z
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1007
                            else (Arg w + Arg z) - 2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1008
  using Arg_add [OF assms]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1009
  by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1010
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1011
lemma Arg_cnj: "Arg(cnj z) = (if z \<in> \<real> \<and> 0 \<le> Re z then Arg z else 2*pi - Arg z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1012
  apply (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1013
  apply (rule trans [of _ "Arg(inverse z)"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1014
  apply (simp add: Arg_eq_iff divide_simps complex_norm_square [symmetric] mult.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1015
  apply (metis norm_eq_zero of_real_power zero_less_power2)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1016
  apply (auto simp: of_real_numeral Arg_inverse)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1017
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1018
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1019
lemma Arg_real: "z \<in> \<real> \<Longrightarrow> Arg z = (if 0 \<le> Re z then 0 else pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1020
  using Arg_eq_0 Arg_eq_0_pi
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1021
  by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1022
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1023
lemma Arg_exp: "0 \<le> Im z \<Longrightarrow> Im z < 2*pi \<Longrightarrow> Arg(exp z) = Im z"
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
  1024
  by (rule Arg_unique [of  "exp(Re z)"]) (auto simp: exp_eq_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
  1025
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
  1026
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
  1027
  obtains r a::real where "z = complex_of_real r * (cos a + \<i> * sin a)" "0 \<le> r" "0 \<le> a" "a < 2*pi"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
  1028
  using Arg cis.ctr cis_conv_exp unfolding Complex_eq by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1029
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1030
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
  1031
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
  1032
  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
  1033
    apply (simp add: norm_mult cmod_unit_one)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1034
    by (metis (no_types, hide_lams) abs_le_D1 distrib_left mult.commute mult.left_commute mult_left_le)
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1035
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1036
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1037
subsection\<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
  1038
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1039
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
  1040
  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
  1041
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1042
lemma field_differentiable_at_tan: "~(cos z = 0) \<Longrightarrow> tan 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
  1043
  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
  1044
  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
  1045
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1046
lemma field_differentiable_within_tan: "~(cos z = 0)
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1047
         \<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
  1048
  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
  1049
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1050
lemma continuous_within_tan: "~(cos z = 0) \<Longrightarrow> continuous (at z within s) tan"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1051
  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
  1052
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1053
lemma continuous_on_tan [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 0)) \<Longrightarrow> continuous_on s tan"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1054
  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
  1055
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1056
lemma holomorphic_on_tan: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 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
  1057
  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
  1058
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1059
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1060
subsection\<open>Complex logarithms (the conventional principal value)\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1061
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
  1062
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
  1063
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
  1064
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
  1065
definition ln_complex :: "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
  1066
  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
  1067
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1068
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
  1069
lemma
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1070
  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
  1071
    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
  1072
      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
  1073
      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
  1074
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1075
  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
  1076
    using complex_unimodular_polar [of "z / (norm z)"] assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1077
    by (auto simp: norm_divide divide_simps)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1078
  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
  1079
    using sincos_principal_value [of "\<psi>"] assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1080
    by (auto simp: norm_divide divide_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
  1081
  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
  1082
    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
  1083
    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
  1084
    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
  1085
    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
  1086
  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
  1087
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1088
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1089
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1090
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
  1091
  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
  1092
    shows "ln(exp z) = z"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1093
  apply (rule exp_complex_eqI)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1094
  using assms mpi_less_Im_Ln  [of "exp z"] Im_Ln_le_pi [of "exp z"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1095
  apply auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1096
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1097
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1098
subsection\<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
  1099
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
  1100
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
  1101
  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
  1102
    shows "ln(of_real z::complex) = of_real(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
  1103
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
  1104
  have "ln(of_real (exp (ln z))::complex) = ln (exp (of_real (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
  1105
    by (simp add: exp_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
  1106
  also have "... = of_real(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
  1107
    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
  1108
    by (subst Ln_exp) 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
  1109
  finally 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
  1110
    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
  1111
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
  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
corollary Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<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
  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
  1115
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1116
corollary Im_Ln_of_real [simp]: "r > 0 \<Longrightarrow> Im (ln (of_real r)) = 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1117
  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
  1118
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
  1119
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
  1120
  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
  1121
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1122
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
  1123
  using Ln_of_real by force
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1124
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1125
lemma Ln_1 [simp]: "ln 1 = (0::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
  1126
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
  1127
  have "ln (exp 0) = (0::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
  1128
    by (metis (mono_tags, hide_lams) Ln_of_real exp_zero ln_one of_real_0 of_real_1 zero_less_one)
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
  1129
  then show ?thesis
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1130
    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
  1131
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
  1132
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1133
  
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1134
lemma Ln_eq_zero_iff [simp]: "x \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln x = 0 \<longleftrightarrow> x = 1" for x::complex
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1135
  by auto (metis exp_Ln exp_zero nonpos_Reals_zero_I)
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1136
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
  1137
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
  1138
  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
  1139
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
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
  1141
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
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
  1143
  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
  1144
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1145
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
  1146
  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
  1147
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1148
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
  1149
  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
  1150
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1151
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
  1152
  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
  1153
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
  1154
corollary 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
  1155
  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
  1156
    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
  1157
  using norm_exp [of "Ln z", simplified exp_Ln [OF z]]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1158
  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
  1159
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
  1160
proposition exists_complex_root:
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1161
  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
  1162
  assumes "n \<noteq> 0"  obtains w where "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
  1163
  apply (cases "z=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
  1164
  using assms apply (simp add: 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
  1165
  apply (rule_tac w = "exp(Ln z / n)" in that)
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
  apply (auto simp: assms exp_of_nat_mult [symmetric])
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1167
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1168
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
  1169
corollary exists_complex_root_nonzero:
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1170
  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
  1171
  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
  1172
  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
  1173
  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
  1174
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1175
subsection\<open>The Unwinding Number and the Ln-product Formula\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1176
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1177
text\<open>Note that in this special case the unwinding number is -1, 0 or 1.\<close>
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1178
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1179
definition unwinding :: "complex \<Rightarrow> complex" where
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1180
   "unwinding(z) = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1181
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1182
lemma unwinding_2pi: "(2*pi) * \<i> * unwinding(z) = z - Ln(exp z)"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1183
  by (simp add: unwinding_def)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1184
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1185
lemma Ln_times_unwinding:
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1186
    "w \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \<i> * unwinding(Ln w + Ln z)"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1187
  using unwinding_2pi by (simp add: exp_add)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1188
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1189
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1190
subsection\<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
  1191
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1192
lemma
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1193
  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
  1194
    shows has_field_derivative_Ln: "(Ln has_field_derivative inverse(z)) (at z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1195
      and Im_Ln_less_pi:           "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
  1196
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1197
  have znz: "z \<noteq> 0"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1198
    using assms by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1199
  then have "Im (Ln z) \<noteq> pi"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1200
    by (metis (no_types) Im_exp Ln_in_Reals assms complex_nonpos_Reals_iff complex_is_Real_iff exp_Ln mult_zero_right not_less pi_neq_zero sin_pi znz)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1201
  then show *: "Im (Ln z) < pi" using assms Im_Ln_le_pi
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1202
    by (simp add: le_neq_trans znz)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1203
  have "(exp has_field_derivative z) (at (Ln z))"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1204
    by (metis znz DERIV_exp exp_Ln)
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1205
  then show "(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
  1206
    apply (rule has_complex_derivative_inverse_strong_x
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1207
              [where s = "{w. -pi < Im(w) \<and> Im(w) < pi}"])
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1208
    using znz *
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1209
    apply (auto simp: Transcendental.continuous_on_exp [OF continuous_on_id] open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt mpi_less_Im_Ln)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1210
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1211
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1212
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1213
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
  1214
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
  1215
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1216
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
  1217
  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
  1218
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1219
lemma field_differentiable_within_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1220
         \<Longrightarrow> Ln 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
  1221
  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
  1222
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1223
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
  1224
  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
  1225
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1226
lemma isCont_Ln' [simp]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1227
   "\<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
  1228
  by (blast intro: isCont_o2 [OF _ continuous_at_Ln])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1229
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1230
lemma continuous_within_Ln: "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
  1231
  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
  1232
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1233
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
  1234
  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
  1235
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1236
lemma holomorphic_on_Ln: "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> Ln 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
  1237
  by (simp add: 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
  1238
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1239
lemma divide_ln_mono:
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1240
  fixes x y::real
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1241
  assumes "3 \<le> x" "x \<le> y"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1242
  shows "x / ln x \<le> y / ln y"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1243
proof (rule exE [OF complex_mvt_line [of x y "\<lambda>z. z / Ln z" "\<lambda>z. 1/(Ln z) - 1/(Ln z)^2"]];
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1244
    clarsimp simp add: closed_segment_Reals closed_segment_eq_real_ivl assms)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1245
  show "\<And>u. \<lbrakk>x \<le> u; u \<le> y\<rbrakk> \<Longrightarrow> ((\<lambda>z. z / Ln z) has_field_derivative 1 / Ln u - 1 / (Ln u)\<^sup>2) (at u)"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1246
    using \<open>3 \<le> x\<close> apply -
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1247
    apply (rule derivative_eq_intros | simp)+
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1248
    apply (force simp: field_simps power_eq_if)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1249
    done
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1250
  show "x / ln x \<le> y / ln y"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1251
    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
  1252
    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
  1253
  proof -
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1254
    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
  1255
      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
  1256
    show ?thesis
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1257
      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
  1258
  qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1259
qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1260
    
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1261
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1262
subsection\<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
  1263
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1264
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
  1265
  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
  1266
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1267
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1268
lemma Re_Ln_pos_lt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1269
  assumes "z \<noteq> 0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1270
    shows "\<bar>Im(Ln z)\<bar> < pi/2 \<longleftrightarrow> 0 < Re(z)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1271
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1272
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1273
    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
  1274
    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
  1275
      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
  1276
      by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1277
    then have "\<bar>Im w\<bar> < pi/2 \<longleftrightarrow> 0 < 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
  1278
      apply (auto simp: Re_exp zero_less_mult_iff cos_gt_zero_pi)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1279
      using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"]
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  1280
      apply (simp add: abs_if split: if_split_asm)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1281
      apply (metis (no_types) cos_minus cos_pi_half eq_divide_eq_numeral1(1) eq_numeral_simps(4)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1282
               less_numeral_extra(3) linorder_neqE_linordered_idom minus_mult_minus minus_mult_right
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1283
               mult_numeral_1_right)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1284
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1285
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1286
  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
  1287
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1288
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1289
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1290
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
  1291
  assumes "z \<noteq> 0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1292
    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
  1293
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1294
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1295
    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
  1296
    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
  1297
      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
  1298
      by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1299
    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
  1300
      apply (auto simp: Re_exp zero_le_mult_iff cos_ge_zero)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1301
      using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  1302
      apply (auto simp: abs_if split: if_split_asm)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1303
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1304
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1305
  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
  1306
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1307
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1308
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1309
lemma Im_Ln_pos_lt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1310
  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
  1311
    shows "0 < Im(Ln z) \<and> Im(Ln z) < pi \<longleftrightarrow> 0 < Im(z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1312
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1313
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1314
    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
  1315
    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
  1316
      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
  1317
      by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1318
    then have "0 < Im w \<and> Im w < pi \<longleftrightarrow> 0 < Im(exp w)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1319
      using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1320
      apply (auto simp: Im_exp zero_less_mult_iff)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1321
      using less_linear apply fastforce
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1322
      using less_linear apply fastforce
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1323
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1324
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1325
  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
  1326
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1327
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1328
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1329
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
  1330
  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
  1331
    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
  1332
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1333
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1334
    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
  1335
    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
  1336
      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
  1337
      by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1338
    then have "0 \<le> Im w \<and> Im w \<le> pi \<longleftrightarrow> 0 \<le> Im(exp w)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1339
      using sin_ge_zero [of "- (Im w)"] sin_ge_zero [of "(Im w)"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1340
      apply (auto simp: Im_exp zero_le_mult_iff sin_ge_zero)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1341
      apply (metis not_le not_less_iff_gr_or_eq pi_not_less_zero sin_eq_0_pi)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1342
      done }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1343
  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
  1344
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1345
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1346
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1347
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
  1348
  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
  1349
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1350
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
  1351
  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
  1352
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1353
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
  1354
lemma Im_Ln_eq_0: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = 0 \<longleftrightarrow> 0 < Re(z) \<and> Im(z) = 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
  1355
by (metis Im_complex_of_real Im_exp Ln_in_Reals Re_Ln_pos_lt Re_Ln_pos_lt_imp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1356
          Re_complex_of_real complex_is_Real_iff exp_Ln exp_of_real pi_gt_zero)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1357
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1358
lemma Im_Ln_eq_pi: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = pi \<longleftrightarrow> Re(z) < 0 \<and> Im(z) = 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
  1359
by (metis Im_Ln_eq_0 Im_Ln_pos_le Im_Ln_pos_lt add.left_neutral complex_eq less_eq_real_def
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1360
    mult_zero_right not_less_iff_gr_or_eq pi_ge_zero pi_neq_zero rcis_zero_arg rcis_zero_mod)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1361
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1362
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1363
subsection\<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
  1364
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1365
lemma cnj_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> cnj(Ln z) = Ln(cnj z)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1366
  apply (cases "z=0", auto)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1367
  apply (rule exp_complex_eqI)
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  1368
  apply (auto simp: abs_if split: if_split_asm)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1369
  using Im_Ln_less_pi Im_Ln_le_pi apply force
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1370
  apply (metis complex_cnj_zero_iff diff_minus_eq_add diff_strict_mono minus_less_iff
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1371
          mpi_less_Im_Ln mult.commute mult_2_right)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1372
  by (metis exp_Ln exp_cnj)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1373
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1374
lemma Ln_inverse: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Ln(inverse z) = -(Ln z)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1375
  apply (cases "z=0", auto)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1376
  apply (rule exp_complex_eqI)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1377
  using mpi_less_Im_Ln [of z] mpi_less_Im_Ln [of "inverse z"]
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  1378
  apply (auto simp: abs_if exp_minus split: if_split_asm)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1379
  apply (metis Im_Ln_less_pi Im_Ln_le_pi add.commute add_mono_thms_linordered_field(3) inverse_nonzero_iff_nonzero mult_2)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1380
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1381
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1382
lemma Ln_minus1 [simp]: "Ln(-1) = \<i> * pi"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1383
  apply (rule exp_complex_eqI)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1384
  using Im_Ln_le_pi [of "-1"] mpi_less_Im_Ln [of "-1"] cis_conv_exp cis_pi
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1385
  apply (auto simp: abs_if)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1386
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1387
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1388
lemma Ln_ii [simp]: "Ln \<i> = \<i> * of_real pi/2"
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1389
  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
  1390
  unfolding exp_Euler
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1391
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1392
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1393
lemma Ln_minus_ii [simp]: "Ln(-\<i>) = - (\<i> * pi/2)"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1394
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1395
  have  "Ln(-\<i>) = Ln(inverse \<i>)"    by simp
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1396
  also have "... = - (Ln \<i>)"         using Ln_inverse by blast
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1397
  also have "... = - (\<i> * pi/2)"     by simp
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1398
  finally show ?thesis .
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1399
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1400
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1401
lemma Ln_times:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1402
  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
  1403
    shows "Ln(w * z) =
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1404
                (if Im(Ln w + Ln z) \<le> -pi then
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1405
                  (Ln(w) + Ln(z)) + \<i> * of_real(2*pi)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1406
                else if Im(Ln w + Ln z) > pi then
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1407
                  (Ln(w) + Ln(z)) - \<i> * of_real(2*pi)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1408
                else 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
  1409
  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
  1410
  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
  1411
  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
  1412
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1413
corollary 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
  1414
    "\<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
  1415
         \<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
  1416
  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
  1417
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1418
corollary Ln_times_of_real:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1419
    "\<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
  1420
  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
  1421
  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
  1422
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1423
corollary Ln_divide_of_real:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1424
    "\<lbrakk>r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(z / of_real r) = Ln(z) - ln r"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1425
using Ln_times_of_real [of "inverse r" z]
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
  1426
by (simp add: ln_inverse Ln_of_real mult.commute divide_inverse of_real_inverse [symmetric]
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1427
         del: of_real_inverse)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1428
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1429
lemma Ln_minus:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1430
  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
  1431
    shows "Ln(-z) = (if Im(z) \<le> 0 \<and> ~(Re(z) < 0 \<and> Im(z) = 0)
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1432
                     then Ln(z) + \<i> * pi
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1433
                     else Ln(z) - \<i> * pi)" (is "_ = ?rhs")
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1434
  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
  1435
        Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1436
    by (fastforce simp: exp_add exp_diff 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
  1437
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1438
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
  1439
  assumes "z \<noteq> 0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1440
    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
  1441
proof (cases "z \<in> \<real>\<^sub>\<le>\<^sub>0")
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1442
  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
  1443
    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
  1444
next
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1445
  case True
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1446
  then have z: "Im z = 0" "Re z < 0"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1447
    using assms
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1448
    apply (auto simp: complex_nonpos_Reals_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1449
    by (metis complex_is_Real_iff le_imp_less_or_eq of_real_0 of_real_Re)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1450
  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
  1451
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1452
  also have "... = Ln (inverse (-z)) + \<i> * complex_of_real pi"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1453
    using assms z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1454
    apply (simp add: Ln_minus)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1455
    apply (simp add: field_simps)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1456
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1457
  also have "... = - Ln (- z) + \<i> * complex_of_real pi"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1458
    apply (subst Ln_inverse)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1459
    using z by (auto simp add: complex_nonneg_Reals_iff)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1460
  also have "... = - (Ln z) + \<i> * 2 * complex_of_real pi"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1461
    apply (subst Ln_minus [OF assms])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1462
    using assms z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1463
    apply simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1464
    done
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1465
  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
  1466
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1467
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1468
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
  1469
  assumes "z \<noteq> 0"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1470
    shows  "Ln(\<i> * z) = (if 0 \<le> Re(z) | Im(z) < 0
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1471
                          then Ln(z) + \<i> * of_real pi/2
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1472
                          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
  1473
  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
  1474
        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
  1475
  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
  1476
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65585
diff changeset
  1477
lemma Ln_of_nat [simp]: "0 < n \<Longrightarrow> Ln (of_nat n) = of_real (ln (of_nat n))"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1478
  by (subst of_real_of_nat_eq[symmetric], subst Ln_of_real[symmetric]) simp_all
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1479
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
  1480
lemma Ln_of_nat_over_of_nat:
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1481
  assumes "m > 0" "n > 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1482
  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
  1483
proof -
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1484
  have "of_nat m / of_nat n = (of_real (of_nat m / of_nat n) :: complex)" by simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1485
  also from assms have "Ln ... = of_real (ln (of_nat m / of_nat n))"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1486
    by (simp add: Ln_of_real[symmetric])
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1487
  also from assms have "... = of_real (ln (of_nat m) - ln (of_nat n))"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1488
    by (simp add: ln_div)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1489
  finally show ?thesis .
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1490
qed
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1491
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1492
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1493
subsection\<open>Relation between Ln and Arg, and hence continuity of Arg\<close>
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1494
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
  1495
lemma Arg_Ln:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1496
  assumes "0 < Arg z" shows "Arg z = Im(Ln(-z)) + pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1497
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
  1498
  case True
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1499
  with assms show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1500
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1501
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1502
  case False
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1503
  then have "z / of_real(norm z) = exp(\<i> * of_real(Arg z))"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1504
    using Arg [of z]
64240
eabf80376aab more standardized names
haftmann
parents: 63918
diff changeset
  1505
    by (metis abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1506
  then have "- z / of_real(norm z) = exp (\<i> * (of_real (Arg z) - pi))"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1507
    using cis_conv_exp cis_pi
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1508
    by (auto simp: exp_diff algebra_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1509
  then have "ln (- z / of_real(norm z)) = ln (exp (\<i> * (of_real (Arg z) - pi)))"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1510
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1511
  also have "... = \<i> * (of_real(Arg z) - pi)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1512
    using Arg [of z] assms pi_not_less_zero
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1513
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1514
  finally have "Arg z =  Im (Ln (- z / of_real (cmod z))) + pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1515
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1516
  also have "... = Im (Ln (-z) - ln (cmod z)) + pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1517
    by (metis diff_0_right minus_diff_eq zero_less_norm_iff Ln_divide_of_real False)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1518
  also have "... = Im (Ln (-z)) + pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1519
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1520
  finally show ?thesis .
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1521
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1522
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
  1523
lemma continuous_at_Arg:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1524
  assumes "z \<notin> \<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
  1525
    shows "continuous (at z) Arg"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1526
proof -
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1527
  have *: "isCont (\<lambda>z. Im (Ln (- z)) + pi) z"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1528
    by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1529
  have [simp]: "\<And>x. \<lbrakk>Im x \<noteq> 0\<rbrakk> \<Longrightarrow> Im (Ln (- x)) + pi = Arg x"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1530
      using Arg_Ln Arg_gt_0 complex_is_Real_iff by auto
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1531
  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
  1532
    using complex_nonneg_Reals_iff not_le by blast
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1533
  then have [simp]: "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg z"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1534
      using "*"  by (simp add: isCont_def) (metis Arg_Ln Arg_gt_0 complex_is_Real_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1535
  show ?thesis
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1536
      apply (simp add: continuous_at)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1537
      apply (rule Lim_transform_within_open [where s= "-\<real>\<^sub>\<ge>\<^sub>0" and f = "\<lambda>z. Im(Ln(-z)) + pi"])
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1538
      apply (auto simp add: not_le Arg_Ln [OF Arg_gt_0] complex_nonneg_Reals_iff closed_def [symmetric])
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1539
      using assms apply (force simp add: complex_nonneg_Reals_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1540
      done
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1541
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1542
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1543
lemma Ln_series:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1544
  fixes z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1545
  assumes "norm z < 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1546
  shows   "(\<lambda>n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\<lambda>n. ?f n * z^n) sums _")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1547
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1548
  let ?F = "\<lambda>z. \<Sum>n. ?f n * z^n" and ?F' = "\<lambda>z. \<Sum>n. diffs ?f n * z^n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1549
  have r: "conv_radius ?f = 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1550
    by (intro conv_radius_ratio_limit_nonzero[of _ 1])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1551
       (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1552
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1553
  have "\<exists>c. \<forall>z\<in>ball 0 1. ln (1 + z) - ?F z = c"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1554
  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
  1555
    fix z :: complex assume z': "z \<in> ball 0 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1556
    hence z: "norm z < 1" by (simp add: dist_0_norm)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  1557
    define t :: complex where "t = of_real (1 + norm z) / 2"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1558
    from z have t: "norm z < norm t" "norm t < 1" unfolding t_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1559
      by (simp_all add: field_simps norm_divide del: of_real_add)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1560
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1561
    have "Re (-z) \<le> norm (-z)" by (rule complex_Re_le_cmod)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1562
    also from z have "... < 1" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1563
    finally have "((\<lambda>z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1564
      by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1565
    moreover have "(?F has_field_derivative ?F' z) (at z)" using t r
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1566
      by (intro termdiffs_strong[of _ t] summable_in_conv_radius) 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
  1567
    ultimately have "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z))
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1568
                       (at z within ball 0 1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1569
      by (intro derivative_intros) (simp_all add: at_within_open[OF z'])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1570
    also have "(\<lambda>n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1571
      by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1572
    from sums_split_initial_segment[OF this, of 1]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1573
      have "(\<lambda>i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1574
    hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1575
    also have "inverse (1 + z) - inverse (1 + z) = 0" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1576
    finally show "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1577
  qed simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1578
  then obtain c where c: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> ln (1 + z) - ?F z = c" by blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1579
  from c[of 0] have "c = 0" by (simp only: powser_zero) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1580
  with c[of z] assms have "ln (1 + z) = ?F z" by (simp add: dist_0_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1581
  moreover have "summable (\<lambda>n. ?f n * z^n)" using assms r
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1582
    by (intro summable_in_conv_radius) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1583
  ultimately show ?thesis by (simp add: sums_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1584
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1585
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1586
lemma Ln_series': "cmod z < 1 \<Longrightarrow> (\<lambda>n. - ((-z)^n) / of_nat n) sums ln (1 + z)"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1587
  by (drule Ln_series) (simp add: power_minus')
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1588
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
  1589
lemma ln_series':
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1590
  assumes "abs (x::real) < 1"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1591
  shows   "(\<lambda>n. - ((-x)^n) / of_nat n) sums ln (1 + x)"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1592
proof -
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1593
  from assms have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1594
    by (intro Ln_series') simp_all
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1595
  also have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) = (\<lambda>n. complex_of_real (- ((-x)^n) / of_nat n))"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1596
    by (rule ext) simp
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
  1597
  also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))"
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1598
    by (subst Ln_of_real [symmetric]) simp_all
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1599
  finally show ?thesis by (subst (asm) sums_of_real_iff)
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1600
qed
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1601
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1602
lemma Ln_approx_linear:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1603
  fixes z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1604
  assumes "norm z < 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1605
  shows   "norm (ln (1 + z) - z) \<le> norm z^2 / (1 - norm z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1606
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1607
  let ?f = "\<lambda>n. (-1)^Suc n / of_nat n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1608
  from assms have "(\<lambda>n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1609
  moreover have "(\<lambda>n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1610
  ultimately have "(\<lambda>n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1611
    by (subst left_diff_distrib, intro sums_diff) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1612
  from sums_split_initial_segment[OF this, of "Suc 1"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1613
    have "(\<lambda>i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1614
    by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1615
  hence "(Ln (1 + z) - z) = (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1616
    by (simp add: sums_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1617
  also have A: "summable (\<lambda>n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1618
    by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1619
       (auto simp: assms field_simps intro!: always_eventually)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1620
  hence "norm (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \<le>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1621
             (\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1622
    by (intro summable_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1623
       (auto simp: norm_power norm_inverse norm_mult mult_ac simp 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
  1624
  also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \<le> norm ((-z)^2 * (-z)^i) * 1" for i
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1625
    by (intro mult_left_mono) (simp_all add: divide_simps)
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1626
  hence "(\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i)) \<le>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1627
           (\<Sum>i. norm (-(z^2) * (-z)^i))" using A assms
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1628
    apply (simp_all only: norm_power norm_inverse norm_divide norm_mult)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1629
    apply (intro suminf_le summable_mult summable_geometric)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1630
    apply (auto simp: norm_power field_simps simp 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
  1631
    done
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1632
  also have "... = norm z^2 * (\<Sum>i. norm z^i)" using assms
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1633
    by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1634
  also have "(\<Sum>i. norm z^i) = inverse (1 - norm z)" using assms
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1635
    by (subst suminf_geometric) (simp_all add: divide_inverse)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1636
  also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1637
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1638
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1639
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  1640
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1641
text\<open>Relation between Arg and arctangent in upper halfplane\<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
  1642
lemma Arg_arctan_upperhalf:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1643
  assumes "0 < Im z"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1644
    shows "Arg z = pi/2 - arctan(Re z / Im z)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1645
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
  1646
  case True with assms show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1647
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1648
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1649
  case False
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1650
  show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1651
    apply (rule Arg_unique [of "norm z"])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1652
    using False assms arctan [of "Re z / Im z"] pi_ge_two pi_half_less_two
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1653
    apply (auto simp: exp_Euler cos_diff sin_diff)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1654
    using norm_complex_def [of z, symmetric]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1655
    apply (simp add: sin_of_real cos_of_real sin_arctan cos_arctan field_simps real_sqrt_divide)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1656
    apply (metis complex_eq mult.assoc ring_class.ring_distribs(2))
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1657
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1658
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1659
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
  1660
lemma Arg_eq_Im_Ln:
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
  1661
  assumes "0 \<le> Im z" "0 < Re z"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1662
    shows "Arg z = Im (Ln z)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1663
proof (cases "z = 0 \<or> Im z = 0")
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1664
  case True then 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
  1665
    using assms Arg_eq_0 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
  1666
    apply auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1667
    by (metis Arg_eq_0_pi Arg_eq_pi Im_Ln_eq_0 Im_Ln_eq_pi less_numeral_extra(3) zero_complex.simps(1))
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1668
next
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
  1669
  case False
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1670
  then have "Arg z > 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1671
    using Arg_gt_0 complex_is_Real_iff by blast
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1672
  then 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
  1673
    using assms False
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1674
    by (subst Arg_Ln) (auto simp: Ln_minus)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1675
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1676
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
  1677
lemma continuous_within_upperhalf_Arg:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1678
  assumes "z \<noteq> 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1679
    shows "continuous (at z within {z. 0 \<le> Im z}) Arg"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1680
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
  1681
  case False then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1682
    using continuous_at_Arg continuous_at_imp_continuous_within by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1683
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1684
  case True
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1685
  then have z: "z \<in> \<real>" "0 < Re z"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1686
    using assms  by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1687
  then have [simp]: "Arg z = 0" "Im (Ln z) = 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1688
    by (auto simp: Arg_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
  1689
  show ?thesis
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1690
  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
  1691
    fix e::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1692
    assume "0 < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1693
    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
  1694
      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
  1695
    ultimately
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1696
    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
  1697
      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
  1698
    { fix x
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1699
      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
  1700
      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
  1701
        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
  1702
      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
  1703
        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
  1704
    }
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1705
    then show "\<exists>d>0. \<forall>x. 0 \<le> Im x \<longrightarrow> x \<noteq> z \<and> cmod (x - z) < d \<longrightarrow> \<bar>Arg x\<bar> < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1706
      apply (rule_tac x="min d (Re z / 2)" in exI)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1707
      using z d
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1708
      apply (auto simp: Arg_eq_Im_Ln)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1709
      done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1710
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1711
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1712
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1713
lemma continuous_on_upperhalf_Arg: "continuous_on ({z. 0 \<le> Im z} - {0}) Arg"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1714
  apply (auto simp: continuous_on_eq_continuous_within)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1715
  by (metis Diff_subset continuous_within_subset continuous_within_upperhalf_Arg)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1716
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
  1717
lemma open_Arg_less_Int:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1718
  assumes "0 \<le> s" "t \<le> 2*pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1719
    shows "open ({y. s < Arg y} \<inter> {y. Arg y < t})"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1720
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1721
  have 1: "continuous_on (UNIV - \<real>\<^sub>\<ge>\<^sub>0) Arg"
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
  1722
    using continuous_at_Arg continuous_at_imp_continuous_within
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1723
    by (auto simp: continuous_on_eq_continuous_within)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1724
  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
  1725
  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
  1726
    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
  1727
    by (metis greaterThan_def lessThan_def open_Int)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1728
  moreover have "{y. s < Arg y} \<inter> {y. Arg y < t} \<subseteq> - \<real>\<^sub>\<ge>\<^sub>0"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1729
    using assms by (auto simp: Arg_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
  1730
  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
  1731
    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
  1732
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1733
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1734
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1735
lemma open_Arg_gt: "open {z. t < Arg z}"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1736
proof (cases "t < 0")
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1737
  case True then have "{z. t < Arg z} = UNIV"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1738
    using Arg_ge_0 less_le_trans by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1739
  then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1740
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1741
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1742
  case False then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1743
    using open_Arg_less_Int [of t "2*pi"] Arg_lt_2pi
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1744
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1745
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1746
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1747
lemma closed_Arg_le: "closed {z. Arg z \<le> t}"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1748
  using open_Arg_gt [of t]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1749
  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
  1750
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1751
subsection\<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
  1752
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
  1753
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
  1754
  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
  1755
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
  1756
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
  1757
  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
  1758
  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
  1759
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
  1760
lemma norm_powr_real: "w \<in> \<real> \<Longrightarrow> 0 < Re w \<Longrightarrow> norm(w powr z) = exp(Re z * ln(Re w))"
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
  1761
  apply (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
  1762
  using Im_Ln_eq_0 complex_is_Real_iff norm_complex_def
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
  1763
  by auto
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
  1764
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
  1765
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
  1766
  fixes x::complex shows "x \<noteq> 0 \<Longrightarrow> x powr (of_nat n) = x^n"
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  1767
  by (induct n) (auto simp: ac_simps powr_add)
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  1768
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  1769
lemma powr_complexnumeral [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
  1770
  fixes x::complex shows "x \<noteq> 0 \<Longrightarrow> x powr (numeral n) = x ^ (numeral n)"
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  1771
  by (metis of_nat_numeral powr_complexpow)
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  1772
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1773
lemma cnj_powr:
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1774
  assumes "Im a = 0 \<Longrightarrow> Re a \<ge> 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1775
  shows   "cnj (a powr b) = cnj a powr cnj b"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1776
proof (cases "a = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1777
  case False
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1778
  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
  1779
  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
  1780
qed simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1781
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
  1782
lemma powr_real_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
  1783
    "\<lbrakk>w \<in> \<real>; z \<in> \<real>; 0 < Re w\<rbrakk> \<Longrightarrow> w powr z = exp(Re z * ln(Re w))"
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
  1784
  apply (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
  1785
  by (metis complex_eq complex_is_Real_iff diff_0 diff_0_right diff_minus_eq_add exp_ln exp_not_eq_zero
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
  1786
       exp_of_real Ln_of_real mult_zero_right of_real_0 of_real_mult)
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
  1787
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
  1788
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
  1789
  fixes x::real and y::real
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  1790
  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
  1791
  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
  1792
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1793
lemma powr_Reals_eq: "\<lbrakk>x \<in> \<real>; y \<in> \<real>; Re x > 0\<rbrakk> \<Longrightarrow> x powr y = of_real (Re x powr Re y)"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1794
  by (metis linear not_le of_real_Re powr_of_real)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1795
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
  1796
lemma norm_powr_real_mono:
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
  1797
    "\<lbrakk>w \<in> \<real>; 1 < Re w\<rbrakk>
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
  1798
     \<Longrightarrow> cmod(w powr z1) \<le> cmod(w powr z2) \<longleftrightarrow> Re z1 \<le> Re z2"
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
  1799
  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
  1800
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
  1801
lemma powr_times_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
  1802
    "\<lbrakk>x \<in> \<real>; y \<in> \<real>; 0 \<le> Re x; 0 \<le> Re y\<rbrakk>
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
  1803
           \<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
  1804
  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
  1805
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1806
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
  1807
  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
  1808
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1809
lemma
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1810
  fixes w::complex
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1811
  shows Reals_powr [simp]: "\<lbrakk>w \<in> \<real>\<^sub>\<ge>\<^sub>0; z \<in> \<real>\<rbrakk> \<Longrightarrow> w powr z \<in> \<real>"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1812
  and nonneg_Reals_powr [simp]: "\<lbrakk>w \<in> \<real>\<^sub>\<ge>\<^sub>0; z \<in> \<real>\<rbrakk> \<Longrightarrow> w powr z \<in> \<real>\<^sub>\<ge>\<^sub>0"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1813
  by (auto simp: nonneg_Reals_def Reals_def powr_of_real)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1814
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1815
lemma powr_neg_real_complex:
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1816
  shows   "(- of_real x) powr a = (-1) powr (of_real (sgn x) * a) * of_real x powr (a :: complex)"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1817
proof (cases "x = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1818
  assume x: "x \<noteq> 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1819
  hence "(-x) powr a = exp (a * ln (-of_real x))" by (simp add: powr_def)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1820
  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
  1821
    by (simp add: Ln_minus Ln_of_real)
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1822
  also from x have "exp (a * ...) = cis pi powr (of_real (sgn x) * a) * of_real x powr a"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1823
    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
  1824
  also note cis_pi
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1825
  finally show ?thesis by simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1826
qed simp_all
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1827
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
  1828
lemma has_field_derivative_powr:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1829
  fixes z :: complex
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1830
  shows "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ((\<lambda>z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)"
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
  1831
  apply (cases "z=0", auto)
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
  1832
  apply (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
  1833
  apply (rule DERIV_transform_at [where d = "norm z" and f = "\<lambda>z. exp (s * Ln 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
  1834
  apply (auto simp: dist_complex_def)
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1835
  apply (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
  1836
  apply (simp add: field_simps exp_diff)
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
  1837
  done
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
  1838
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1839
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
  1840
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1841
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
  1842
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
  1843
    "w \<noteq> 0 \<Longrightarrow> ((\<lambda>z. w powr z) has_field_derivative Ln w * w powr z) (at 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
  1844
  apply (simp add: powr_def)
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  1845
  apply (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
  1846
  done
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
  1847
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
  1848
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
  1849
  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
  1850
  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
  1851
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
  1852
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
  1853
lemma holomorphic_on_powr_right [holomorphic_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
  1854
    "f holomorphic_on s \<Longrightarrow> w \<noteq> 0 \<Longrightarrow> (\<lambda>z. w powr (f z)) holomorphic_on s"
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
  1855
  unfolding holomorphic_on_def field_differentiable_def
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  1856
  by (metis (full_types) DERIV_chain' has_field_derivative_powr_right)
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
  1857
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
  1858
lemma norm_powr_real_powr:
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1859
  "w \<in> \<real> \<Longrightarrow> 0 \<le> Re w \<Longrightarrow> cmod (w powr z) = Re w powr Re z"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63589
diff changeset
  1860
  by (cases "w = 0") (auto simp add: norm_powr_real powr_def Im_Ln_eq_0
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1861
                                     complex_is_Real_iff in_Reals_norm complex_eq_iff)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1862
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1863
lemma tendsto_ln_complex [tendsto_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1864
  assumes "(f \<longlongrightarrow> a) F" "a \<notin> \<real>\<^sub>\<le>\<^sub>0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1865
  shows   "((\<lambda>z. ln (f z :: complex)) \<longlongrightarrow> ln a) F"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1866
  using tendsto_compose[OF continuous_at_Ln[of a, unfolded isCont_def] assms(1)] assms(2) by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1867
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1868
lemma tendsto_powr_complex:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1869
  fixes f g :: "_ \<Rightarrow> complex"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1870
  assumes a: "a \<notin> \<real>\<^sub>\<le>\<^sub>0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1871
  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
  1872
  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
  1873
proof -
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1874
  from a have [simp]: "a \<noteq> 0" by auto
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1875
  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
  1876
    by (auto intro!: tendsto_intros simp: powr_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1877
  also {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1878
    have "eventually (\<lambda>z. z \<noteq> 0) (nhds a)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1879
      by (intro t1_space_nhds) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1880
    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
  1881
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1882
  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
  1883
    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
  1884
  finally show ?thesis .
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1885
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1886
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1887
lemma tendsto_powr_complex_0:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1888
  fixes f g :: "'a \<Rightarrow> complex"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1889
  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
  1890
  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
  1891
proof (rule tendsto_norm_zero_cancel)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1892
  define h where
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1893
    "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
  1894
  {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1895
    fix z :: 'a assume z: "f z \<noteq> 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1896
    define c where "c = abs (Im (g z)) * pi"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1897
    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
  1898
      have "abs (Im (Ln (f z))) \<le> pi" by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1899
    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
  1900
      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
  1901
    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
  1902
    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
  1903
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1904
  hence le: "norm (f z powr g z) \<le> h z" for z by (cases "f x = 0") (simp_all add: h_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1905
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1906
  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
  1907
    by (rule tendsto_mono[OF _ g]) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1908
  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
  1909
    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
  1910
  moreover {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1911
    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
  1912
      by (auto simp: filterlim_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1913
    hence "filterlim (\<lambda>x. norm (f x)) (principal {0<..})
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1914
             (inf F (principal {z. f z \<noteq> 0}))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1915
      by (rule filterlim_mono) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1916
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1917
  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
  1918
    by (simp add: filterlim_inf at_within_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1919
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1920
  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
  1921
    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
  1922
          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
  1923
  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
  1924
          -\<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
  1925
    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
  1926
  have C: "(h \<longlongrightarrow> 0) F" unfolding h_def
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1927
    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
  1928
       (insert B, auto simp: filterlim_uminus_at_bot algebra_simps)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1929
  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
  1930
    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
  1931
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1932
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1933
lemma tendsto_powr_complex' [tendsto_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1934
  fixes f g :: "_ \<Rightarrow> complex"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1935
  assumes fz: "a \<notin> \<real>\<^sub>\<le>\<^sub>0 \<or> (a = 0 \<and> Re b > 0)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1936
  assumes fg: "(f \<longlongrightarrow> a) F" "(g \<longlongrightarrow> b) F"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1937
  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
  1938
proof (cases "a = 0")
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1939
  case True
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1940
  with assms show ?thesis by (auto intro!: tendsto_powr_complex_0)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1941
next
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1942
  case False
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1943
  with assms show ?thesis by (auto intro!: tendsto_powr_complex elim!: nonpos_Reals_cases)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1944
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1945
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1946
lemma continuous_powr_complex:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1947
  assumes "f (netlimit F) \<notin> \<real>\<^sub>\<le>\<^sub>0" "continuous F f" "continuous F g"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1948
  shows   "continuous F (\<lambda>z. f z powr g z :: complex)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1949
  using assms unfolding continuous_def by (intro tendsto_powr_complex) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1950
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1951
lemma isCont_powr_complex [continuous_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1952
  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
  1953
  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
  1954
  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
  1955
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1956
lemma continuous_on_powr_complex [continuous_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1957
  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
  1958
  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
  1959
  assumes "continuous_on A f" "continuous_on A g"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1960
  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
  1961
  unfolding continuous_on_def
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1962
proof
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1963
  fix z assume z: "z \<in> A"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1964
  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
  1965
  proof (cases "f z = 0")
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1966
    case False
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1967
    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
  1968
    with assms(3,4) z show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1969
      by (intro tendsto_powr_complex')
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1970
         (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
  1971
  next
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1972
    case True
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1973
    with assms z show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1974
      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
  1975
  qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  1976
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
  1977
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1978
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1979
subsection\<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
  1980
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1981
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
  1982
  fixes s::complex
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1983
  assumes "0 < Re s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1984
    shows "((\<lambda>n. Ln n / (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
  1985
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
  1986
  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
  1987
  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
  1988
  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
  1989
  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
  1990
    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
  1991
      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
  1992
  next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1993
    fix x::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1994
    assume x: "2 / (e * (Re s)\<^sup>2) \<le> x"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1995
    then have "x>0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1996
    using e assms
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1997
      by (metis less_le_trans mult_eq_0_iff mult_pos_pos pos_less_divide_eq power2_eq_square
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1998
                zero_less_numeral)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1999
    then show "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
  2000
      using e assms x
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2001
      apply (auto simp: field_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2002
      apply (rule_tac y = "e * (x\<^sup>2 * (Re s)\<^sup>2)" in le_less_trans)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2003
      apply (auto simp: power2_eq_square field_simps add_pos_pos)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2004
      done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2005
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2006
  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
  2007
    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
  2008
  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
  2009
    using assms
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2010
    by (force intro: less_le_trans [OF _ exp_lower_taylor_quadratic])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2011
  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
  2012
    using e   by (auto simp: field_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2013
  with e show "\<exists>no. \<forall>n\<ge>no. norm (Ln (of_nat n) / of_nat n powr s) < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2014
    apply (auto simp: norm_divide norm_powr_real divide_simps)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61808
diff changeset
  2015
    apply (rule_tac x="nat \<lceil>exp xo\<rceil>" in exI)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2016
    apply clarify
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2017
    apply (drule_tac x="ln n" in spec)
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
  2018
    apply (auto simp: exp_less_mono nat_ceiling_le_eq not_le)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2019
    apply (metis exp_less_mono exp_ln not_le of_nat_0_less_iff)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2020
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2021
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2022
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2023
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
  2024
  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
  2025
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2026
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
  2027
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2028
  assumes "0 < s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2029
    shows "((\<lambda>n. ln n / (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
  2030
  using lim_Ln_over_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
  2031
  apply (subst filterlim_sequentially_Suc [symmetric])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2032
  apply (simp add: lim_sequentially dist_norm
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
  2033
          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
  2034
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2035
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2036
lemma lim_ln_over_n: "((\<lambda>n. ln(real_of_nat n) / of_nat n) \<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
  2037
  using lim_ln_over_power [of 1, THEN filterlim_sequentially_Suc [THEN iffD2]]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2038
  apply (subst filterlim_sequentially_Suc [symmetric])
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
  2039
  apply (simp add: lim_sequentially dist_norm)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2040
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2041
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2042
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
  2043
  assumes "0 < Re s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2044
    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
  2045
proof -
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2046
  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
  2047
    using ln_272_gt_1
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2048
    by (force intro: order_trans [of _ "ln (272/100)"])
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  2049
  moreover have "(\<lambda>n. cmod (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
  2050
    using lim_Ln_over_power [OF assms]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2051
    by (metis tendsto_norm_zero_iff)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2052
  ultimately show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2053
    apply (auto intro!: Lim_null_comparison [where g = "\<lambda>n. norm (Ln(of_nat n) / of_nat n powr s)"])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2054
    apply (auto simp: norm_divide divide_simps eventually_sequentially)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2055
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2056
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2057
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2058
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
  2059
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2060
  assumes "0 < s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2061
    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
  2062
  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
  2063
  apply (subst filterlim_sequentially_Suc [symmetric])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2064
  apply (simp add: lim_sequentially dist_norm)
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
  2065
  apply (simp add: 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
  2066
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2067
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2068
lemma lim_1_over_Ln: "((\<lambda>n. 1 / Ln(of_nat n)) \<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
  2069
proof (clarsimp simp add: lim_sequentially dist_norm norm_divide divide_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2070
  fix r::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2071
  assume "0 < r"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2072
  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
  2073
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2074
  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
  2075
    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
  2076
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2077
  then have "exp (inverse r) < of_nat n"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2078
    by (simp add: divide_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2079
  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
  2080
    by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2081
  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
  2082
    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
  2083
  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
  2084
    using neq0_conv by fastforce
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2085
  ultimately show "\<exists>no. \<forall>n. Ln (of_nat n) \<noteq> 0 \<longrightarrow> no \<le> n \<longrightarrow> 1 < r * cmod (Ln (of_nat n))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2086
    using n \<open>0 < r\<close>
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2087
    apply (rule_tac x=n in exI)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2088
    apply (auto simp: divide_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2089
    apply (erule less_le_trans, auto)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2090
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2091
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2092
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2093
lemma lim_1_over_ln: "((\<lambda>n. 1 / ln(real_of_nat n)) \<longlongrightarrow> 0) sequentially"
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  2094
  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
  2095
  apply (subst filterlim_sequentially_Suc [symmetric])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2096
  apply (simp add: lim_sequentially dist_norm)
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
  2097
  apply (simp add: 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
  2098
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2099
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2100
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
  2101
proof (rule Lim_transform_eventually)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2102
  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
  2103
  proof (rule Lim_transform_bound)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2104
    show "(inverse o real) \<longlonglongrightarrow> 0"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2105
      by (metis comp_def seq_harmonic tendsto_explicit)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2106
    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
  2107
    proof
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2108
      fix n::nat
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2109
      assume n: "3 \<le> n"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2110
      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
  2111
        by auto
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2112
      with ln3_gt_1 have "1/ ln n \<le> 1"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2113
        by (simp add: divide_simps)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2114
      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
  2115
        by (simp add: ln_add_one_self_le_self)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2116
      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
  2117
        by (intro mult_mono) (use n in auto)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2118
      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
  2119
        by (simp add: field_simps ln0)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2120
      qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2121
  qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2122
  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
  2123
    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
  2124
  show "\<forall>\<^sub>F k in sequentially. 1 + ln (1 + 1 / k) / ln k = ln(Suc k) / ln k"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2125
    by (simp add: divide_simps ln_div eventually_sequentiallyI [of 2])
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2126
qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2127
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2128
lemma lim_ln_over_ln1: "(\<lambda>n. ln n / ln(Suc n)) \<longlonglongrightarrow> 1"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2129
proof -
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2130
  have "(\<lambda>n. inverse (ln(Suc n) / ln n)) \<longlonglongrightarrow> inverse 1"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2131
    by (rule tendsto_inverse [OF lim_ln1_over_ln]) auto
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2132
  then show ?thesis
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2133
    by simp
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2134
qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2135
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
  2136
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2137
subsection\<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
  2138
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2139
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
  2140
  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
  2141
    shows "csqrt z = exp(Ln(z) / 2)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2142
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2143
  have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))"
64240
eabf80376aab more standardized names
haftmann
parents: 63918
diff changeset
  2144
    by (metis exp_double nonzero_mult_div_cancel_left times_divide_eq_right zero_neq_numeral)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2145
  also have "... = z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2146
    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
  2147
  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
  2148
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2149
  also have "... = exp (Ln z / 2)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2150
    apply (subst csqrt_square)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2151
    using cos_gt_zero_pi [of "(Im (Ln z) / 2)"] Im_Ln_le_pi mpi_less_Im_Ln assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2152
    apply (auto simp: Re_exp Im_exp zero_less_mult_iff zero_le_mult_iff, fastforce+)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2153
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2154
  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
  2155
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2156
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2157
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2158
lemma csqrt_inverse:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2159
  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
  2160
    shows "csqrt (inverse z) = inverse (csqrt z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2161
proof (cases "z=0", simp)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2162
  assume "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
  2163
  then show ?thesis
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2164
    using assms csqrt_exp_Ln Ln_inverse exp_minus
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2165
    by (simp add: csqrt_exp_Ln Ln_inverse exp_minus)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2166
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2167
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2168
lemma cnj_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2169
  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
  2170
    shows "cnj(csqrt z) = csqrt(cnj z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2171
proof (cases "z=0", simp)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2172
  assume "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
  2173
  then show ?thesis
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2174
     by (simp add: assms cnj_Ln 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
  2175
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2176
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2177
lemma has_field_derivative_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2178
  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
  2179
    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
  2180
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2181
  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
  2182
    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
  2183
  then have *: "inverse z = inverse (2*z) * 2"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2184
    by (simp add: divide_simps)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2185
  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
  2186
    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
  2187
  have "Im z = 0 \<Longrightarrow> 0 < Re z"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2188
    using assms complex_nonpos_Reals_iff not_less by blast
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2189
  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
  2190
    by (force intro: derivative_eq_intros * simp add: assms)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2191
  then show ?thesis
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2192
    apply (rule DERIV_transform_at[where d = "norm z"])
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2193
    apply (intro z derivative_eq_intros | simp add: assms)+
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2194
    using z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2195
    apply (metis csqrt_exp_Ln dist_0_norm less_irrefl)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2196
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2197
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2198
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2199
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
  2200
    "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
  2201
  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
  2202
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2203
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
  2204
    "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
  2205
  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
  2206
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2207
lemma continuous_at_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2208
    "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
  2209
  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
  2210
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  2211
corollary isCont_csqrt' [simp]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2212
   "\<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
  2213
  by (blast intro: isCont_o2 [OF _ continuous_at_csqrt])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  2214
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2215
lemma continuous_within_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2216
    "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
  2217
  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
  2218
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2219
lemma continuous_on_csqrt [continuous_intros]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2220
    "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on s csqrt"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2221
  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
  2222
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2223
lemma holomorphic_on_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2224
    "(\<And>z. z \<in> s \<Longrightarrow> z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> csqrt 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
  2225
  by (simp add: field_differentiable_within_csqrt 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
  2226
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2227
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
  2228
    "closed s \<Longrightarrow> a \<notin> s ==> continuous (at a within s) f"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2229
  using open_Compl
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2230
  by (force simp add: continuous_def eventually_at_topological filterlim_iff open_Collect_neg)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2231
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2232
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
  2233
    "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
  2234
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
  2235
  case True
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2236
  then have "Im z = 0" "Re z < 0 \<or> z = 0"
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
  2237
    using cnj.code complex_cnj_zero_iff  by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2238
  then show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2239
    apply (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2240
    apply (auto simp: continuous_within_eps_delta norm_conv_dist [symmetric])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2241
    apply (rule_tac x="e^2" in exI)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2242
    apply (auto simp: Reals_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2243
    by (metis linear not_less real_sqrt_less_iff real_sqrt_pow2_iff real_sqrt_power)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2244
next
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2245
  case False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2246
    then show ?thesis   by (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
  2247
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2248
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2249
subsection\<open>Complex arctangent\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2250
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2251
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
  2252
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2253
definition Arctan :: "complex \<Rightarrow> complex" where
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2254
    "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
  2255
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2256
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
  2257
  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
  2258
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2259
lemma Ln_conv_Arctan:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2260
  assumes "z \<noteq> -1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2261
  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
  2262
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2263
  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
  2264
             \<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
  2265
    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
  2266
  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
  2267
  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
  2268
  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
  2269
    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
  2270
  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
  2271
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2272
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2273
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
  2274
  by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2275
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2276
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
  2277
  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
  2278
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2279
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
  2280
  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
  2281
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2282
lemma tan_Arctan:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2283
  assumes "z\<^sup>2 \<noteq> -1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2284
    shows [simp]:"tan(Arctan z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2285
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2286
  have "1 + \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2287
    by (metis assms complex_i_mult_minus i_squared minus_unique power2_eq_square power2_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2288
  moreover
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2289
  have "1 - \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2290
    by (metis assms complex_i_mult_minus i_squared power2_eq_square power2_minus right_minus_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2291
  ultimately
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2292
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2293
    by (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus csqrt_exp_Ln [symmetric]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2294
                  divide_simps power2_eq_square [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2295
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2296
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2297
lemma Arctan_tan [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2298
  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
  2299
    shows "Arctan(tan z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2300
proof -
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2301
  have ge_pi2: "\<And>n::int. \<bar>of_int (2*n + 1) * pi/2\<bar> \<ge> pi/2"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2302
    by (case_tac n rule: int_cases) (auto simp: abs_mult)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2303
  have "exp (\<i>*z)*exp (\<i>*z) = -1 \<longleftrightarrow> exp (2*\<i>*z) = -1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2304
    by (metis distrib_right exp_add mult_2)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2305
  also have "... \<longleftrightarrow> exp (2*\<i>*z) = exp (\<i>*pi)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2306
    using cis_conv_exp cis_pi by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2307
  also have "... \<longleftrightarrow> exp (2*\<i>*z - \<i>*pi) = 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2308
    by (metis (no_types) diff_add_cancel diff_minus_eq_add exp_add exp_minus_inverse mult.commute)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2309
  also have "... \<longleftrightarrow> Re(\<i>*2*z - \<i>*pi) = 0 \<and> (\<exists>n::int. Im(\<i>*2*z - \<i>*pi) = of_int (2 * n) * pi)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2310
    by (simp add: exp_eq_1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2311
  also have "... \<longleftrightarrow> Im z = 0 \<and> (\<exists>n::int. 2 * Re z = of_int (2*n + 1) * pi)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2312
    by (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2313
  also have "... \<longleftrightarrow> False"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2314
    using assms ge_pi2
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2315
    apply (auto simp: algebra_simps)
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
  2316
    by (metis abs_mult_pos not_less of_nat_less_0_iff of_nat_numeral)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2317
  finally have *: "exp (\<i>*z)*exp (\<i>*z) + 1 \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2318
    by (auto simp: add.commute minus_unique)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2319
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2320
    using assms *
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2321
    apply (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps
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
  2322
                     i_times_eq_iff power2_eq_square [symmetric])
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2323
    apply (rule Ln_unique)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2324
    apply (auto simp: divide_simps exp_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2325
    apply (simp add: algebra_simps exp_double [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2326
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2327
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2328
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2329
lemma
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2330
  assumes "Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2331
  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
  2332
    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
  2333
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2334
  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
  2335
    using assms
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
  2336
    by (metis abs_one complex_i_mult_minus diff_0_right diff_minus_eq_add imaginary_unit.simps
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2337
              less_irrefl minus_diff_eq mult.right_neutral right_minus_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2338
  have "z \<noteq> -\<i>" using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2339
    by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2340
  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
  2341
    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
  2342
  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
  2343
    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
  2344
  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
  2345
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2346
    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
  2347
              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
  2348
  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
  2349
    using nz1 nz2 by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2350
  have "Im ((1 - \<i>*z) / (1 + \<i>*z)) = 0 \<Longrightarrow> 0 < Re ((1 - \<i>*z) / (1 + \<i>*z))"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2351
    apply (simp add: divide_complex_def)
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  2352
    apply (simp add: divide_simps split: if_split_asm)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2353
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2354
    apply (auto simp: algebra_simps abs_square_less_1 [unfolded power2_eq_square])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2355
    done
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2356
  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
  2357
    by (auto simp add: complex_nonpos_Reals_iff)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2358
  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
  2359
    unfolding Arctan_def divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2360
    using mpi_less_Im_Ln [OF nzi]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2361
    apply (auto simp: abs_if intro!: Im_Ln_less_pi * [unfolded divide_complex_def])
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2362
    done
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2363
  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
  2364
    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
  2365
    apply (rule DERIV_cong)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2366
    apply (intro derivative_eq_intros | simp add: nz0 *)+
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2367
    using nz0 nz1 zz
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2368
    apply (simp add: divide_simps power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2369
    apply (auto simp: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2370
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2371
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2372
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2373
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
  2374
  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
  2375
  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
  2376
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2377
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
  2378
    "(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
  2379
  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
  2380
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2381
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
  2382
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
  2383
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2384
lemma continuous_at_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2385
    "(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
  2386
  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
  2387
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2388
lemma continuous_within_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2389
    "(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
  2390
  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
  2391
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2392
lemma continuous_on_Arctan [continuous_intros]:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2393
    "(\<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
  2394
  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
  2395
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2396
lemma holomorphic_on_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2397
    "(\<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
  2398
  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
  2399
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2400
lemma Arctan_series:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2401
  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
  2402
  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
  2403
  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
  2404
  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
  2405
  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
  2406
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2407
  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
  2408
  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
  2409
  proof (cases "u = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2410
    assume u: "u \<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
  2411
    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
  2412
              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
  2413
    proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2414
      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
  2415
      have "ereal (norm (h u n) / norm (h u (Suc n))) =
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2416
             ereal (inverse (norm u)^2) * ereal ((of_nat (2*Suc n+1) / of_nat (Suc n)) /
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2417
                 (of_nat (2*Suc n-1) / of_nat (Suc 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
  2418
      by (simp add: h_def norm_mult norm_power norm_divide divide_simps
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2419
                    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
  2420
      also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2421
        by (auto simp: divide_simps simp del: of_nat_Suc) simp_all?
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2422
      also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc 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
  2423
        by (auto simp: divide_simps simp del: of_nat_Suc) simp_all?
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2424
      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
  2425
              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
  2426
    qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2427
    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
  2428
      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
  2429
    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
  2430
      by (intro lim_imp_Liminf) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2431
    moreover from power_strict_mono[OF that, of 2] u have "inverse (norm u)^2 > 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2432
      by (simp add: divide_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2433
    ultimately have A: "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2434
    from u have "summable (h u)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2435
      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
  2436
         (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
  2437
               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
  2438
    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
  2439
      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
  2440
         (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
  2441
  qed (simp add: h_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2442
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2443
  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
  2444
  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
  2445
    fix u :: complex assume "u \<in> ball 0 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2446
    hence u: "norm u < 1" by (simp add: dist_0_norm)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2447
    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
  2448
    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
  2449
    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
  2450
    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
  2451
      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
  2452
    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
  2453
      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
  2454
    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
  2455
      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
  2456
         (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
  2457
    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
  2458
    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
  2459
      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
  2460
    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
  2461
    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
  2462
      show "((\<lambda>u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2463
      by (simp_all add: dist_0_norm at_within_open[OF _ open_ball])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2464
  qed simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2465
  then obtain c where c: "\<And>u. norm u < 1 \<Longrightarrow> Arctan u - G u = c" by (auto simp: dist_0_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2466
  from this[of 0] have "c = 0" by (simp add: G_def g_def powser_zero)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2467
  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
  2468
  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
  2469
  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
  2470
                              (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
  2471
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2472
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2473
text \<open>A quickly-converging series for the logarithm, based on the arctangent.\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2474
lemma ln_series_quadratic:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2475
  assumes x: "x > (0::real)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2476
  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
  2477
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2478
  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
  2479
  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
  2480
  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
  2481
  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
  2482
    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
  2483
  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
  2484
  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
  2485
    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
  2486
  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
  2487
                 (\<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
  2488
    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
  2489
  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
  2490
    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
  2491
  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
  2492
    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
  2493
  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
  2494
    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
  2495
  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
  2496
    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
  2497
  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
  2498
  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
  2499
  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
  2500
  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
  2501
qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2502
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2503
subsection \<open>Real arctangent\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2504
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
  2505
lemma norm_exp_i_times [simp]: "norm (exp(\<i> * of_real y)) = 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2506
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2507
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2508
lemma norm_exp_imaginary: "norm(exp z) = 1 \<Longrightarrow> Re z = 0"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  2509
  by simp
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2510
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2511
lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2512
  unfolding Arctan_def divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2513
  apply (simp add: complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2514
  apply (rule norm_exp_imaginary)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2515
  apply (subst exp_Ln, auto)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2516
  apply (simp_all add: cmod_def complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2517
  apply (auto simp: divide_simps)
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
  2518
  apply (metis power_one sum_power2_eq_zero_iff zero_neq_one, algebra)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2519
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2520
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2521
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
  2522
proof (rule arctan_unique)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2523
  show "- (pi / 2) < Re (Arctan (complex_of_real x))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2524
    apply (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2525
    apply (rule Im_Ln_less_pi)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2526
    apply (auto simp: Im_complex_div_lemma complex_nonpos_Reals_iff)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2527
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2528
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2529
  have *: " (1 - \<i>*x) / (1 + \<i>*x) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2530
    by (simp add: divide_simps) ( simp add: complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2531
  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
  2532
    using mpi_less_Im_Ln [OF *]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2533
    by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2534
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2535
  have "tan (Re (Arctan (of_real x))) = Re (tan (Arctan (of_real x)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2536
    apply (auto simp: tan_def Complex.Re_divide Re_sin Re_cos Im_sin Im_cos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2537
    apply (simp add: field_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2538
    by (simp add: power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2539
  also have "... = x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2540
    apply (subst tan_Arctan, auto)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2541
    by (metis diff_0_right minus_diff_eq mult_zero_left not_le of_real_1 of_real_eq_iff of_real_minus of_real_power power2_eq_square real_minus_mult_self_le zero_less_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2542
  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
  2543
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2544
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2545
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
  2546
  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
  2547
  by (simp add: complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2548
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2549
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
  2550
  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
  2551
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2552
declare arctan_one [simp]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2553
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2554
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
  2555
  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
  2556
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2557
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
  2558
  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
  2559
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2560
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
  2561
  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
  2562
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2563
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
  2564
  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
  2565
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2566
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
  2567
  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
  2568
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2569
lemma arctan_add_raw:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2570
  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
  2571
    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
  2572
proof (rule arctan_unique [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2573
  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
  2574
    using assms by linarith+
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2575
  show "tan (arctan x + arctan y) = (x + y) / (1 - x * y)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2576
    using cos_gt_zero_pi [OF 12]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2577
    by (simp add: arctan tan_add)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2578
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2579
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2580
lemma arctan_inverse:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2581
  assumes "0 < x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2582
    shows "arctan(inverse x) = pi/2 - arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2583
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2584
  have "arctan(inverse x) = arctan(inverse(tan(arctan x)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2585
    by (simp add: arctan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2586
  also have "... = arctan (tan (pi / 2 - arctan x))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2587
    by (simp add: tan_cot)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2588
  also have "... = pi/2 - arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2589
  proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2590
    have "0 < pi - arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2591
    using arctan_ubound [of x] pi_gt_zero by linarith
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2592
    with assms show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2593
      by (simp add: Transcendental.arctan_tan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2594
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2595
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2596
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2597
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2598
lemma arctan_add_small:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2599
  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
  2600
    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
  2601
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
  2602
  case True then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2603
    by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2604
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2605
  case False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2606
  then have *: "\<bar>arctan x\<bar> < pi / 2 - \<bar>arctan y\<bar>" using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2607
    apply (auto simp add: abs_arctan arctan_inverse [symmetric] arctan_less_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2608
    apply (simp add: divide_simps abs_mult)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2609
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2610
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2611
    apply (rule arctan_add_raw)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2612
    using * by linarith
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2613
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2614
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2615
lemma abs_arctan_le:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2616
  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
  2617
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2618
  { fix w::complex and z::complex
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2619
    assume *: "w \<in> \<real>" "z \<in> \<real>"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2620
    have "cmod (Arctan w - Arctan z) \<le> 1 * cmod (w-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
  2621
      apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1])
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2622
      apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2623
      apply (force simp add: Reals_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2624
      apply (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2625
      using * by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2626
  }
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2627
  then have "cmod (Arctan (of_real x) - Arctan 0) \<le> 1 * cmod (of_real x -0)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2628
    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
  2629
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2630
    by (simp add: Arctan_of_real)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2631
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2632
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2633
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
  2634
  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
  2635
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2636
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
  2637
  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
  2638
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2639
lemma arctan_bounds:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2640
  assumes "0 \<le> x" "x < 1"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2641
  shows arctan_lower_bound:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2642
    "(\<Sum>k<2 * n. (- 1) ^ k * (1 / real (k * 2 + 1) * x ^ (k * 2 + 1))) \<le> arctan x"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2643
    (is "(\<Sum>k<_. (- 1)^ k * ?a k) \<le> _")
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2644
    and arctan_upper_bound:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2645
    "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
  2646
proof -
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2647
  have tendsto_zero: "?a \<longlonglongrightarrow> 0"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2648
    using assms
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2649
    apply -
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2650
    apply (rule tendsto_eq_rhs[where x="0 * 0"])
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2651
    subgoal by (intro tendsto_mult real_tendsto_divide_at_top)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2652
        (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2653
          intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2654
           tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2655
    subgoal by simp
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2656
    done
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2657
  have nonneg: "0 \<le> ?a n" for n
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2658
    by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2659
  have le: "?a (Suc n) \<le> ?a n" for n
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2660
    by (rule mult_mono[OF _ power_decreasing]) (auto simp: divide_simps assms less_imp_le)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2661
  from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2662
    summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2663
    assms
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2664
  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
  2665
    by (auto simp: arctan_series)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2666
qed
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2667
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2668
subsection \<open>Bounds on pi using real arctangent\<close>
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2669
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2670
lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2671
  using machin
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2672
  by simp
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2673
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2674
lemma pi_approx: "3.141592653588 \<le> pi" "pi \<le> 3.1415926535899"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2675
  unfolding pi_machin
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2676
  using arctan_bounds[of "1/5"   4]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2677
        arctan_bounds[of "1/239" 4]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2678
  by (simp_all add: eval_nat_numeral)
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
  2679
    
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2680
corollary pi_gt3: "pi > 3"
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2681
  using pi_approx by simp
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  2682
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2683
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2684
subsection\<open>Inverse Sine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2685
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2686
definition Arcsin :: "complex \<Rightarrow> complex" where
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2687
   "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
  2688
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2689
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
  2690
  using power2_csqrt [of "1 - z\<^sup>2"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2691
  apply auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2692
  by (metis complex_i_mult_minus diff_add_cancel diff_minus_eq_add diff_self mult.assoc mult.left_commute numeral_One power2_csqrt power2_eq_square zero_neq_numeral)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2693
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2694
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
  2695
  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
  2696
  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
  2697
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2698
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
  2699
  by (simp add: Arcsin_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2700
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2701
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
  2702
  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
  2703
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2704
lemma one_minus_z2_notin_nonpos_Reals:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2705
  assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2706
  shows "1 - z\<^sup>2 \<notin> \<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
  2707
    using assms
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2708
    apply (auto simp: complex_nonpos_Reals_iff Re_power2 Im_power2)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2709
    using power2_less_0 [of "Im z"] apply force
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2710
    using abs_square_less_1 not_le by blast
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2711
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2712
lemma isCont_Arcsin_lemma:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2713
  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
  2714
    shows False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2715
proof (cases "Im z = 0")
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2716
  case True
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2717
  then show ?thesis
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2718
    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
  2719
next
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2720
  case False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2721
  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
  2722
  proof (clarsimp simp add: cmod_def)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2723
    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
  2724
    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
  2725
      by simp
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2726
    then show False using False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2727
      by (simp add: power2_eq_square algebra_simps)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2728
  qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2729
  moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2730
    using le0
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2731
    apply simp
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2732
    apply (drule sqrt_le_D)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2733
    using cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1]
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2734
    apply (simp add: norm_power Re_power2 norm_minus_commute [of 1])
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2735
    done
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2736
  ultimately show False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2737
    by (simp add: Re_power2 Im_power2 cmod_power2)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2738
qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2739
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2740
lemma isCont_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2741
  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
  2742
    shows "isCont Arcsin z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2743
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2744
  have *: "\<i> * z + 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
  2745
    by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2746
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2747
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2748
    apply (simp add: Arcsin_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2749
    apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2750
    apply (simp add: one_minus_z2_notin_nonpos_Reals assms)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2751
    apply (rule *)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2752
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2753
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2754
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2755
lemma isCont_Arcsin' [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2756
  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
  2757
  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
  2758
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2759
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
  2760
proof -
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2761
  have "\<i>*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \<longleftrightarrow> (\<i>*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61806
diff changeset
  2762
    by (simp add: algebra_simps)  \<comment>\<open>Cancelling a factor of 2\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2763
  moreover have "... \<longleftrightarrow> (\<i>*z) + csqrt (1 - z\<^sup>2) = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2764
    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
  2765
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2766
    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
  2767
    apply (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2768
    apply (simp add: power2_eq_square [symmetric] algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2769
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2770
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2771
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2772
lemma Re_eq_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2773
    "\<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
  2774
      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
  2775
  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
  2776
  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
  2777
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2778
lemma Re_less_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2779
  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
  2780
    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
  2781
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2782
  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
  2783
    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
  2784
  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
  2785
    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
  2786
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2787
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2788
lemma Arcsin_sin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2789
    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
  2790
      shows "Arcsin(sin z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2791
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2792
  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
  2793
    by (simp add: sin_exp_eq Arcsin_def exp_minus power_divide)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2794
  also have "... = - (\<i> * Ln (csqrt (((exp (\<i>*z) + inverse (exp (\<i>*z)))/2)\<^sup>2) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2795
    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
  2796
  also have "... = - (\<i> * Ln (((exp (\<i>*z) + inverse (exp (\<i>*z)))/2) - (inverse (exp (\<i>*z)) - exp (\<i>*z)) / 2))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2797
    apply (subst csqrt_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2798
    using assms Re_eq_pihalf_lemma Re_less_pihalf_lemma
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2799
    apply auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2800
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2801
  also have "... =  - (\<i> * Ln (exp (\<i>*z)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2802
    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
  2803
  also have "... = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2804
    apply (subst Complex_Transcendental.Ln_exp)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2805
    using assms
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  2806
    apply (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
  2807
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2808
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2809
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2810
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2811
lemma Arcsin_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2812
    "\<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
  2813
  by (metis Arcsin_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2814
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2815
lemma Arcsin_0 [simp]: "Arcsin 0 = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2816
  by (metis Arcsin_sin norm_zero pi_half_gt_zero real_norm_def sin_zero zero_complex.simps(1))
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2817
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2818
lemma Arcsin_1 [simp]: "Arcsin 1 = pi/2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2819
  by (metis Arcsin_sin Im_complex_of_real Re_complex_of_real numeral_One of_real_numeral pi_half_ge_zero real_sqrt_abs real_sqrt_pow2 real_sqrt_power sin_of_real sin_pi_half)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2820
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2821
lemma Arcsin_minus_1 [simp]: "Arcsin(-1) = - (pi/2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2822
  by (metis Arcsin_1 Arcsin_sin Im_complex_of_real Re_complex_of_real abs_of_nonneg of_real_minus pi_half_ge_zero power2_minus real_sqrt_abs sin_Arcsin sin_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2823
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2824
lemma has_field_derivative_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2825
  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
  2826
    shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2827
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2828
  have "(sin (Arcsin z))\<^sup>2 \<noteq> 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2829
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2830
    apply atomize
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2831
    apply (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2832
    apply (metis abs_minus_cancel abs_one abs_power2 numeral_One numeral_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2833
    by (metis abs_minus_cancel abs_one abs_power2 one_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2834
  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
  2835
    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
  2836
  then show ?thesis
63492
a662e8139804 More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents: 63296
diff changeset
  2837
    apply (rule has_complex_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]])
a662e8139804 More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents: 63296
diff changeset
  2838
    apply (auto intro: isCont_Arcsin assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2839
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2840
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2841
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2842
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
  2843
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
  2844
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2845
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
  2846
    "(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
  2847
  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
  2848
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2849
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
  2850
    "(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
  2851
  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
  2852
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2853
lemma continuous_within_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2854
    "(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
  2855
  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
  2856
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2857
lemma continuous_on_Arcsin [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2858
    "(\<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
  2859
  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
  2860
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2861
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
  2862
  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
  2863
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2864
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2865
subsection\<open>Inverse Cosine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2866
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2867
definition Arccos :: "complex \<Rightarrow> complex" where
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2868
   "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
  2869
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2870
lemma Arccos_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Im(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
  2871
  using Arcsin_range_lemma [of "-z"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2872
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2873
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2874
lemma Arccos_body_lemma: "z + \<i> * 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
  2875
  using Arcsin_body_lemma [of z]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2876
  by (metis complex_i_mult_minus diff_add_cancel minus_diff_eq minus_unique mult.assoc mult.left_commute
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2877
           power2_csqrt power2_eq_square zero_neq_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2878
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2879
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
  2880
  by (simp add: Arccos_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2881
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2882
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
  2883
  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
  2884
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2885
text\<open>A very tricky argument to find!\<close>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2886
lemma isCont_Arccos_lemma:
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2887
  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
  2888
    shows False
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2889
proof (cases "Im z = 0")
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2890
  case True
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2891
  then show ?thesis
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2892
    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
  2893
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2894
  case False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2895
  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
  2896
    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
  2897
    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
  2898
  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
  2899
  proof (clarsimp simp add: cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2900
    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
  2901
    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
  2902
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2903
    then show False using False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2904
      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
  2905
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2906
  moreover have "(Im z)\<^sup>2 = ((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
  2907
    apply (subst Imz)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2908
    using abs_Re_le_cmod [of "1-z\<^sup>2"]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2909
    apply (simp add: Re_power2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2910
    done
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2911
  ultimately show False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2912
    by (simp add: cmod_power2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2913
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2914
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2915
lemma isCont_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2916
  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
  2917
    shows "isCont Arccos z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2918
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2919
  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
  2920
    by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2921
  with assms show ?thesis
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2922
    apply (simp add: Arccos_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2923
    apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2924
    apply (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
  2925
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2926
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2927
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2928
lemma isCont_Arccos' [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2929
  shows "isCont f z \<Longrightarrow> (Im (f z) = 0 \<Longrightarrow> \<bar>Re (f z)\<bar> < 1) \<Longrightarrow> isCont (\<lambda>x. Arccos (f x)) z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2930
  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
  2931
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2932
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
  2933
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2934
  have "z*2 + \<i> * (2 * csqrt (1 - z\<^sup>2)) = 0 \<longleftrightarrow> z*2 + \<i> * csqrt (1 - z\<^sup>2)*2 = 0"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61806
diff changeset
  2935
    by (simp add: algebra_simps)  \<comment>\<open>Cancelling a factor of 2\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2936
  moreover have "... \<longleftrightarrow> z + \<i> * csqrt (1 - z\<^sup>2) = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2937
    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
  2938
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2939
    apply (simp add: cos_exp_eq Arccos_def Arccos_body_lemma exp_minus field_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2940
    apply (simp add: power2_eq_square [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2941
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2942
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2943
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2944
lemma Arccos_cos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2945
    assumes "0 < Re z & Re z < pi \<or>
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2946
             Re z = 0 & 0 \<le> Im z \<or>
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2947
             Re z = pi & Im z \<le> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2948
      shows "Arccos(cos z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2949
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
  2950
  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
  2951
    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
  2952
  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
  2953
    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
  2954
  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
  2955
                           \<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
  2956
    by (simp add: cos_exp_eq Arccos_def exp_minus power_divide)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2957
  also have "... = - (\<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
  2958
                              \<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
  2959
    apply (subst csqrt_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2960
    using assms Re_sin_pos [of z] Im_sin_nonneg [of z] Im_sin_nonneg2 [of z]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2961
    apply (auto simp: * Re_sin Im_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2962
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2963
  also have "... =  - (\<i> * Ln (exp (\<i>*z)))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2964
    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
  2965
  also have "... = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2966
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2967
    apply (subst Complex_Transcendental.Ln_exp, auto)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2968
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2969
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2970
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2971
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2972
lemma Arccos_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2973
    "\<lbrakk>cos z = w;
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2974
      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
  2975
      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
  2976
      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
  2977
  using Arccos_cos by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2978
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2979
lemma Arccos_0 [simp]: "Arccos 0 = pi/2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2980
  by (rule Arccos_unique) (auto simp: of_real_numeral)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2981
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2982
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
  2983
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2984
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2985
lemma Arccos_minus1: "Arccos(-1) = pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2986
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2987
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2988
lemma has_field_derivative_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2989
  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
  2990
    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
  2991
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2992
  have "(cos (Arccos z))\<^sup>2 \<noteq> 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2993
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2994
    apply atomize
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2995
    apply (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2996
    apply (metis abs_minus_cancel abs_one abs_power2 numeral_One numeral_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2997
    apply (metis left_minus less_irrefl power_one sum_power2_gt_zero_iff zero_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2998
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2999
  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
  3000
    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
  3001
  then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)"
63492
a662e8139804 More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents: 63296
diff changeset
  3002
    apply (rule has_complex_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]])
a662e8139804 More advanced theorems about retracts, homotopies., etc
paulson <lp15@cam.ac.uk>
parents: 63296
diff changeset
  3003
    apply (auto intro: isCont_Arccos assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3004
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3005
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3006
    by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3007
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3008
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3009
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
  3010
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
  3011
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3012
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
  3013
    "(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
  3014
  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
  3015
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3016
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
  3017
    "(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
  3018
  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
  3019
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3020
lemma continuous_within_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3021
    "(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
  3022
  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
  3023
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3024
lemma continuous_on_Arccos [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3025
    "(\<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
  3026
  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
  3027
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3028
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
  3029
  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
  3030
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3031
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3032
subsection\<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
  3033
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3034
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
  3035
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3036
  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
  3037
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3038
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
  3039
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3040
  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
  3041
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3042
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
  3043
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3044
  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
  3045
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3046
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
  3047
  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
  3048
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3049
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
  3050
proof -
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3051
  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
  3052
    using norm_cos_squared [of "Arccos w"] real_le_abs_sinh [of "Im (Arccos w)"]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3053
    apply (simp only: abs_le_square_iff)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3054
    apply (simp add: divide_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3055
    done
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3056
  also have "... \<le> (cmod w)\<^sup>2"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3057
    by (auto simp: cmod_power2)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3058
  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
  3059
    using abs_le_square_iff by force
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3060
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
  3061
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3062
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
  3063
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3064
  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
  3065
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3066
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
  3067
  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
  3068
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3069
lemma norm_Arccos_bounded:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3070
  fixes w :: complex
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3071
  shows "norm (Arccos w) \<le> pi + norm w"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3072
proof -
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3073
  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
  3074
    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
  3075
  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
  3076
    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
  3077
  then have "cmod (Arccos w) \<le> pi + cmod (cos (Arccos w))"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3078
    apply (simp add: norm_le_square)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3079
    by (metis dot_square_norm norm_ge_zero norm_le_square pi_ge_zero triangle_lemma)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3080
  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
  3081
    by auto
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3082
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3083
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3084
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3085
subsection\<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
  3086
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3087
lemma cos_Arcsin_nonzero:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3088
  assumes "z\<^sup>2 \<noteq> 1" shows "cos(Arcsin z) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3089
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3090
  have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = z\<^sup>2 * (z\<^sup>2 - 1)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3091
    by (simp add: power_mult_distrib algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3092
  have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> z\<^sup>2 - 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3093
  proof
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3094
    assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = z\<^sup>2 - 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3095
    then have "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (z\<^sup>2 - 1)\<^sup>2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3096
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3097
    then have "z\<^sup>2 * (z\<^sup>2 - 1) = (z\<^sup>2 - 1)*(z\<^sup>2 - 1)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3098
      using eq power2_eq_square by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3099
    then show False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3100
      using assms by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3101
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3102
  then have "1 + \<i> * z * (csqrt (1 - z * z)) \<noteq> z\<^sup>2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3103
    by (metis add_minus_cancel power2_eq_square uminus_add_conv_diff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3104
  then have "2*(1 + \<i> * z * (csqrt (1 - z * z))) \<noteq> 2*z\<^sup>2"  (*FIXME cancel_numeral_factor*)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3105
    by (metis mult_cancel_left zero_neq_numeral)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3106
  then have "(\<i> * z + csqrt (1 - z\<^sup>2))\<^sup>2 \<noteq> -1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3107
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3108
    apply (auto simp: power2_sum)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3109
    apply (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
  3110
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3111
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3112
    apply (simp add: cos_exp_eq Arcsin_def exp_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3113
    apply (simp add: divide_simps Arcsin_body_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3114
    apply (metis add.commute minus_unique power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3115
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3116
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3117
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3118
lemma sin_Arccos_nonzero:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3119
  assumes "z\<^sup>2 \<noteq> 1" shows "sin(Arccos z) \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3120
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3121
  have eq: "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = -(z\<^sup>2) * (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3122
    by (simp add: power_mult_distrib algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3123
  have "\<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> 1 - z\<^sup>2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3124
  proof
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3125
    assume "\<i> * z * (csqrt (1 - z\<^sup>2)) = 1 - z\<^sup>2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3126
    then have "(\<i> * z * (csqrt (1 - z\<^sup>2)))\<^sup>2 = (1 - z\<^sup>2)\<^sup>2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3127
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3128
    then have "-(z\<^sup>2) * (1 - z\<^sup>2) = (1 - z\<^sup>2)*(1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3129
      using eq power2_eq_square by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3130
    then have "-(z\<^sup>2) = (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3131
      using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3132
      by (metis add.commute add.right_neutral diff_add_cancel mult_right_cancel)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3133
    then show False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3134
      using assms by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3135
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3136
  then have "z\<^sup>2 + \<i> * z * (csqrt (1 - z\<^sup>2)) \<noteq> 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3137
    by (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3138
  then have "2*(z\<^sup>2 + \<i> * z * (csqrt (1 - z\<^sup>2))) \<noteq> 2*1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3139
    by (metis mult_cancel_left2 zero_neq_numeral)  (*FIXME cancel_numeral_factor*)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3140
  then have "(z + \<i> * csqrt (1 - z\<^sup>2))\<^sup>2 \<noteq> 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3141
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3142
    apply (auto simp: Power.comm_semiring_1_class.power2_sum power_mult_distrib)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3143
    apply (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
  3144
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3145
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3146
    apply (simp add: sin_exp_eq Arccos_def exp_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3147
    apply (simp add: divide_simps Arccos_body_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3148
    apply (simp add: power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3149
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3150
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3151
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3152
lemma cos_sin_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3153
  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
  3154
    shows "cos z = csqrt(1 - (sin z)\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3155
  apply (rule csqrt_unique [THEN sym])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3156
  apply (simp add: cos_squared_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3157
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3158
  apply (auto simp: Re_cos Im_cos add_pos_pos mult_le_0_iff zero_le_mult_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3159
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3160
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3161
lemma sin_cos_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3162
  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
  3163
    shows "sin z = csqrt(1 - (cos z)\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3164
  apply (rule csqrt_unique [THEN sym])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3165
  apply (simp add: sin_squared_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3166
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3167
  apply (auto simp: Re_sin Im_sin add_pos_pos mult_le_0_iff zero_le_mult_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3168
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3169
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3170
lemma Arcsin_Arccos_csqrt_pos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3171
    "(0 < Re z | Re z = 0 & 0 \<le> Im z) \<Longrightarrow> Arcsin z = Arccos(csqrt(1 - z\<^sup>2))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3172
  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
  3173
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3174
lemma Arccos_Arcsin_csqrt_pos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3175
    "(0 < Re z | Re z = 0 & 0 \<le> Im z) \<Longrightarrow> Arccos z = Arcsin(csqrt(1 - z\<^sup>2))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3176
  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
  3177
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3178
lemma sin_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3179
    "0 < Re z | Re z = 0 & 0 \<le> Im z \<Longrightarrow> sin(Arccos z) = csqrt(1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3180
  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
  3181
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3182
lemma cos_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3183
    "0 < Re z | Re z = 0 & 0 \<le> Im z \<Longrightarrow> cos(Arcsin z) = csqrt(1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3184
  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
  3185
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3186
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3187
subsection\<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
  3188
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3189
lemma Im_Arcsin_of_real:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3190
  assumes "\<bar>x\<bar> \<le> 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3191
    shows "Im (Arcsin (of_real x)) = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3192
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3193
  have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3194
    by (simp add: of_real_sqrt del: csqrt_of_real_nonneg)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3195
  then have "cmod (\<i> * of_real x + csqrt (1 - (of_real x)\<^sup>2))^2 = 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3196
    using assms abs_square_le_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3197
    by (force simp add: Complex.cmod_power2)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3198
  then have "cmod (\<i> * of_real x + csqrt (1 - (of_real x)\<^sup>2)) = 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3199
    by (simp add: norm_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3200
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3201
    by (simp add: Im_Arcsin exp_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3202
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3203
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3204
corollary Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3205
  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
  3206
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3207
lemma arcsin_eq_Re_Arcsin:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3208
  assumes "\<bar>x\<bar> \<le> 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3209
    shows "arcsin x = Re (Arcsin (of_real x))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3210
unfolding arcsin_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3211
proof (rule the_equality, safe)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3212
  show "- (pi / 2) \<le> Re (Arcsin (complex_of_real x))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3213
    using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3214
    by (auto simp: Complex.in_Reals_norm Re_Arcsin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3215
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3216
  show "Re (Arcsin (complex_of_real x)) \<le> pi / 2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3217
    using Re_Ln_pos_le [OF Arcsin_body_lemma, of "of_real x"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3218
    by (auto simp: Complex.in_Reals_norm Re_Arcsin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3219
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3220
  show "sin (Re (Arcsin (complex_of_real x))) = x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3221
    using Re_sin [of "Arcsin (of_real x)"] Arcsin_body_lemma [of "of_real x"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3222
    by (simp add: Im_Arcsin_of_real assms)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3223
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3224
  fix x'
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3225
  assume "- (pi / 2) \<le> x'" "x' \<le> pi / 2" "x = sin x'"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3226
  then show "x' = Re (Arcsin (complex_of_real (sin x')))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3227
    apply (simp add: sin_of_real [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3228
    apply (subst Arcsin_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3229
    apply (auto simp: )
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3230
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3231
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3232
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3233
lemma of_real_arcsin: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arcsin x) = Arcsin(of_real x)"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3234
  by (metis Im_Arcsin_of_real add.right_neutral arcsin_eq_Re_Arcsin complex_eq mult_zero_right of_real_0)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3235
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3236
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3237
subsection\<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
  3238
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3239
lemma Im_Arccos_of_real:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3240
  assumes "\<bar>x\<bar> \<le> 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3241
    shows "Im (Arccos (of_real x)) = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3242
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3243
  have "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 1))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3244
    by (simp add: of_real_sqrt del: csqrt_of_real_nonneg)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3245
  then have "cmod (of_real x + \<i> * csqrt (1 - (of_real x)\<^sup>2))^2 = 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3246
    using assms abs_square_le_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3247
    by (force simp add: Complex.cmod_power2)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3248
  then have "cmod (of_real x + \<i> * csqrt (1 - (of_real x)\<^sup>2)) = 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3249
    by (simp add: norm_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3250
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3251
    by (simp add: Im_Arccos exp_minus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3252
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3253
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3254
corollary Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3255
  by (metis Im_Arccos_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
  3256
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3257
lemma arccos_eq_Re_Arccos:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3258
  assumes "\<bar>x\<bar> \<le> 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3259
    shows "arccos x = Re (Arccos (of_real x))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3260
unfolding arccos_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3261
proof (rule the_equality, safe)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3262
  show "0 \<le> Re (Arccos (complex_of_real x))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3263
    using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3264
    by (auto simp: Complex.in_Reals_norm Re_Arccos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3265
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3266
  show "Re (Arccos (complex_of_real x)) \<le> pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3267
    using Im_Ln_pos_le [OF Arccos_body_lemma, of "of_real x"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3268
    by (auto simp: Complex.in_Reals_norm Re_Arccos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3269
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3270
  show "cos (Re (Arccos (complex_of_real x))) = x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3271
    using Re_cos [of "Arccos (of_real x)"] Arccos_body_lemma [of "of_real x"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3272
    by (simp add: Im_Arccos_of_real assms)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3273
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3274
  fix x'
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3275
  assume "0 \<le> x'" "x' \<le> pi" "x = cos x'"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3276
  then show "x' = Re (Arccos (complex_of_real (cos x')))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3277
    apply (simp add: cos_of_real [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3278
    apply (subst Arccos_cos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3279
    apply (auto simp: )
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3280
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3281
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3282
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3283
lemma of_real_arccos: "\<bar>x\<bar> \<le> 1 \<Longrightarrow> of_real(arccos x) = Arccos(of_real x)"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3284
  by (metis Im_Arccos_of_real add.right_neutral arccos_eq_Re_Arccos complex_eq mult_zero_right of_real_0)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  3285
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3286
subsection\<open>Some interrelationships among the real inverse trig functions.\<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
  3287
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
  3288
lemma arccos_arctan:
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
  3289
  assumes "-1 < x" "x < 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
  3290
    shows "arccos x = pi/2 - arctan(x / sqrt(1 - x\<^sup>2))"
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
  3291
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
  3292
  have "arctan(x / sqrt(1 - x\<^sup>2)) - (pi/2 - arccos x) = 0"
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
  3293
  proof (rule sin_eq_0_pi)
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
  3294
    show "- pi < arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos 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
  3295
      using arctan_lbound [of "x / sqrt(1 - x\<^sup>2)"]  arccos_bounded [of x] assms
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
  3296
      by (simp add: algebra_simps)
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
  3297
  next
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
  3298
    show "arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x) < pi"
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
  3299
      using arctan_ubound [of "x / sqrt(1 - x\<^sup>2)"]  arccos_bounded [of x] assms
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
  3300
      by (simp add: algebra_simps)
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
  3301
  next
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
  3302
    show "sin (arctan (x / sqrt (1 - x\<^sup>2)) - (pi / 2 - arccos x)) = 0"
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
  3303
      using assms
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
  3304
      by (simp add: algebra_simps sin_diff cos_add sin_arccos sin_arctan cos_arctan
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
  3305
                    power2_eq_square square_eq_1_iff)
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
  3306
  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
  3307
  then 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
  3308
    by 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
  3309
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
  3310
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
  3311
lemma arcsin_plus_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
  3312
  assumes "-1 \<le> x" "x \<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
  3313
    shows "arcsin x + arccos x = pi/2"
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
  3314
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
  3315
  have "arcsin x = pi/2 - arccos 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
  3316
    apply (rule sin_inj_pi)
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
  3317
    using assms arcsin [OF assms] arccos [OF assms]
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
  3318
    apply (auto simp: algebra_simps sin_diff)
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
  3319
    done
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
  3320
  then 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
  3321
    by (simp add: algebra_simps)
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
  3322
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
  3323
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
  3324
lemma arcsin_arccos_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = pi/2 - arccos 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
  3325
  using arcsin_plus_arccos by force
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
  3326
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
  3327
lemma arccos_arcsin_eq: "-1 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = pi/2 - arcsin 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
  3328
  using arcsin_plus_arccos by force
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
  3329
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
  3330
lemma arcsin_arctan: "-1 < x \<Longrightarrow> x < 1 \<Longrightarrow> arcsin x = arctan(x / sqrt(1 - x\<^sup>2))"
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
  3331
  by (simp add: arccos_arctan arcsin_arccos_eq)
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
  3332
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  3333
lemma csqrt_1_diff_eq: "csqrt (1 - (of_real x)\<^sup>2) = (if x^2 \<le> 1 then sqrt (1 - x^2) else \<i> * sqrt (x^2 - 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
  3334
  by ( simp add: of_real_sqrt del: csqrt_of_real_nonneg)
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
  3335
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
  3336
lemma arcsin_arccos_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arcsin x = arccos(sqrt(1 - x\<^sup>2))"
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
  3337
  apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_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
  3338
  apply (subst Arcsin_Arccos_csqrt_pos)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  3339
  apply (auto simp: power_le_one csqrt_1_diff_eq)
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
  3340
  done
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
  3341
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
  3342
lemma arcsin_arccos_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arcsin x = -arccos(sqrt(1 - x\<^sup>2))"
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
  3343
  using arcsin_arccos_sqrt_pos [of "-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
  3344
  by (simp add: arcsin_minus)
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
  3345
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
  3346
lemma arccos_arcsin_sqrt_pos: "0 \<le> x \<Longrightarrow> x \<le> 1 \<Longrightarrow> arccos x = arcsin(sqrt(1 - x\<^sup>2))"
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
  3347
  apply (simp add: abs_square_le_1 arcsin_eq_Re_Arcsin arccos_eq_Re_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
  3348
  apply (subst Arccos_Arcsin_csqrt_pos)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  3349
  apply (auto simp: power_le_one csqrt_1_diff_eq)
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
  3350
  done
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
  3351
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
  3352
lemma arccos_arcsin_sqrt_neg: "-1 \<le> x \<Longrightarrow> x \<le> 0 \<Longrightarrow> arccos x = pi - arcsin(sqrt(1 - x\<^sup>2))"
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
  3353
  using arccos_arcsin_sqrt_pos [of "-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
  3354
  by (simp add: arccos_minus)
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
  3355
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3356
subsection\<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
  3357
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
  3358
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
  3359
    "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
  3360
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
  3361
  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
  3362
        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
  3363
    by (rule continuous_on_cong [OF refl]) (simp add: arcsin_eq_Re_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
  3364
  also have "... = ?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
  3365
    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
  3366
  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
  3367
    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
  3368
          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
  3369
    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
  3370
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
  3371
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
  3372
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
  3373
    "continuous (at z within {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
  3374
proof (cases "z \<in> {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
  3375
  case True then 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
  3376
    using continuous_on_Arcsin_real continuous_on_eq_continuous_within
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
  3377
    by blast
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
  3378
next
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
  3379
  case False
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
  3380
  with closed_real_abs_le [of 1] 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
  3381
    by (rule continuous_within_closed_nontrivial)
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
  3382
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
  3383
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
  3384
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
  3385
    "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
  3386
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
  3387
  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
  3388
        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
  3389
    by (rule continuous_on_cong [OF refl]) (simp add: arccos_eq_Re_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
  3390
  also have "... = ?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
  3391
    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
  3392
  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
  3393
    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
  3394
          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
  3395
    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
  3396
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
  3397
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
  3398
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
  3399
    "continuous (at z within {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
  3400
proof (cases "z \<in> {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
  3401
  case True then 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
  3402
    using continuous_on_Arccos_real continuous_on_eq_continuous_within
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
  3403
    by blast
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
  3404
next
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
  3405
  case False
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
  3406
  with closed_real_abs_le [of 1] 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
  3407
    by (rule continuous_within_closed_nontrivial)
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
  3408
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
  3409
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
  3410
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3411
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
  3412
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
  3413
lemma complex_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
  3414
  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
  3415
  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
  3416
    shows "exp(2 * of_real pi * \<i> * of_nat j / of_nat n)^n = 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
  3417
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
  3418
  have *: "of_nat j * (complex_of_real pi * 2) = complex_of_real (2 * real j * pi)"
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
  3419
    by (simp add: of_real_numeral)
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
  3420
  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
  3421
    apply (simp add: 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
  3422
    apply (simp only: * cos_of_real sin_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
  3423
    apply (simp add: )
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
  3424
    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
  3425
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
  3426
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
  3427
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
  3428
  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
  3429
  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
  3430
    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
  3431
           \<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
  3432
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
  3433
    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
  3434
               \<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
  3435
          (\<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
  3436
              (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
  3437
      by (simp add: algebra_simps)
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
  3438
    also have "... \<longleftrightarrow> (\<exists>z::int. of_nat j = of_nat k + of_nat n * (of_int z :: 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
  3439
      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
  3440
    also have "... \<longleftrightarrow> (\<exists>z::int. of_nat j = of_nat k + of_nat n * 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
  3441
      apply (rule HOL.iff_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
  3442
      apply (auto 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
  3443
      using of_int_eq_iff apply fastforce
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
  3444
      by (metis of_int_add of_int_mult of_int_of_nat_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
  3445
    also have "... \<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
  3446
      by (auto simp: mod_eq_dvd_iff dvd_def algebra_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
  3447
    also have "... \<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
  3448
      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
  3449
    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
  3450
             \<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
  3451
   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
  3452
  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
  3453
    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
  3454
    by (simp add: exp_eq divide_simps mult_ac of_real_numeral *)
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
  3455
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
  3456
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
  3457
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
  3458
    "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
  3459
              {..<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
  3460
  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
  3461
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
  3462
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
  3463
  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
  3464
  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
  3465
    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
  3466
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
  3467
  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
  3468
    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
  3469
  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
  3470
     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
  3471
     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
  3472
  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
  3473
    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
  3474
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
  3475
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
  3476
lemma finite_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
  3477
     "finite {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j::nat. 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
  3478
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
  3479
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
  3480
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
  3481
     "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
  3482
  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
  3483
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
  3484
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
  3485
  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
  3486
    shows "{z::complex. z^n = 1} = {exp(2 * of_real pi * \<i> * of_nat j / of_nat n) | j::nat. 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
  3487
  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
  3488
  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
  3489
  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
  3490
  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
  3491
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
  3492
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
  3493
  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
  3494
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
  3495
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
  3496
    "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
  3497
  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
  3498
  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
  3499
  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
  3500
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3501
subsection\<open> Formulation of loop homotopy in terms of maps out of type complex\<close>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3502
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3503
lemma homotopic_circlemaps_imp_homotopic_loops:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3504
  assumes "homotopic_with (\<lambda>h. True) (sphere 0 1) S f g"
64508
874555896035 more symbols;
wenzelm
parents: 64394
diff changeset
  3505
   shows "homotopic_loops S (f \<circ> exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>))
874555896035 more symbols;
wenzelm
parents: 64394
diff changeset
  3506
                            (g \<circ> exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>))"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3507
proof -
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3508
  have "homotopic_with (\<lambda>f. True) {z. cmod z = 1} S f g"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3509
    using assms by (auto simp: sphere_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3510
  moreover have "continuous_on {0..1} (exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3511
     by (intro continuous_intros)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3512
  moreover have "(exp \<circ> (\<lambda>t. 2 * of_real pi * of_real t * \<i>)) ` {0..1} \<subseteq> {z. cmod z = 1}"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3513
    by (auto simp: norm_mult)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3514
  ultimately
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3515
  show ?thesis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3516
    apply (simp add: homotopic_loops_def comp_assoc)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3517
    apply (rule homotopic_with_compose_continuous_right)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3518
      apply (auto simp: pathstart_def pathfinish_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3519
    done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3520
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3521
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3522
lemma homotopic_loops_imp_homotopic_circlemaps:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3523
  assumes "homotopic_loops S p q"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3524
    shows "homotopic_with (\<lambda>h. True) (sphere 0 1) S
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3525
                          (p \<circ> (\<lambda>z. (Arg z / (2 * pi))))
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3526
                          (q \<circ> (\<lambda>z. (Arg z / (2 * pi))))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3527
proof -
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3528
  obtain h where conth: "continuous_on ({0..1::real} \<times> {0..1}) h"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3529
             and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3530
             and h0: "(\<forall>x. h (0, x) = p x)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3531
             and h1: "(\<forall>x. h (1, x) = q x)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3532
             and h01: "(\<forall>t\<in>{0..1}. h (t, 1) = h (t, 0)) "
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3533
    using assms
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3534
    by (auto simp: homotopic_loops_def sphere_def homotopic_with_def pathstart_def pathfinish_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3535
  define j where "j \<equiv> \<lambda>z. if 0 \<le> Im (snd z)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3536
                          then h (fst z, Arg (snd z) / (2 * pi))
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3537
                          else h (fst z, 1 - Arg (cnj (snd z)) / (2 * pi))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3538
  have Arg_eq: "1 - Arg (cnj y) / (2 * pi) = Arg y / (2 * pi) \<or> Arg y = 0 \<and> Arg (cnj y) = 0" if "cmod y = 1" for y
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3539
    using that Arg_eq_0_pi Arg_eq_pi by (force simp: Arg_cnj divide_simps)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3540
  show ?thesis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3541
  proof (simp add: homotopic_with; intro conjI ballI exI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3542
    show "continuous_on ({0..1} \<times> sphere 0 1) (\<lambda>w. h (fst w, Arg (snd w) / (2 * pi)))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3543
    proof (rule continuous_on_eq)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3544
      show j: "j x = h (fst x, Arg (snd x) / (2 * pi))" if "x \<in> {0..1} \<times> sphere 0 1" for x
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3545
        using Arg_eq that h01 by (force simp: j_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3546
      have eq:  "S = S \<inter> (UNIV \<times> {z. 0 \<le> Im z}) \<union> S \<inter> (UNIV \<times> {z. Im z \<le> 0})" for S :: "(real*complex)set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3547
        by auto
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3548
      have c1: "continuous_on ({0..1} \<times> sphere 0 1 \<inter> UNIV \<times> {z. 0 \<le> Im z}) (\<lambda>x. h (fst x, Arg (snd x) / (2 * pi)))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3549
        apply (intro continuous_intros continuous_on_compose2 [OF conth]  continuous_on_compose2 [OF continuous_on_upperhalf_Arg])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3550
            apply (auto simp: Arg)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3551
        apply (meson Arg_lt_2pi linear not_le)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3552
        done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3553
      have c2: "continuous_on ({0..1} \<times> sphere 0 1 \<inter> UNIV \<times> {z. Im z \<le> 0}) (\<lambda>x. h (fst x, 1 - Arg (cnj (snd x)) / (2 * pi)))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3554
        apply (intro continuous_intros continuous_on_compose2 [OF conth]  continuous_on_compose2 [OF continuous_on_upperhalf_Arg])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3555
            apply (auto simp: Arg)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3556
        apply (meson Arg_lt_2pi linear not_le)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3557
        done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3558
      show "continuous_on ({0..1} \<times> sphere 0 1) j"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3559
        apply (simp add: j_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3560
        apply (subst eq)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3561
        apply (rule continuous_on_cases_local)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3562
            apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3563
        using Arg_eq h01
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3564
        by force
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3565
    qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3566
    have "(\<lambda>w. h (fst w, Arg (snd w) / (2 * pi))) ` ({0..1} \<times> sphere 0 1) \<subseteq> h ` ({0..1} \<times> {0..1})"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3567
      by (auto simp: Arg_ge_0 Arg_lt_2pi less_imp_le)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3568
    also have "... \<subseteq> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3569
      using him by blast
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3570
    finally show "(\<lambda>w. h (fst w, Arg (snd w) / (2 * pi))) ` ({0..1} \<times> sphere 0 1) \<subseteq> S" .
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3571
  qed (auto simp: h0 h1)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3572
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3573
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3574
lemma simply_connected_homotopic_loops:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3575
  "simply_connected S \<longleftrightarrow>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3576
       (\<forall>p q. homotopic_loops S p p \<and> homotopic_loops S q q \<longrightarrow> homotopic_loops S p q)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3577
unfolding simply_connected_def using homotopic_loops_refl by metis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3578
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3579
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3580
lemma simply_connected_eq_homotopic_circlemaps1:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3581
  fixes f :: "complex \<Rightarrow> 'a::topological_space" and g :: "complex \<Rightarrow> 'a"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3582
  assumes S: "simply_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3583
      and contf: "continuous_on (sphere 0 1) f" and fim: "f ` (sphere 0 1) \<subseteq> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3584
      and contg: "continuous_on (sphere 0 1) g" and gim: "g ` (sphere 0 1) \<subseteq> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3585
    shows "homotopic_with (\<lambda>h. True) (sphere 0 1) S f g"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3586
proof -
64508
874555896035 more symbols;
wenzelm
parents: 64394
diff changeset
  3587
  have "homotopic_loops S (f \<circ> exp \<circ> (\<lambda>t. of_real(2 * pi * t) * \<i>)) (g \<circ> exp \<circ> (\<lambda>t. of_real(2 * pi *  t) * \<i>))"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3588
    apply (rule S [unfolded simply_connected_homotopic_loops, rule_format])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3589
    apply (simp add: homotopic_circlemaps_imp_homotopic_loops homotopic_with_refl contf fim contg gim)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3590
    done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3591
  then show ?thesis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3592
    apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3593
      apply (auto simp: o_def complex_norm_eq_1_exp mult.commute)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3594
    done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3595
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3596
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3597
lemma simply_connected_eq_homotopic_circlemaps2a:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3598
  fixes h :: "complex \<Rightarrow> 'a::topological_space"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3599
  assumes conth: "continuous_on (sphere 0 1) h" and him: "h ` (sphere 0 1) \<subseteq> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3600
      and hom: "\<And>f g::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3601
                \<lbrakk>continuous_on (sphere 0 1) f; f ` (sphere 0 1) \<subseteq> S;
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3602
                continuous_on (sphere 0 1) g; g ` (sphere 0 1) \<subseteq> S\<rbrakk>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3603
                \<Longrightarrow> homotopic_with (\<lambda>h. True) (sphere 0 1) S f g"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3604
            shows "\<exists>a. homotopic_with (\<lambda>h. True) (sphere 0 1) S h (\<lambda>x. a)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3605
    apply (rule_tac x="h 1" in exI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3606
    apply (rule hom)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3607
    using assms
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3608
    by (auto simp: continuous_on_const)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3609
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3610
lemma simply_connected_eq_homotopic_circlemaps2b:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3611
  fixes S :: "'a::real_normed_vector set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3612
  assumes "\<And>f g::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3613
                \<lbrakk>continuous_on (sphere 0 1) f; f ` (sphere 0 1) \<subseteq> S;
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3614
                continuous_on (sphere 0 1) g; g ` (sphere 0 1) \<subseteq> S\<rbrakk>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3615
                \<Longrightarrow> homotopic_with (\<lambda>h. True) (sphere 0 1) S f g"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3616
  shows "path_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3617
proof (clarsimp simp add: path_connected_eq_homotopic_points)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3618
  fix a b
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3619
  assume "a \<in> S" "b \<in> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3620
  then show "homotopic_loops S (linepath a a) (linepath b b)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3621
    using homotopic_circlemaps_imp_homotopic_loops [OF assms [of "\<lambda>x. a" "\<lambda>x. b"]]
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3622
    by (auto simp: o_def continuous_on_const linepath_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3623
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3624
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3625
lemma simply_connected_eq_homotopic_circlemaps3:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3626
  fixes h :: "complex \<Rightarrow> 'a::real_normed_vector"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3627
  assumes "path_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3628
      and hom: "\<And>f::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3629
                  \<lbrakk>continuous_on (sphere 0 1) f; f `(sphere 0 1) \<subseteq> S\<rbrakk>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3630
                  \<Longrightarrow> \<exists>a. homotopic_with (\<lambda>h. True) (sphere 0 1) S f (\<lambda>x. a)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3631
    shows "simply_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3632
proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3633
  fix p
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3634
  assume p: "path p" "path_image p \<subseteq> S" "pathfinish p = pathstart p"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3635
  then have "homotopic_loops S p p"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3636
    by (simp add: homotopic_loops_refl)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3637
  then obtain a where homp: "homotopic_with (\<lambda>h. True) (sphere 0 1) S (p \<circ> (\<lambda>z. Arg z / (2 * pi))) (\<lambda>x. a)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3638
    by (metis homotopic_with_imp_subset2 homotopic_loops_imp_homotopic_circlemaps homotopic_with_imp_continuous hom)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3639
  show "\<exists>a. a \<in> S \<and> homotopic_loops S p (linepath a a)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3640
  proof (intro exI conjI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3641
    show "a \<in> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3642
      using homotopic_with_imp_subset2 [OF homp]
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3643
      by (metis dist_0_norm image_subset_iff mem_sphere norm_one)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3644
    have teq: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3645
               \<Longrightarrow> t = Arg (exp (2 * of_real pi * of_real t * \<i>)) / (2 * pi) \<or> t=1 \<and> Arg (exp (2 * of_real pi * of_real t * \<i>)) = 0"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3646
      apply (rule disjCI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3647
      using Arg_of_real [of 1] apply (auto simp: Arg_exp)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3648
      done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3649
    have "homotopic_loops S p (p \<circ> (\<lambda>z. Arg z / (2 * pi)) \<circ> exp \<circ> (\<lambda>t. 2 * complex_of_real pi * complex_of_real t * \<i>))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3650
      apply (rule homotopic_loops_eq [OF p])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3651
      using p teq apply (fastforce simp: pathfinish_def pathstart_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3652
      done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3653
    then
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3654
    show "homotopic_loops S p (linepath a a)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3655
      by (simp add: linepath_refl  homotopic_loops_trans [OF _ homotopic_circlemaps_imp_homotopic_loops [OF homp, simplified K_record_comp]])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3656
  qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3657
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3658
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3659
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3660
proposition simply_connected_eq_homotopic_circlemaps:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3661
  fixes S :: "'a::real_normed_vector set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3662
  shows "simply_connected S \<longleftrightarrow>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3663
         (\<forall>f g::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3664
              continuous_on (sphere 0 1) f \<and> f ` (sphere 0 1) \<subseteq> S \<and>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3665
              continuous_on (sphere 0 1) g \<and> g ` (sphere 0 1) \<subseteq> S
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3666
              \<longrightarrow> homotopic_with (\<lambda>h. True) (sphere 0 1) S f g)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3667
  apply (rule iffI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3668
   apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3669
  by (simp add: simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b simply_connected_eq_homotopic_circlemaps3)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3670
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3671
proposition simply_connected_eq_contractible_circlemap:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3672
  fixes S :: "'a::real_normed_vector set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3673
  shows "simply_connected S \<longleftrightarrow>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3674
         path_connected S \<and>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3675
         (\<forall>f::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3676
              continuous_on (sphere 0 1) f \<and> f `(sphere 0 1) \<subseteq> S
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3677
              \<longrightarrow> (\<exists>a. homotopic_with (\<lambda>h. True) (sphere 0 1) S f (\<lambda>x. a)))"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3678
  apply (rule iffI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3679
   apply (simp add: simply_connected_eq_homotopic_circlemaps1 simply_connected_eq_homotopic_circlemaps2a simply_connected_eq_homotopic_circlemaps2b)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3680
  using simply_connected_eq_homotopic_circlemaps3 by blast
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3681
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3682
corollary homotopy_eqv_simple_connectedness:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3683
  fixes S :: "'a::real_normed_vector set" and T :: "'b::real_normed_vector set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3684
  shows "S homotopy_eqv T \<Longrightarrow> simply_connected S \<longleftrightarrow> simply_connected T"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3685
  by (simp add: simply_connected_eq_homotopic_circlemaps homotopy_eqv_homotopic_triviality)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  3686
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3687
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3688
subsection\<open>Homeomorphism of simple closed curves to circles\<close>
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3689
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3690
proposition homeomorphic_simple_path_image_circle:
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3691
  fixes a :: complex and \<gamma> :: "real \<Rightarrow> 'a::t2_space"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3692
  assumes "simple_path \<gamma>" and loop: "pathfinish \<gamma> = pathstart \<gamma>" and "0 < r"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3693
  shows "(path_image \<gamma>) homeomorphic sphere a r"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3694
proof -
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3695
  have "homotopic_loops (path_image \<gamma>) \<gamma> \<gamma>"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3696
    by (simp add: assms homotopic_loops_refl simple_path_imp_path)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3697
  then have hom: "homotopic_with (\<lambda>h. True) (sphere 0 1) (path_image \<gamma>)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3698
               (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi)))"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3699
    by (rule homotopic_loops_imp_homotopic_circlemaps)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3700
  have "\<exists>g. homeomorphism (sphere 0 1) (path_image \<gamma>) (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) g"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3701
  proof (rule homeomorphism_compact)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3702
    show "continuous_on (sphere 0 1) (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi)))"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3703
      using hom homotopic_with_imp_continuous by blast
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3704
    show "inj_on (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) (sphere 0 1)"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3705
    proof
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3706
      fix x y
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3707
      assume xy: "x \<in> sphere 0 1" "y \<in> sphere 0 1"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3708
         and eq: "(\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) x = (\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) y"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3709
      then have "(Arg x / (2*pi)) = (Arg y / (2*pi))"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3710
      proof -
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3711
        have "(Arg x / (2*pi)) \<in> {0..1}" "(Arg y / (2*pi)) \<in> {0..1}"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3712
          using Arg_ge_0 Arg_lt_2pi dual_order.strict_iff_order by fastforce+
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3713
        with eq show ?thesis
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3714
          using \<open>simple_path \<gamma>\<close> Arg_lt_2pi unfolding simple_path_def o_def
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3715
          by (metis eq_divide_eq_1 not_less_iff_gr_or_eq)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3716
      qed
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3717
      with xy show "x = y"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3718
        by (metis Arg Arg_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3719
    qed
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3720
    have "\<And>z. cmod z = 1 \<Longrightarrow> \<exists>x\<in>{0..1}. \<gamma> (Arg z / (2*pi)) = \<gamma> x"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3721
       by (metis Arg_ge_0 Arg_lt_2pi atLeastAtMost_iff divide_less_eq_1 less_eq_real_def zero_less_mult_iff pi_gt_zero zero_le_divide_iff zero_less_numeral)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3722
     moreover have "\<exists>z\<in>sphere 0 1. \<gamma> x = \<gamma> (Arg z / (2*pi))" if "0 \<le> x" "x \<le> 1" for x
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3723
     proof (cases "x=1")
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3724
       case True
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3725
       then show ?thesis
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3726
         apply (rule_tac x=1 in bexI)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3727
         apply (metis loop Arg_of_real divide_eq_0_iff of_real_1 pathfinish_def pathstart_def \<open>0 \<le> x\<close>, auto)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3728
         done
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3729
     next
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3730
       case False
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3731
       then have *: "(Arg (exp (\<i>*(2* of_real pi* of_real x))) / (2*pi)) = x"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3732
         using that by (auto simp: Arg_exp divide_simps)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3733
       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
  3734
         by (rule_tac x="exp(\<i> * of_real(2*pi*x))" in bexI) (auto simp: *)
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3735
    qed
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3736
    ultimately show "(\<gamma> \<circ> (\<lambda>z. Arg z / (2*pi))) ` sphere 0 1 = path_image \<gamma>"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3737
      by (auto simp: path_image_def image_iff)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3738
    qed auto
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3739
    then have "path_image \<gamma> homeomorphic sphere (0::complex) 1"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3740
      using homeomorphic_def homeomorphic_sym by blast
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3741
  also have "... homeomorphic sphere a r"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3742
    by (simp add: assms homeomorphic_spheres)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3743
  finally show ?thesis .
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3744
qed
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3745
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3746
lemma homeomorphic_simple_path_images:
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3747
  fixes \<gamma>1 :: "real \<Rightarrow> 'a::t2_space" and \<gamma>2 :: "real \<Rightarrow> 'b::t2_space"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3748
  assumes "simple_path \<gamma>1" and loop: "pathfinish \<gamma>1 = pathstart \<gamma>1"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3749
  assumes "simple_path \<gamma>2" and loop: "pathfinish \<gamma>2 = pathstart \<gamma>2"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3750
  shows "(path_image \<gamma>1) homeomorphic (path_image \<gamma>2)"
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3751
  by (meson assms homeomorphic_simple_path_image_circle homeomorphic_sym homeomorphic_trans loop pi_gt_zero)
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  3752
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  3753
end