src/HOL/Analysis/Complex_Transcendental.thy
author nipkow
Thu, 27 Dec 2018 19:48:28 +0100
changeset 69508 2a4c8a2a3f8e
parent 69180 922833cc6839
child 69529 4ab9657b3257
permissions -rw-r--r--
tuned headers; ~ -> \<not>
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
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    12
subsection\<open>Möbius transformations\<close>
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    13
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    14
(* TODO: Figure out what to do with Möbius transformations *)
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    15
definition%important "moebius a b c d = (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    16
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    17
theorem moebius_inverse:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    18
  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
    19
  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
    20
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    21
  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
    22
    by (simp add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    23
  with assms show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    24
    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
    25
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    26
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    27
lemma moebius_inverse':
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    28
  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
    29
  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
    30
  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
    31
  by (auto simp: algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
    32
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    33
lemma cmod_add_real_less:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    34
  assumes "Im z \<noteq> 0" "r\<noteq>0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    35
    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
    36
proof (cases z)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    37
  case (Complex x y)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    38
  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
    39
    apply (rule real_less_rsqrt)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    40
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    41
    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
    42
    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
    43
  then show ?thesis using assms Complex
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    44
    apply (simp add: cmod_def)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    45
    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
    46
    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
    47
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    48
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    49
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    50
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
    51
  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
    52
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    53
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    54
lemma cmod_square_less_1_plus:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    55
  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
    56
    shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    57
proof (cases "Im z = 0 \<or> Re z = 0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    58
  case True
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
    59
  with assms abs_square_less_1 show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    60
    by (force simp add: Re_power2 Im_power2 cmod_def)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    61
next
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    62
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    63
  with cmod_diff_real_less [of "1 - z\<^sup>2" "1"] show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    64
    by (simp add: norm_power Im_power2)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
    65
qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    66
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
    67
subsection%unimportant\<open>The Exponential Function\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    69
lemma norm_exp_i_times [simp]: "norm (exp(\<i> * of_real y)) = 1"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    70
  by simp
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    71
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    72
lemma norm_exp_imaginary: "norm(exp z) = 1 \<Longrightarrow> Re z = 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    73
  by simp
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
    74
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    75
lemma field_differentiable_within_exp: "exp field_differentiable (at z within s)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    76
  using DERIV_exp field_differentiable_at_within field_differentiable_def by blast
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    78
lemma continuous_within_exp:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  fixes z::"'a::{real_normed_field,banach}"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
  shows "continuous (at z within s) exp"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
by (simp add: continuous_at_imp_continuous_within)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
62381
a6479cb85944 New and revised material for (multivariate) analysis
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
    83
lemma holomorphic_on_exp [holomorphic_intros]: "exp holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
    84
  by (simp add: field_differentiable_within_exp holomorphic_on_def)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
66480
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    86
lemma holomorphic_on_exp' [holomorphic_intros]:
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    87
  "f holomorphic_on s \<Longrightarrow> (\<lambda>x. exp (f x)) holomorphic_on s"
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    88
  using holomorphic_on_compose[OF _ holomorphic_on_exp] by (simp add: o_def)
4b8d1df8933b HOL-Analysis: Convergent FPS and infinite sums
Manuel Eberl <eberlm@in.tum.de>
parents: 66466
diff changeset
    89
67968
a5ad4c015d1c removed dots at the end of (sub)titles
nipkow
parents: 67706
diff changeset
    90
subsection\<open>Euler and de Moivre formulas\<close>
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    91
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    92
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
    93
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
    94
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
    95
  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
    96
    using sin_converges sums_mult by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
  then show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
    by (simp add: scaleR_conv_of_real field_simps)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   101
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
   102
proof -
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   103
  have "(\<lambda>n. (cos_coeff n + \<i> * sin_coeff n) * z^n) = (\<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
   104
  proof
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
    fix n
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   106
    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
   107
      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
   108
  qed
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   109
  also have "... sums (exp (\<i> * z))"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
    by (rule exp_converges)
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   111
  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
   112
  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
   113
    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
   114
    by (simp add: field_simps scaleR_conv_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  ultimately show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
    using sums_unique2 by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   119
corollary%unimportant 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
   120
  using exp_Euler [of "-z"]
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
  by simp
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   123
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
   124
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   126
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
   127
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   128
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   129
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
   130
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   131
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   132
theorem Euler: "exp(z) = of_real(exp(Re z)) *
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   133
              (of_real(cos(Im z)) + \<i> * of_real(sin(Im z)))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   134
by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real Complex_eq)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   135
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   136
lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   137
  by (simp add: sin_exp_eq field_simps Re_divide Im_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   138
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   139
lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   140
  by (simp add: sin_exp_eq field_simps Im_divide Re_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   141
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   142
lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   143
  by (simp add: cos_exp_eq field_simps Re_divide Re_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   144
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   145
lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   146
  by (simp add: cos_exp_eq field_simps Im_divide Im_exp)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   147
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   148
lemma Re_sin_pos: "0 < Re z \<Longrightarrow> Re z < pi \<Longrightarrow> Re (sin z) > 0"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   149
  by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   150
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   151
lemma Im_sin_nonneg: "Re z = 0 \<Longrightarrow> 0 \<le> Im z \<Longrightarrow> 0 \<le> Im (sin z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   152
  by (simp add: Re_sin Im_sin algebra_simps)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   153
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   154
lemma Im_sin_nonneg2: "Re z = pi \<Longrightarrow> Im z \<le> 0 \<Longrightarrow> 0 \<le> Im (sin z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   155
  by (simp add: Re_sin Im_sin algebra_simps)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   156
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   157
subsection%unimportant\<open>Relationships between real and complex trigonometric and hyperbolic functions\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   158
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   159
lemma real_sin_eq [simp]: "Re(sin(of_real x)) = sin x"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   160
  by (simp add: sin_of_real)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   161
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   162
lemma real_cos_eq [simp]: "Re(cos(of_real x)) = cos x"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
  by (simp add: cos_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   165
lemma DeMoivre: "(cos z + \<i> * sin z) ^ n = cos(n * z) + \<i> * sin(n * z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   166
  by (metis exp_Euler [symmetric] exp_of_nat_mult mult.left_commute)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   167
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   168
lemma exp_cnj: "cnj (exp z) = exp (cnj z)"
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
proof -
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
  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
   171
    by auto
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
  also have "... sums (exp (cnj z))"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
    by (rule exp_converges)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
  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
   175
  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
   176
    by (metis exp_converges sums_cnj)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   177
  ultimately show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   178
    using sums_unique2
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   179
    by blast
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   180
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   181
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   182
lemma cnj_sin: "cnj(sin z) = sin(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   183
  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
   184
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   185
lemma cnj_cos: "cnj(cos z) = cos(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   186
  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
   187
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   188
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
   189
  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
   190
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   191
lemma field_differentiable_within_sin: "sin field_differentiable (at z within S)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   192
  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
   193
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   194
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
   195
  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
   196
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   197
lemma field_differentiable_within_cos: "cos field_differentiable (at z within S)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
   198
  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
   199
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   200
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
   201
  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
   202
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   203
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
   204
  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
   205
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   206
lemma holomorphic_on_sin' [holomorphic_intros]:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   207
  assumes "f holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   208
  shows   "(\<lambda>x. sin (f x)) holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   209
  using holomorphic_on_compose[OF assms holomorphic_on_sin] by (simp add: o_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   210
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   211
lemma holomorphic_on_cos' [holomorphic_intros]:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   212
  assumes "f holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   213
  shows   "(\<lambda>x. cos (f x)) holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   214
  using holomorphic_on_compose[OF assms holomorphic_on_cos] by (simp add: o_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
   215
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   216
subsection%unimportant\<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
   217
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   218
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
   219
  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
   220
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   221
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
   222
                 (is "?lhs = ?rhs")
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   223
proof
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   224
  assume "exp z = 1"
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   225
  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
   226
    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
   227
  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
   228
    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
   229
next
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   230
  assume ?rhs then show ?lhs
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   231
    using Im_exp Re_exp complex_eq_iff
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   232
    by (simp add: cos_one_2pi_int cos_one_sin_zero mult.commute)
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   233
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   234
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   235
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
   236
                (is "?lhs = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   237
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   238
  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
   239
    by (simp add: exp_diff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   240
  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
   241
    by (simp add: exp_eq_1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   242
  also have "... \<longleftrightarrow> ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   243
    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
   244
  finally show ?thesis .
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   245
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   246
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   247
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
   248
  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
   249
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   250
lemma exp_integer_2pi:
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   251
  assumes "n \<in> \<int>"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   252
  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
   253
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   254
  have "exp((2 * n * pi) * \<i>) = exp 0"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   255
    using assms unfolding Ints_def exp_eq by auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   256
  also have "... = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   257
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   258
  finally show ?thesis .
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   259
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   260
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   261
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
   262
  by (simp add: exp_eq)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   263
66466
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   264
lemma exp_integer_2pi_plus1:
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   265
  assumes "n \<in> \<int>"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   266
  shows "exp(((2 * n + 1) * pi) * \<i>) = - 1"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   267
proof -
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   268
  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
   269
    by (auto simp: Ints_def)
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   270
  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
   271
    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
   272
  also have "... = - 1"
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   273
    by simp
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   274
  finally show ?thesis .
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   275
qed
aec5d9c88d69 More lemmas for HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 66453
diff changeset
   276
64287
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   277
lemma inj_on_exp_pi:
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   278
  fixes z::complex shows "inj_on exp (ball z pi)"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   279
proof (clarsimp simp: inj_on_def exp_eq)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   280
  fix y n
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   281
  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
   282
         "dist z y < pi"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   283
  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
   284
    using dist_commute_lessI dist_triangle_less_add by blast
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   285
  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
   286
    by (simp add: dist_norm)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   287
  then show "n = 0"
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   288
    by (auto simp: norm_mult)
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   289
qed
d85d88722745 more from moretop.ml
paulson <lp15@cam.ac.uk>
parents: 64267
diff changeset
   290
68585
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   291
lemma cmod_add_squared:
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   292
  fixes r1 r2::real
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   293
  assumes "r1 \<ge> 0" "r2 \<ge> 0"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   294
  shows "(cmod (r1 * exp (\<i> * \<theta>1) + r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 + 2 * r1 * r2 * cos (\<theta>1 - \<theta>2)" (is "(cmod (?z1 + ?z2))\<^sup>2 = ?rhs")
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   295
proof -
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   296
  have "(cmod (?z1 + ?z2))\<^sup>2 = (?z1 + ?z2) * cnj (?z1 + ?z2)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   297
    by (rule complex_norm_square)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   298
  also have "\<dots> = (?z1 * cnj ?z1 + ?z2 * cnj ?z2) + (?z1 * cnj ?z2 + cnj ?z1 * ?z2)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   299
    by (simp add: algebra_simps)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   300
  also have "\<dots> = (norm ?z1)\<^sup>2 + (norm ?z2)\<^sup>2 + 2 * Re (?z1 * cnj ?z2)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   301
    unfolding complex_norm_square [symmetric] cnj_add_mult_eq_Re by simp
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   302
  also have "\<dots> = ?rhs"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   303
    by (simp add: norm_mult) (simp add: exp_Euler complex_is_Real_iff [THEN iffD1] cos_diff algebra_simps)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   304
  finally show ?thesis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   305
    using of_real_eq_iff by blast
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   306
qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   307
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   308
lemma cmod_diff_squared:
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   309
  fixes r1 r2::real
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   310
  assumes "r1 \<ge> 0" "r2 \<ge> 0"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   311
  shows "(cmod (r1 * exp (\<i> * \<theta>1) - r2 * exp (\<i> * \<theta>2)))\<^sup>2 = r1\<^sup>2 + r2\<^sup>2 - 2*r1*r2*cos (\<theta>1 - \<theta>2)" (is "(cmod (?z1 - ?z2))\<^sup>2 = ?rhs")
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   312
proof -
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   313
  have "exp (\<i> * (\<theta>2 + pi)) = - exp (\<i> * \<theta>2)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   314
    by (simp add: exp_Euler cos_plus_pi sin_plus_pi)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   315
  then have "(cmod (?z1 - ?z2))\<^sup>2 = cmod (?z1 + r2 * exp (\<i> * (\<theta>2 + pi))) ^2"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   316
    by simp
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   317
  also have "\<dots> = r1\<^sup>2 + r2\<^sup>2 + 2*r1*r2*cos (\<theta>1 - (\<theta>2 + pi))"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   318
    using assms cmod_add_squared by blast
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   319
  also have "\<dots> = ?rhs"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   320
    by (simp add: add.commute diff_add_eq_diff_diff_swap)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   321
  finally show ?thesis .
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   322
qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   323
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   324
lemma polar_convergence:
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   325
  fixes R::real
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   326
  assumes "\<And>j. r j > 0" "R > 0"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   327
  shows "((\<lambda>j. r j * exp (\<i> * \<theta> j)) \<longlonglongrightarrow> (R * exp (\<i> * \<Theta>))) \<longleftrightarrow>
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   328
         (r \<longlonglongrightarrow> R) \<and> (\<exists>k. (\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>)"    (is "(?z \<longlonglongrightarrow> ?Z) = ?rhs")
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   329
proof
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   330
  assume L: "?z \<longlonglongrightarrow> ?Z"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   331
  have rR: "r \<longlonglongrightarrow> R"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   332
    using tendsto_norm [OF L] assms by (auto simp: norm_mult abs_of_pos)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   333
  moreover obtain k where "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   334
  proof -
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   335
    have "cos (\<theta> j - \<Theta>) = ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)" for j
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   336
      apply (subst cmod_diff_squared)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   337
      using assms by (auto simp: divide_simps less_le)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   338
    moreover have "(\<lambda>j. ((r j)\<^sup>2 + R\<^sup>2 - (norm(?z j - ?Z))\<^sup>2) / (2 * R * r j)) \<longlonglongrightarrow> ((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R))"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   339
      by (intro L rR tendsto_intros) (use \<open>R > 0\<close> in force)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   340
    moreover have "((R\<^sup>2 + R\<^sup>2 - (norm(?Z - ?Z))\<^sup>2) / (2 * R * R)) = 1"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   341
      using \<open>R > 0\<close> by (simp add: power2_eq_square divide_simps)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   342
    ultimately have "(\<lambda>j. cos (\<theta> j - \<Theta>)) \<longlonglongrightarrow> 1"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   343
      by auto
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   344
    then show ?thesis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   345
      using that cos_diff_limit_1 by blast
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   346
  qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   347
  ultimately show ?rhs
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   348
    by metis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   349
next
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   350
  assume R: ?rhs
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   351
  show "?z \<longlonglongrightarrow> ?Z"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   352
  proof (rule tendsto_mult)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   353
    show "(\<lambda>x. complex_of_real (r x)) \<longlonglongrightarrow> of_real R"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   354
      using R by (auto simp: tendsto_of_real_iff)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   355
    obtain k where "(\<lambda>j. \<theta> j - of_int (k j) * (2 * pi)) \<longlonglongrightarrow> \<Theta>"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   356
      using R by metis
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   357
    then have "(\<lambda>j. complex_of_real (\<theta> j - of_int (k j) * (2 * pi))) \<longlonglongrightarrow> of_real \<Theta>"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   358
      using tendsto_of_real_iff by force
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   359
    then have "(\<lambda>j.  exp (\<i> * of_real (\<theta> j - of_int (k j) * (2 * pi)))) \<longlonglongrightarrow> exp (\<i> * \<Theta>)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   360
      using tendsto_mult [OF tendsto_const] isCont_exp isCont_tendsto_compose by blast
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   361
    moreover have "exp (\<i> * of_real (\<theta> j - of_int (k j) * (2 * pi))) = exp (\<i> * \<theta> j)" for j
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   362
      unfolding exp_eq
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   363
      by (rule_tac x="- k j" in exI) (auto simp: algebra_simps)
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   364
    ultimately show "(\<lambda>j. exp (\<i> * \<theta> j)) \<longlonglongrightarrow> exp (\<i> * \<Theta>)"
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   365
      by auto
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   366
  qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   367
qed
1657b9a5dd5d more on infinite products
paulson <lp15@cam.ac.uk>
parents: 68535
diff changeset
   368
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   369
lemma sin_cos_eq_iff: "sin y = sin x \<and> cos y = cos x \<longleftrightarrow> (\<exists>n::int. y = x + 2 * pi * n)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   370
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   371
  { 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
   372
    then have "cos (y-x) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   373
      using cos_add [of y "-x"] by simp
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   374
    then have "\<exists>n::int. y-x = 2 * pi * n"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   375
      using cos_one_2pi_int by auto }
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   376
  then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   377
  apply (auto simp: sin_add cos_add)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   378
  apply (metis add.commute diff_add_cancel)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   379
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   380
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   381
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   382
lemma exp_i_ne_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   383
  assumes "0 < x" "x < 2*pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   384
  shows "exp(\<i> * of_real x) \<noteq> 1"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   385
proof
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   386
  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
   387
  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
   388
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   389
  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
   390
    by (simp only: Ints_def exp_eq) auto
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   391
  then have "of_real x = (of_int (2 * n) * pi)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   392
    by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   393
  then have "x = (of_int (2 * n) * pi)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   394
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   395
  then show False using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   396
    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
   397
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   398
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   399
lemma sin_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   400
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   401
  shows "sin z = 0 \<longleftrightarrow> (\<exists>n::int. z = of_real(n * pi))"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   402
  by (simp add: sin_exp_eq exp_eq)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   403
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   404
lemma cos_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   405
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   406
  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
   407
  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
   408
  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
   409
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   410
lemma cos_eq_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   411
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   412
  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
   413
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   414
  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
   415
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   416
  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
   417
    by (simp only: cos_double_sin)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   418
  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
   419
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   420
  show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   421
    by (auto simp: sin_eq_0)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   422
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   423
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   424
lemma csin_eq_1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   425
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   426
  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
   427
  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
   428
  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
   429
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   430
lemma csin_eq_minus1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   431
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   432
  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
   433
        (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   434
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   435
  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
   436
    by (simp add: equation_minus_iff)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   437
  also have "... \<longleftrightarrow> (\<exists>n::int. -z = of_real(2 * n * pi) + of_real pi/2)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   438
    by (simp only: csin_eq_1)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   439
  also have "... \<longleftrightarrow> (\<exists>n::int. z = - of_real(2 * n * pi) - of_real pi/2)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   440
    apply (rule iff_exI)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   441
    by (metis (no_types) is_num_normalize(8) minus_minus of_real_def real_vector.scale_minus_left uminus_add_conv_diff)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   442
  also have "... = ?rhs"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   443
    apply safe
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   444
    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
   445
    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
   446
    apply (simp_all add: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   447
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   448
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   449
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   450
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   451
lemma ccos_eq_minus1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   452
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   453
  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
   454
  using csin_eq_1 [of "z - of_real pi/2"]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   455
  by (simp add: sin_diff algebra_simps equation_minus_iff)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   456
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   457
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
   458
                (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   459
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   460
  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
   461
    by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   462
  also have "... \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   463
    by (simp only: csin_eq_1)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   464
  also have "... \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + of_real pi/2)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   465
    by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   466
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   467
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   468
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   469
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   470
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   471
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
   472
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   473
  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
   474
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   475
  also have "... \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   476
    by (simp only: csin_eq_minus1)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   477
  also have "... \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + 3/2*pi)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   478
    by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   479
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   480
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   481
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   482
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   483
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   484
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
   485
                      (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   486
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   487
  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
   488
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   489
  also have "... \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + pi)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   490
    by (simp only: ccos_eq_minus1)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   491
  also have "... \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + pi)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   492
    by (rule iff_exI) (auto simp: algebra_simps intro: injD [OF inj_of_real [where 'a = complex]])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   493
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   494
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   495
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   496
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   497
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
   498
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
   499
  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
   500
  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
   501
  apply (simp add: real_sqrt_mult)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   502
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   503
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   504
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   505
lemma complex_sin_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   506
  fixes w :: complex
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   507
  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
   508
        (is "?lhs = ?rhs")
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   509
proof
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   510
  assume ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   511
  then have "sin w - sin z = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   512
    by (auto simp: algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   513
  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
   514
    by (auto simp: sin_diff_sin)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   515
  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
   516
    using mult_eq_0_iff by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   517
  then show ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   518
  proof cases
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   519
    case 1
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   520
    then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   521
      by (simp add: sin_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   522
  next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   523
    case 2
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   524
    then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   525
      by (simp add: cos_eq_0 algebra_simps) (metis Ints_of_int of_real_of_int_eq)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   526
  qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   527
next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   528
  assume ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   529
  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
   530
                               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
   531
    using Ints_cases by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   532
  then show ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   533
    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
   534
    apply (auto simp: algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   535
    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
   536
              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
   537
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   538
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   539
lemma complex_cos_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   540
  fixes w :: complex
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   541
  shows "cos w = cos z \<longleftrightarrow> (\<exists>n \<in> \<int>. w = z + of_real(2*n*pi) \<or> w = -z + of_real(2*n*pi))"
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   542
        (is "?lhs = ?rhs")
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   543
proof
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   544
  assume ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   545
  then have "cos w - cos z = 0"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   546
    by (auto simp: algebra_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   547
  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
   548
    by (auto simp: cos_diff_cos)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   549
  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
   550
    using mult_eq_0_iff by blast
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   551
  then show ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   552
  proof cases
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   553
    case 1
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   554
    then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   555
      apply (simp add: sin_eq_0 algebra_simps)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   556
      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
   557
  next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   558
    case 2
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   559
    then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   560
      apply (clarsimp simp: sin_eq_0 algebra_simps)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   561
      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
   562
  qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   563
next
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   564
  assume ?rhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   565
  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
   566
                               w = -z + of_real(2*n*pi)"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   567
    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
   568
  then show ?lhs
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   569
    using Periodic_Fun.cos.plus_of_int [of z n]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   570
    apply (simp add: algebra_simps)
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   571
    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
   572
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   573
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   574
lemma sin_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   575
   "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
   576
  using complex_sin_eq [of x y]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   577
  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
   578
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   579
lemma cos_eq:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   580
   "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
   581
  using complex_cos_eq [of x y]
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
   582
  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
   583
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   584
lemma sinh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   585
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   586
  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
   587
  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
   588
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
   589
lemma sin_i_times:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   590
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   591
  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
   592
  using sinh_complex by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   593
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   594
lemma sinh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   595
  fixes x :: real
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   596
  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
   597
  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
   598
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   599
lemma cosh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   600
  fixes z :: complex
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   601
  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
   602
  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
   603
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   604
lemma cosh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   605
  fixes x :: real
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
   606
  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
   607
  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
   608
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
   609
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
   610
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   611
lemma norm_cos_squared:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   612
    "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
   613
  apply (cases z)
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   614
  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
   615
  apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   616
  apply (simp only: left_diff_distrib [symmetric] power_mult_distrib sin_squared_eq)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   617
  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
   618
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   619
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   620
lemma norm_sin_squared:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   621
    "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
   622
  apply (cases z)
65274
db2de50de28e Removed [simp] status for Complex_eq. Also tidied some proofs
paulson <lp15@cam.ac.uk>
parents: 65064
diff changeset
   623
  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
   624
  apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   625
  apply (simp only: left_diff_distrib [symmetric] power_mult_distrib cos_squared_eq)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   626
  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
   627
  done
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   628
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   629
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
   630
  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
   631
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   632
lemma norm_cos_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   633
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   634
  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
   635
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   636
  have "Im z \<le> cmod z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   637
    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
   638
  with exp_uminus_Im show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   639
    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
   640
    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
   641
    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
   642
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   643
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   644
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   645
lemma norm_cos_plus1_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   646
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   647
  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
   648
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   649
  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
   650
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   651
  have *: "Im z \<le> cmod z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   652
    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
   653
  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
   654
    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
   655
  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
   656
    by (simp add: cos_exp_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   657
  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
   658
    by (simp add: field_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   659
  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
   660
    by (simp add: norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   661
  finally show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   662
    by (metis exp_eq_one_iff exp_le_cancel_iff mult_2 norm_cos_le norm_ge_zero norm_one norm_triangle_mono)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   663
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   664
67578
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   665
lemma sinh_conv_sin: "sinh z = -\<i> * sin (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   666
  by (simp add: sinh_field_def sin_i_times exp_minus)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   667
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   668
lemma cosh_conv_cos: "cosh z = cos (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   669
  by (simp add: cosh_field_def cos_i_times exp_minus)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   670
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   671
lemma tanh_conv_tan: "tanh z = -\<i> * tan (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   672
  by (simp add: tanh_def sinh_conv_sin cosh_conv_cos tan_def)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   673
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   674
lemma sin_conv_sinh: "sin z = -\<i> * sinh (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   675
  by (simp add: sinh_conv_sin)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   676
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   677
lemma cos_conv_cosh: "cos z = cosh (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   678
  by (simp add: cosh_conv_cos)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   679
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   680
lemma tan_conv_tanh: "tan z = -\<i> * tanh (\<i>*z)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   681
  by (simp add: tan_def sin_conv_sinh cos_conv_cosh tanh_def)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   682
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   683
lemma sinh_complex_eq_iff:
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   684
  "sinh (z :: complex) = sinh w \<longleftrightarrow>
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   685
      (\<exists>n\<in>\<int>. z = w - 2 * \<i> * of_real n * of_real pi \<or>
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   686
              z = -(2 * complex_of_real n + 1) * \<i> * complex_of_real pi - w)" (is "_ = ?rhs")
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   687
proof -
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   688
  have "sinh z = sinh w \<longleftrightarrow> sin (\<i> * z) = sin (\<i> * w)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   689
    by (simp add: sinh_conv_sin)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   690
  also have "\<dots> \<longleftrightarrow> ?rhs"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   691
    by (subst complex_sin_eq) (force simp: field_simps complex_eq_iff)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   692
  finally show ?thesis .
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   693
qed
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   694
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
   695
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   696
subsection%unimportant\<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
   697
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   698
declare power_Suc [simp del]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   699
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   700
lemma Taylor_exp_field:
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   701
  fixes z::"'a::{banach,real_normed_field}"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   702
  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
   703
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
   704
  show "convex (closed_segment 0 z)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   705
    by (rule convex_closed_segment [of 0 z])
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   706
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   707
  fix k x
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   708
  assume "x \<in> closed_segment 0 z" "k \<le> n"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   709
  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
   710
    using DERIV_exp DERIV_subset by blast
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   711
next
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   712
  fix x
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   713
  assume x: "x \<in> closed_segment 0 z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   714
  have "norm (exp x) \<le> exp (norm x)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   715
    by (rule norm_exp)
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   716
  also have "norm x \<le> norm z"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   717
    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
   718
  finally show "norm (exp x) \<le> exp (norm z)"
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   719
    by simp
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   720
qed auto
66252
b73f94b366b7 some generalizations complex=>real_normed_field
immler
parents: 65719
diff changeset
   721
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   722
lemma Taylor_exp:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   723
  "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
   724
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
   725
  show "convex (closed_segment 0 z)"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   726
    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
   727
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   728
  fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   729
  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
   730
  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
   731
    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
   732
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   733
  fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   734
  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
   735
  then show "Re x \<le> \<bar>Re z\<bar>"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   736
    apply (clarsimp simp: closed_segment_def scaleR_conv_of_real)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   737
    by (meson abs_ge_self abs_ge_zero linear mult_left_le_one_le mult_nonneg_nonpos order_trans)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   738
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   739
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   740
lemma
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   741
  assumes "0 \<le> u" "u \<le> 1"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   742
  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
   743
    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
   744
proof -
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   745
  have mono: "\<And>u w z::real. w \<le> u \<Longrightarrow> z \<le> u \<Longrightarrow> (w + z)/2 \<le> u"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   746
    by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   747
  have *: "(cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2 \<le> exp \<bar>Im z\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   748
  proof (rule mono)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   749
    show "cmod (exp (\<i> * (u * z))) \<le> exp \<bar>Im z\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   750
      apply (simp add: abs_if mult_left_le_one_le assms not_less)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   751
      by (meson assms(1) mult_nonneg_nonneg neg_le_0_iff_le order_trans)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   752
    show "cmod (exp (- (\<i> * (u * z)))) \<le> exp \<bar>Im z\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   753
      apply (simp add: abs_if mult_left_le_one_le assms)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   754
      by (meson \<open>0 \<le> u\<close> less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   755
  qed
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   756
  have "cmod (sin (u *\<^sub>R z)) = cmod (exp (\<i> * (u * z)) - exp (- (\<i> * (u * z)))) / 2"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   757
    by (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   758
  also have "... \<le> (cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   759
    by (intro divide_right_mono norm_triangle_ineq4) simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   760
  also have "... \<le> exp \<bar>Im z\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   761
    by (rule *)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   762
  finally show "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" .
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   763
  have "cmod (cos (u *\<^sub>R z)) = cmod (exp (\<i> * (u * z)) + exp (- (\<i> * (u * z)))) / 2"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   764
    by (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   765
  also have "... \<le> (cmod (exp (\<i> * (u * z))) + cmod (exp (- (\<i> * (u * z)))) ) / 2"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   766
    by (intro divide_right_mono norm_triangle_ineq) simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   767
  also have "... \<le> exp \<bar>Im z\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   768
    by (rule *)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   769
  finally show "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" .
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   770
qed
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   771
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   772
lemma Taylor_sin:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   773
  "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
   774
   \<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
   775
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   776
  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
   777
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   778
  have *: "cmod (sin z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   779
                 (\<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
   780
           \<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
   781
  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
   782
                                 "\<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
   783
                                 "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
   784
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   785
    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
   786
            (- 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
   787
            (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
   788
      apply (auto simp: power_Suc)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   789
      apply (intro derivative_eq_intros | simp)+
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   790
      done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   791
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   792
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   793
    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
   794
    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
   795
      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
   796
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   797
  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
   798
            = (-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
   799
    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
   800
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   801
    apply (rule order_trans [OF _ *])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   802
    apply (simp add: **)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   803
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   804
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   805
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   806
lemma Taylor_cos:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   807
  "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
   808
   \<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
   809
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   810
  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
   811
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   812
  have *: "cmod (cos z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   813
                 (\<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
   814
           \<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
   815
  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
   816
simplified])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   817
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   818
    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
   819
    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
   820
            (- 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
   821
             (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
   822
      apply (auto simp: power_Suc)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   823
      apply (intro derivative_eq_intros | simp)+
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   824
      done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   825
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   826
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   827
    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
   828
    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
   829
      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
   830
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   831
  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
   832
            = (-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
   833
    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
   834
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   835
    apply (rule order_trans [OF _ *])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   836
    apply (simp add: **)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   837
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   838
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   839
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
   840
declare power_Suc [simp]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   841
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   842
text\<open>32-bit Approximation to e\<close>
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   843
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
   844
  using Taylor_exp [of 1 14] exp_le
64267
b9a1486e79be setsum -> sum
nipkow
parents: 64240
diff changeset
   845
  apply (simp add: sum_distrib_right in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
66611
c375b64a6c24 adapted to better linear arith
nipkow
parents: 66480
diff changeset
   846
  apply (simp only: pos_le_divide_eq [symmetric])
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   847
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   848
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   849
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
   850
  using e_approx_32
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
   851
  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
   852
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   853
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
   854
  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
   855
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   856
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
   857
lemma ln3_gt_1: "ln 3 > (1::real)"
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
   858
  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
   859
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   860
subsection\<open>The argument of a complex number (HOL Light version)\<close>
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   861
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   862
definition%important is_Arg :: "[complex,real] \<Rightarrow> bool"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   863
  where "is_Arg z r \<equiv> z = of_real(norm z) * exp(\<i> * of_real r)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   864
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   865
definition%important Arg2pi :: "complex \<Rightarrow> real"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   866
  where "Arg2pi z \<equiv> if z = 0 then 0 else THE t. 0 \<le> t \<and> t < 2*pi \<and> is_Arg z t"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   867
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   868
lemma is_Arg_2pi_iff: "is_Arg z (r + of_int k * (2 * pi)) \<longleftrightarrow> is_Arg z r"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   869
  by (simp add: algebra_simps is_Arg_def)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   870
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   871
lemma is_Arg_eqI:
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   872
  assumes r: "is_Arg z r" and s: "is_Arg z s" and rs: "abs (r-s) < 2*pi" and "z \<noteq> 0"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   873
  shows "r = s"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   874
proof -
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   875
  have zr: "z = (cmod z) * exp (\<i> * r)" and zs: "z = (cmod z) * exp (\<i> * s)"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   876
    using r s by (auto simp: is_Arg_def)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   877
  with \<open>z \<noteq> 0\<close> have eq: "exp (\<i> * r) = exp (\<i> * s)"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   878
    by (metis mult_eq_0_iff vector_space_over_itself.scale_left_imp_eq)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   879
  have  "\<i> * r = \<i> * s"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   880
    by (rule exp_complex_eqI) (use rs in \<open>auto simp: eq exp_complex_eqI\<close>)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   881
  then show ?thesis
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   882
    by simp
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   883
qed
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
   884
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   885
text\<open>This function returns the angle of a complex number from its representation in polar coordinates.
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   886
Due to periodicity, its range is arbitrary. @{term Arg2pi} follows HOL Light in adopting the interval $[0,2\pi)$.
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   887
But we have the same periodicity issue with logarithms, and it is usual to adopt the same interval
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   888
for the complex logarithm and argument functions. Further on down, we shall define both functions for the interval $(-\pi,\pi]$.
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   889
The present version is provided for compatibility.\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   890
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   891
lemma Arg2pi_0 [simp]: "Arg2pi(0) = 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   892
  by (simp add: Arg2pi_def)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   893
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   894
lemma Arg2pi_unique_lemma:
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   895
  assumes z:  "is_Arg z t"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   896
      and z': "is_Arg z t'"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   897
      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
   898
      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
   899
      and nz: "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   900
  shows "t' = t"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   901
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   902
  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
   903
    by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   904
  have "of_real (cmod z) * exp (\<i> * of_real t') = of_real (cmod z) * exp (\<i> * of_real t)"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   905
    by (metis z z' is_Arg_def)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   906
  then have "exp (\<i> * of_real t') = exp (\<i> * of_real t)"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   907
    by (metis nz mult_left_cancel mult_zero_left z is_Arg_def)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   908
  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
   909
    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
   910
    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
   911
  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
   912
    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
   913
  then have "n=0"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   914
    by (cases n) (use t t' in \<open>auto simp: mult_less_0_iff algebra_simps\<close>)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   915
  then show "t' = t"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   916
    by (simp add: n)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   917
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   918
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   919
lemma Arg2pi: "0 \<le> Arg2pi z \<and> Arg2pi z < 2*pi \<and> is_Arg z (Arg2pi z)"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   920
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   921
  case True then show ?thesis
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   922
    by (simp add: Arg2pi_def is_Arg_def)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   923
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   924
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   925
  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
   926
             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
   927
    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
   928
    by blast
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   929
  have z: "is_Arg z t"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   930
    unfolding is_Arg_def
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   931
    apply (rule complex_eqI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   932
    using t False ReIm
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   933
    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
   934
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   935
  show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   936
    apply (simp add: Arg2pi_def False)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   937
    apply (rule theI [where a=t])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   938
    using t z False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   939
    apply (auto intro: Arg2pi_unique_lemma)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   940
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   941
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   942
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
   943
corollary%unimportant
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   944
  shows Arg2pi_ge_0: "0 \<le> Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   945
    and Arg2pi_lt_2pi: "Arg2pi z < 2*pi"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   946
    and Arg2pi_eq: "z = of_real(norm z) * exp(\<i> * of_real(Arg2pi z))"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   947
  using Arg2pi is_Arg_def by auto
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   948
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   949
lemma complex_norm_eq_1_exp: "norm z = 1 \<longleftrightarrow> exp(\<i> * of_real (Arg2pi z)) = z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   950
  by (metis Arg2pi_eq cis_conv_exp mult.left_neutral norm_cis of_real_1)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   951
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   952
lemma Arg2pi_unique: "\<lbrakk>of_real r * exp(\<i> * of_real a) = z; 0 < r; 0 \<le> a; a < 2*pi\<rbrakk> \<Longrightarrow> Arg2pi z = a"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   953
  by (rule Arg2pi_unique_lemma [unfolded is_Arg_def, OF _ Arg2pi_eq]) (use Arg2pi [of z] in \<open>auto simp: norm_mult\<close>)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   954
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   955
lemma Arg2pi_minus: "z \<noteq> 0 \<Longrightarrow> Arg2pi (-z) = (if Arg2pi z < pi then Arg2pi z + pi else Arg2pi z - pi)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   956
  apply (rule Arg2pi_unique [of "norm z"])
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   957
  apply (rule complex_eqI)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   958
  using Arg2pi_ge_0 [of z] Arg2pi_eq [of z] Arg2pi_lt_2pi [of z]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   959
  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
   960
  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
   961
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   962
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   963
lemma Arg2pi_times_of_real [simp]:
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   964
  assumes "0 < r" shows "Arg2pi (of_real r * z) = Arg2pi z"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   965
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   966
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   967
  show ?thesis
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
   968
    by (rule Arg2pi_unique [of "r * norm z"]) (use Arg2pi False assms is_Arg_def in auto)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   969
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   970
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   971
lemma Arg2pi_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg2pi (z * of_real r) = Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   972
  by (metis Arg2pi_times_of_real mult.commute)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   973
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   974
lemma Arg2pi_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg2pi (z / of_real r) = Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   975
  by (metis Arg2pi_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   976
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   977
lemma Arg2pi_le_pi: "Arg2pi z \<le> pi \<longleftrightarrow> 0 \<le> Im z"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   978
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   979
  case False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   980
  have "0 \<le> Im z \<longleftrightarrow> 0 \<le> Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg2pi z)))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   981
    by (metis Arg2pi_eq)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   982
  also have "... = (0 \<le> Im (exp (\<i> * complex_of_real (Arg2pi z))))"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   983
    using False  by (simp add: zero_le_mult_iff)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   984
  also have "... \<longleftrightarrow> Arg2pi z \<le> pi"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   985
    by (simp add: Im_exp) (metis Arg2pi_ge_0 Arg2pi_lt_2pi sin_lt_zero sin_ge_zero not_le)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   986
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   987
    by blast
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   988
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   989
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   990
lemma Arg2pi_lt_pi: "0 < Arg2pi z \<and> Arg2pi z < pi \<longleftrightarrow> 0 < Im z"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   991
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   992
  case False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   993
  have "0 < Im z \<longleftrightarrow> 0 < Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg2pi z)))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   994
    by (metis Arg2pi_eq)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   995
  also have "... = (0 < Im (exp (\<i> * complex_of_real (Arg2pi z))))"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
   996
    using False by (simp add: zero_less_mult_iff)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   997
  also have "... \<longleftrightarrow> 0 < Arg2pi z \<and> Arg2pi z < pi"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
   998
    using Arg2pi_ge_0 Arg2pi_lt_2pi sin_le_zero sin_gt_zero
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   999
    apply (auto simp: Im_exp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1000
    using le_less apply fastforce
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1001
    using not_le by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1002
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1003
    by blast
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1004
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1005
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1006
lemma Arg2pi_eq_0: "Arg2pi z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1007
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1008
  case False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1009
  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 (Arg2pi z)))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1010
    by (metis Arg2pi_eq)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1011
  also have "... \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg2pi z)))"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1012
    using False  by (simp add: zero_le_mult_iff)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1013
  also have "... \<longleftrightarrow> Arg2pi z = 0"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1014
  proof -
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1015
    have [simp]: "Arg2pi z = 0 \<Longrightarrow> z \<in> \<real>"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1016
      using Arg2pi_eq [of z] by (auto simp: Reals_def)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1017
    moreover have "\<lbrakk>z \<in> \<real>; 0 \<le> cos (Arg2pi z)\<rbrakk> \<Longrightarrow> Arg2pi z = 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1018
      by (metis Arg2pi_lt_pi Arg2pi_ge_0 Arg2pi_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1019
    ultimately show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1020
      by (auto simp: Re_exp)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1021
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1022
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1023
    by blast
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1024
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1025
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1026
corollary%unimportant Arg2pi_gt_0:
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1027
  assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1028
    shows "Arg2pi z > 0"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1029
  using Arg2pi_eq_0 Arg2pi_ge_0 assms dual_order.strict_iff_order
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1030
  unfolding nonneg_Reals_def by fastforce
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1031
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1032
lemma Arg2pi_eq_pi: "Arg2pi z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1033
    using Arg2pi_le_pi [of z] Arg2pi_lt_pi [of z] Arg2pi_eq_0 [of z] Arg2pi_ge_0 [of z] 
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1034
    by (fastforce simp: complex_is_Real_iff)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1035
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1036
lemma Arg2pi_eq_0_pi: "Arg2pi z = 0 \<or> Arg2pi z = pi \<longleftrightarrow> z \<in> \<real>"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1037
  using Arg2pi_eq_0 Arg2pi_eq_pi not_le by auto
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1038
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1039
lemma Arg2pi_of_real: "Arg2pi (of_real r) = (if r<0 then pi else 0)"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1040
  using Arg2pi_eq_0_pi Arg2pi_eq_pi by fastforce
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1041
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1042
lemma Arg2pi_real: "z \<in> \<real> \<Longrightarrow> Arg2pi z = (if 0 \<le> Re z then 0 else pi)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1043
  using Arg2pi_eq_0 Arg2pi_eq_0_pi by auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1044
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1045
lemma Arg2pi_inverse: "Arg2pi(inverse z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1046
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1047
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1048
  show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1049
    apply (rule Arg2pi_unique [of "inverse (norm z)"])
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1050
    using Arg2pi_eq False Arg2pi_ge_0 [of z] Arg2pi_lt_2pi [of z] Arg2pi_eq_0 [of z]
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1051
    by (auto simp: Arg2pi_real in_Reals_norm exp_diff field_simps)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1052
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1053
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1054
lemma Arg2pi_eq_iff:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1055
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1056
     shows "Arg2pi w = Arg2pi z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1057
  using assms Arg2pi_eq [of z] Arg2pi_eq [of w]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1058
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1059
  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
  1060
  apply (simp add: divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1061
  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
  1062
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1063
lemma Arg2pi_inverse_eq_0: "Arg2pi(inverse z) = 0 \<longleftrightarrow> Arg2pi z = 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1064
  by (metis Arg2pi_eq_0 Arg2pi_inverse inverse_inverse_eq)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1065
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1066
lemma Arg2pi_divide:
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1067
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg2pi w \<le> Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1068
    shows "Arg2pi(z / w) = Arg2pi z - Arg2pi w"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1069
  apply (rule Arg2pi_unique [of "norm(z / w)"])
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1070
  using assms Arg2pi_eq Arg2pi_ge_0 [of w] Arg2pi_lt_2pi [of z]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1071
  apply (auto simp: exp_diff norm_divide field_simps)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1072
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1073
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1074
lemma Arg2pi_le_div_sum:
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1075
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg2pi w \<le> Arg2pi z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1076
    shows "Arg2pi z = Arg2pi w + Arg2pi(z / w)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1077
  by (simp add: Arg2pi_divide assms)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1078
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1079
lemma Arg2pi_le_div_sum_eq:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1080
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1081
    shows "Arg2pi w \<le> Arg2pi z \<longleftrightarrow> Arg2pi z = Arg2pi w + Arg2pi(z / w)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1082
  using assms by (auto simp: Arg2pi_ge_0 intro: Arg2pi_le_div_sum)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1083
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1084
lemma Arg2pi_diff:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1085
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1086
    shows "Arg2pi w - Arg2pi z = (if Arg2pi z \<le> Arg2pi w then Arg2pi(w / z) else Arg2pi(w/z) - 2*pi)"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1087
  using assms Arg2pi_divide Arg2pi_inverse [of "w/z"] Arg2pi_eq_0_pi
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1088
  by (force simp add: Arg2pi_ge_0 Arg2pi_divide not_le split: if_split_asm)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1089
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1090
lemma Arg2pi_add:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1091
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1092
    shows "Arg2pi w + Arg2pi z = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi(w * z) else Arg2pi(w * z) + 2*pi)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1093
  using assms Arg2pi_diff [of "w*z" z] Arg2pi_le_div_sum_eq [of z "w*z"]
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1094
  apply (auto simp: Arg2pi_ge_0 Arg2pi_divide not_le)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1095
  apply (metis Arg2pi_lt_2pi add.commute)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1096
  apply (metis (no_types) Arg2pi add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1097
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1098
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1099
lemma Arg2pi_times:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1100
  assumes "w \<noteq> 0" "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1101
    shows "Arg2pi (w * z) = (if Arg2pi w + Arg2pi z < 2*pi then Arg2pi w + Arg2pi z
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1102
                            else (Arg2pi w + Arg2pi z) - 2*pi)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1103
  using Arg2pi_add [OF assms]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1104
  by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1105
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1106
lemma Arg2pi_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg2pi (cnj z) = Arg2pi (inverse z)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1107
  apply (simp add: Arg2pi_eq_iff divide_simps complex_norm_square [symmetric] mult.commute)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1108
  by (metis of_real_power zero_less_norm_iff zero_less_power)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1109
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1110
lemma Arg2pi_cnj: "Arg2pi(cnj z) = (if z \<in> \<real> then Arg2pi z else 2*pi - Arg2pi z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1111
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1112
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1113
  then show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1114
    by (simp add: Arg2pi_cnj_eq_inverse Arg2pi_inverse)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1115
qed auto
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
  1116
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1117
lemma Arg2pi_exp: "0 \<le> Im z \<Longrightarrow> Im z < 2*pi \<Longrightarrow> Arg2pi(exp z) = Im z"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1118
  by (rule Arg2pi_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar)
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
  1119
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
  1120
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
  1121
  obtains r a::real where "z = complex_of_real r * (cos a + \<i> * sin a)" "0 \<le> r" "0 \<le> a" "a < 2*pi"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1122
  using Arg2pi cis.ctr cis_conv_exp unfolding Complex_eq is_Arg_def by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1123
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1124
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
  1125
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
  1126
  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
  1127
    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
  1128
    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
  1129
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
  1130
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1131
subsection%unimportant\<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
  1132
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1133
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
  1134
  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
  1135
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1136
lemma field_differentiable_at_tan: "cos z \<noteq> 0 \<Longrightarrow> tan field_differentiable at z"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1137
  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
  1138
  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
  1139
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1140
lemma field_differentiable_within_tan: "cos z \<noteq> 0
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1141
         \<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
  1142
  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
  1143
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1144
lemma continuous_within_tan: "cos z \<noteq> 0 \<Longrightarrow> continuous (at z within s) tan"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1145
  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
  1146
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1147
lemma continuous_on_tan [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> cos z \<noteq> 0) \<Longrightarrow> continuous_on s tan"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1148
  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
  1149
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1150
lemma holomorphic_on_tan: "(\<And>z. z \<in> s \<Longrightarrow> cos z \<noteq> 0) \<Longrightarrow> tan holomorphic_on s"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1151
  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
  1152
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1153
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1154
subsection\<open>The principal branch of the Complex logarithm\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1155
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
  1156
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
  1157
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
  1158
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1159
definition%important ln_complex :: "complex \<Rightarrow> complex"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1160
  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
  1161
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1162
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
  1163
lemma
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1164
  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
  1165
    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
  1166
      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
  1167
      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
  1168
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1169
  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
  1170
    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
  1171
    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
  1172
  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
  1173
    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
  1174
    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
  1175
  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
  1176
    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
  1177
    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
  1178
    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
  1179
    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
  1180
  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
  1181
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1182
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1183
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1184
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
  1185
  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
  1186
    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
  1187
  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
  1188
  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
  1189
  apply auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1190
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1191
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1192
subsection%unimportant\<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
  1193
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
  1194
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
  1195
  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
  1196
    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
  1197
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
  1198
  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
  1199
    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
  1200
  also have "... = of_real(ln z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1201
    using assms by (subst Ln_exp) 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
  1202
  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
  1203
    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
  1204
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
  1205
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1206
corollary%unimportant Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<real>"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1207
  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
  1208
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1209
corollary%unimportant Im_Ln_of_real [simp]: "r > 0 \<Longrightarrow> Im (ln (of_real r)) = 0"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1210
  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
  1211
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
  1212
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
  1213
  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
  1214
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1215
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
  1216
  using Ln_of_real by force
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1217
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1218
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
  1219
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
  1220
  have "ln (exp 0) = (0::complex)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1221
    by (simp add: del: exp_zero)
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
  1222
  then show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1223
    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
  1224
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
  1225
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1226
65585
a043de9ad41e Some fixes related to compactE_image
paulson <lp15@cam.ac.uk>
parents: 65583
diff changeset
  1227
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
  1228
  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
  1229
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
  1230
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
  1231
  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
  1232
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
  1233
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
  1234
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
  1235
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
  1236
  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
  1237
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1238
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
  1239
  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
  1240
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1241
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
  1242
  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
  1243
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1244
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
  1245
  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
  1246
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1247
corollary%unimportant 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
  1248
  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
  1249
    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
  1250
  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
  1251
  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
  1252
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1253
proposition%unimportant exists_complex_root:
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1254
  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
  1255
  assumes "n \<noteq> 0"  obtains w where "z = w ^ n"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1256
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1257
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1258
  then show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1259
    by (rule_tac w = "exp(Ln z / n)" in that) (simp add: assms exp_of_nat_mult [symmetric])
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1260
qed (use assms in auto)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1261
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1262
corollary%unimportant exists_complex_root_nonzero:
62843
313d3b697c9a Mostly renaming (from HOL Light to Isabelle conventions), with a couple of new results
paulson <lp15@cam.ac.uk>
parents: 62534
diff changeset
  1263
  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
  1264
  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
  1265
  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
  1266
  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
  1267
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1268
subsection%unimportant\<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
  1269
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1270
lemma
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1271
  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
  1272
    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
  1273
      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
  1274
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1275
  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
  1276
    using assms by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1277
  then have "Im (Ln z) \<noteq> pi"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1278
    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
  1279
  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
  1280
    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
  1281
  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
  1282
    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
  1283
  then show "(Ln has_field_derivative inverse(z)) (at z)"
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 67976
diff changeset
  1284
    apply (rule has_field_derivative_inverse_strong_x
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 67976
diff changeset
  1285
              [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
  1286
    using znz *
68255
009f783d1bac small clean-up of Complex_Analysis_Basics
paulson <lp15@cam.ac.uk>
parents: 67976
diff changeset
  1287
    apply (auto simp: 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
  1288
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1289
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1290
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1291
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
  1292
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
  1293
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1294
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
  1295
  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
  1296
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1297
lemma field_differentiable_within_Ln: "z \<notin> \<real>\<^sub>\<le>\<^sub>0
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1298
         \<Longrightarrow> Ln field_differentiable (at z within S)"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1299
  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
  1300
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1301
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
  1302
  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
  1303
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1304
lemma isCont_Ln' [simp]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1305
   "\<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
  1306
  by (blast intro: isCont_o2 [OF _ continuous_at_Ln])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1307
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1308
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
  1309
  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
  1310
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1311
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
  1312
  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
  1313
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1314
lemma continuous_on_Ln' [continuous_intros]:
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1315
  "continuous_on S f \<Longrightarrow> (\<And>z. z \<in> S \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> continuous_on S (\<lambda>x. Ln (f x))"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1316
  by (rule continuous_on_compose2[OF continuous_on_Ln, of "UNIV - nonpos_Reals" S f]) auto
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1317
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1318
lemma holomorphic_on_Ln [holomorphic_intros]: "(\<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
  1319
  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
  1320
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1321
lemma holomorphic_on_Ln' [holomorphic_intros]:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1322
  "(\<And>z. z \<in> A \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0) \<Longrightarrow> f holomorphic_on A \<Longrightarrow> (\<lambda>z. Ln (f z)) holomorphic_on A"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1323
  using holomorphic_on_compose_gen[OF _ holomorphic_on_Ln, of f A "- \<real>\<^sub>\<le>\<^sub>0"]
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1324
  by (auto simp: o_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68585
diff changeset
  1325
67371
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1326
lemma tendsto_Ln [tendsto_intros]:
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1327
  fixes L F
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1328
  assumes "(f \<longlongrightarrow> L) F" "L \<notin> \<real>\<^sub>\<le>\<^sub>0"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1329
  shows   "((\<lambda>x. Ln (f x)) \<longlongrightarrow> Ln L) F"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1330
proof -
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1331
  have "nhds L \<ge> filtermap f F"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1332
    using assms(1) by (simp add: filterlim_def)
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1333
  moreover have "\<forall>\<^sub>F y in nhds L. y \<in> - \<real>\<^sub>\<le>\<^sub>0"
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1334
    using eventually_nhds_in_open[of "- \<real>\<^sub>\<le>\<^sub>0" L] assms by (auto simp: open_Compl)
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1335
  ultimately have "\<forall>\<^sub>F y in filtermap f F. y \<in> - \<real>\<^sub>\<le>\<^sub>0" by (rule filter_leD)
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1336
  moreover have "continuous_on (-\<real>\<^sub>\<le>\<^sub>0) Ln" by (rule continuous_on_Ln) auto
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1337
  ultimately show ?thesis using continuous_on_tendsto_compose[of "- \<real>\<^sub>\<le>\<^sub>0" Ln f L F] assms
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1338
    by (simp add: eventually_filtermap)
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1339
qed
2d9cf74943e1 moved in some material from Euler-MacLaurin
paulson <lp15@cam.ac.uk>
parents: 67278
diff changeset
  1340
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1341
lemma divide_ln_mono:
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1342
  fixes x y::real
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1343
  assumes "3 \<le> x" "x \<le> y"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1344
  shows "x / ln x \<le> y / ln y"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1345
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
  1346
    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
  1347
  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)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1348
    using \<open>3 \<le> x\<close> by (force intro!: derivative_eq_intros simp: field_simps power_eq_if)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1349
  show "x / ln x \<le> y / ln y"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1350
    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
  1351
    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
  1352
  proof -
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1353
    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
  1354
      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
  1355
    show ?thesis
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1356
      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
  1357
  qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  1358
qed
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1359
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1360
theorem Ln_series:
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1361
  fixes z :: complex
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1362
  assumes "norm z < 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1363
  shows   "(\<lambda>n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\<lambda>n. ?f n * z^n) sums _")
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1364
proof -
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1365
  let ?F = "\<lambda>z. \<Sum>n. ?f n * z^n" and ?F' = "\<lambda>z. \<Sum>n. diffs ?f n * z^n"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1366
  have r: "conv_radius ?f = 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1367
    by (intro conv_radius_ratio_limit_nonzero[of _ 1])
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1368
       (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1369
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1370
  have "\<exists>c. \<forall>z\<in>ball 0 1. ln (1 + z) - ?F z = c"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1371
  proof (rule has_field_derivative_zero_constant)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1372
    fix z :: complex assume z': "z \<in> ball 0 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1373
    hence z: "norm z < 1" by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1374
    define t :: complex where "t = of_real (1 + norm z) / 2"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1375
    from z have t: "norm z < norm t" "norm t < 1" unfolding t_def
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1376
      by (simp_all add: field_simps norm_divide del: of_real_add)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1377
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1378
    have "Re (-z) \<le> norm (-z)" by (rule complex_Re_le_cmod)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1379
    also from z have "... < 1" by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1380
    finally have "((\<lambda>z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1381
      by (auto intro!: derivative_eq_intros simp: complex_nonpos_Reals_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1382
    moreover have "(?F has_field_derivative ?F' z) (at z)" using t r
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1383
      by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1384
    ultimately have "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z))
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1385
                       (at z within ball 0 1)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1386
      by (intro derivative_intros) (simp_all add: at_within_open[OF z'])
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1387
    also have "(\<lambda>n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1388
      by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1389
    from sums_split_initial_segment[OF this, of 1]
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1390
      have "(\<lambda>i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1391
    hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1392
    also have "inverse (1 + z) - inverse (1 + z) = 0" by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1393
    finally show "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" .
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1394
  qed simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1395
  then obtain c where c: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> ln (1 + z) - ?F z = c" by blast
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1396
  from c[of 0] have "c = 0" by (simp only: powser_zero) simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1397
  with c[of z] assms have "ln (1 + z) = ?F z" by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1398
  moreover have "summable (\<lambda>n. ?f n * z^n)" using assms r
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1399
    by (intro summable_in_conv_radius) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1400
  ultimately show ?thesis by (simp add: sums_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1401
qed
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1402
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1403
lemma Ln_series': "cmod z < 1 \<Longrightarrow> (\<lambda>n. - ((-z)^n) / of_nat n) sums ln (1 + z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1404
  by (drule Ln_series) (simp add: power_minus')
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1405
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1406
lemma ln_series':
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1407
  assumes "abs (x::real) < 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1408
  shows   "(\<lambda>n. - ((-x)^n) / of_nat n) sums ln (1 + x)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1409
proof -
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1410
  from assms have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) sums ln (1 + complex_of_real x)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1411
    by (intro Ln_series') simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1412
  also have "(\<lambda>n. - ((-of_real x)^n) / of_nat n) = (\<lambda>n. complex_of_real (- ((-x)^n) / of_nat n))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1413
    by (rule ext) simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1414
  also from assms have "ln (1 + complex_of_real x) = of_real (ln (1 + x))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1415
    by (subst Ln_of_real [symmetric]) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1416
  finally show ?thesis by (subst (asm) sums_of_real_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1417
qed
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1418
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1419
lemma Ln_approx_linear:
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1420
  fixes z :: complex
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1421
  assumes "norm z < 1"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1422
  shows   "norm (ln (1 + z) - z) \<le> norm z^2 / (1 - norm z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1423
proof -
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1424
  let ?f = "\<lambda>n. (-1)^Suc n / of_nat n"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1425
  from assms have "(\<lambda>n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1426
  moreover have "(\<lambda>n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1427
  ultimately have "(\<lambda>n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1428
    by (subst left_diff_distrib, intro sums_diff) simp_all
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1429
  from sums_split_initial_segment[OF this, of "Suc 1"]
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1430
    have "(\<lambda>i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1431
    by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1432
  hence "(Ln (1 + z) - z) = (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1433
    by (simp add: sums_iff)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1434
  also have A: "summable (\<lambda>n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1435
    by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]])
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1436
       (auto simp: assms field_simps intro!: always_eventually)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1437
  hence "norm (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \<le>
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1438
             (\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1439
    by (intro summable_norm)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1440
       (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1441
  also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \<le> norm ((-z)^2 * (-z)^i) * 1" for i
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1442
    by (intro mult_left_mono) (simp_all add: divide_simps)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1443
  hence "(\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1444
         \<le> (\<Sum>i. norm (-(z^2) * (-z)^i))"
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1445
    using A assms
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1446
    apply (simp_all only: norm_power norm_inverse norm_divide norm_mult)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1447
    apply (intro suminf_le summable_mult summable_geometric)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1448
    apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1449
    done
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1450
  also have "... = norm z^2 * (\<Sum>i. norm z^i)" using assms
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1451
    by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1452
  also have "(\<Sum>i. norm z^i) = inverse (1 - norm z)" using assms
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1453
    by (subst suminf_geometric) (simp_all add: divide_inverse)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1454
  also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse)
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1455
  finally show ?thesis .
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1456
qed
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1457
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1458
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1459
subsection%unimportant\<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
  1460
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1461
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
  1462
  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
  1463
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1464
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1465
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
  1466
  assumes "z \<noteq> 0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1467
    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
  1468
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1469
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1470
    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
  1471
    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
  1472
      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
  1473
      by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1474
    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
  1475
      using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1476
      apply (auto simp: Re_exp zero_less_mult_iff cos_gt_zero_pi abs_if split: if_split_asm)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1477
      apply (metis cos_minus cos_pi_half divide_minus_left less_irrefl linorder_neqE_linordered_idom nonzero_mult_div_cancel_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
  1478
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1479
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1480
  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
  1481
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1482
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1483
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1484
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
  1485
  assumes "z \<noteq> 0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1486
    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
  1487
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1488
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1489
    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
  1490
    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
  1491
      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
  1492
      by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1493
    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
  1494
      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
  1495
      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
  1496
      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
  1497
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1498
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1499
  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
  1500
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1501
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1502
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1503
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
  1504
  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
  1505
    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
  1506
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1507
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1508
    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
  1509
    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
  1510
      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
  1511
      by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1512
    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
  1513
      using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"]
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1514
      apply (simp add: Im_exp zero_less_mult_iff)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1515
      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
  1516
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1517
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1518
  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
  1519
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1520
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1521
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1522
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
  1523
  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
  1524
    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
  1525
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1526
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1527
    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
  1528
    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
  1529
      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
  1530
      by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1531
    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
  1532
      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
  1533
      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
  1534
      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
  1535
      done }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1536
  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
  1537
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1538
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1539
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1540
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
  1541
  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
  1542
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1543
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
  1544
  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
  1545
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1546
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
  1547
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
  1548
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
  1549
          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
  1550
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1551
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
  1552
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
  1553
    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
  1554
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1555
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1556
subsection%unimportant\<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
  1557
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1558
lemma cnj_Ln: assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" shows "cnj(Ln z) = Ln(cnj z)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1559
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1560
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1561
  show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1562
  proof (rule exp_complex_eqI)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1563
    have "\<bar>Im (cnj (Ln z)) - Im (Ln (cnj z))\<bar> \<le> \<bar>Im (cnj (Ln z))\<bar> + \<bar>Im (Ln (cnj z))\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1564
      by (rule abs_triangle_ineq4)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1565
    also have "... < pi + pi"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1566
    proof -
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1567
      have "\<bar>Im (cnj (Ln z))\<bar> < pi"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1568
        by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1569
      moreover have "\<bar>Im (Ln (cnj z))\<bar> \<le> pi"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1570
        by (meson abs_le_iff complex_cnj_zero_iff less_eq_real_def minus_less_iff  False Im_Ln_le_pi mpi_less_Im_Ln)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1571
      ultimately show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1572
        by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1573
    qed
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1574
    finally show "\<bar>Im (cnj (Ln z)) - Im (Ln (cnj z))\<bar> < 2 * pi"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1575
      by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1576
    show "exp (cnj (Ln z)) = exp (Ln (cnj z))"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1577
      by (metis False complex_cnj_zero_iff exp_Ln exp_cnj)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1578
  qed
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1579
qed (use assms in auto)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1580
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1581
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1582
lemma Ln_inverse: assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0" shows "Ln(inverse z) = -(Ln z)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1583
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1584
  case False
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1585
  show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1586
  proof (rule exp_complex_eqI)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1587
    have "\<bar>Im (Ln (inverse z)) - Im (- Ln z)\<bar> \<le> \<bar>Im (Ln (inverse z))\<bar> + \<bar>Im (- Ln z)\<bar>"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1588
      by (rule abs_triangle_ineq4)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1589
    also have "... < pi + pi"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1590
    proof -
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1591
      have "\<bar>Im (Ln (inverse z))\<bar> < pi"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1592
        by (simp add: False Im_Ln_less_pi abs_if assms minus_less_iff mpi_less_Im_Ln)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1593
      moreover have "\<bar>Im (- Ln z)\<bar> \<le> pi"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1594
        using False Im_Ln_le_pi mpi_less_Im_Ln by fastforce
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1595
      ultimately show ?thesis
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1596
        by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1597
    qed
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1598
    finally show "\<bar>Im (Ln (inverse z)) - Im (- Ln z)\<bar> < 2 * pi"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1599
      by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1600
    show "exp (Ln (inverse z)) = exp (- Ln z)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1601
      by (simp add: False exp_minus)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  1602
  qed
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1603
qed (use assms in auto)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1604
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1605
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
  1606
  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
  1607
  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
  1608
  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
  1609
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1610
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1611
lemma Ln_ii [simp]: "Ln \<i> = \<i> * of_real pi/2"
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1612
  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
  1613
  unfolding exp_Euler
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1614
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1615
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1616
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
  1617
proof -
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1618
  have  "Ln(-\<i>) = Ln(inverse \<i>)"    by simp
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1619
  also have "... = - (Ln \<i>)"         using Ln_inverse by blast
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1620
  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
  1621
  finally show ?thesis .
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1622
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1623
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1624
lemma Ln_times:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1625
  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
  1626
    shows "Ln(w * z) =
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1627
           (if Im(Ln w + Ln z) \<le> -pi then (Ln(w) + Ln(z)) + \<i> * of_real(2*pi)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1628
            else if Im(Ln w + Ln z) > pi then (Ln(w) + Ln(z)) - \<i> * of_real(2*pi)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1629
            else Ln(w) + Ln(z))"
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1630
  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
  1631
  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
  1632
  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
  1633
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1634
corollary%unimportant 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
  1635
    "\<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
  1636
         \<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
  1637
  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
  1638
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1639
corollary%unimportant Ln_times_of_real:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1640
    "\<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
  1641
  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
  1642
  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
  1643
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1644
corollary%unimportant Ln_times_Reals:
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1645
    "\<lbrakk>r \<in> Reals; Re r > 0; z \<noteq> 0\<rbrakk> \<Longrightarrow> Ln(r * z) = ln (Re r) + Ln(z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1646
  using Ln_Reals_eq Ln_times_of_real by fastforce
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1647
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1648
corollary%unimportant Ln_divide_of_real:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1649
    "\<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
  1650
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
  1651
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
  1652
         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
  1653
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1654
corollary%unimportant Ln_prod:
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1655
  fixes f :: "'a \<Rightarrow> complex"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1656
  assumes "finite A" "\<And>x. x \<in> A \<Longrightarrow> f x \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1657
  shows "\<exists>n. Ln (prod f A) = (\<Sum>x \<in> A. Ln (f x) + (of_int (n x) * (2 * pi)) * \<i>)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1658
  using assms
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1659
proof (induction A)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1660
  case (insert x A)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1661
  then obtain n where n: "Ln (prod f A) = (\<Sum>x\<in>A. Ln (f x) + of_real (of_int (n x) * (2 * pi)) * \<i>)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1662
    by auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1663
  define D where "D \<equiv> Im (Ln (f x)) + Im (Ln (prod f A))"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1664
  define q::int where "q \<equiv> (if D \<le> -pi then 1 else if D > pi then -1 else 0)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1665
  have "prod f A \<noteq> 0" "f x \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1666
    by (auto simp: insert.hyps insert.prems)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1667
  with insert.hyps pi_ge_zero show ?case
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1668
    by (rule_tac x="n(x:=q)" in exI) (force simp: Ln_times q_def D_def n intro!: sum.cong)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1669
qed auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1670
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1671
lemma Ln_minus:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1672
  assumes "z \<noteq> 0"
69508
2a4c8a2a3f8e tuned headers; ~ -> \<not>
nipkow
parents: 69180
diff changeset
  1673
    shows "Ln(-z) = (if Im(z) \<le> 0 \<and> \<not>(Re(z) < 0 \<and> Im(z) = 0)
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1674
                     then Ln(z) + \<i> * pi
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1675
                     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
  1676
  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
  1677
        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
  1678
    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
  1679
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1680
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
  1681
  assumes "z \<noteq> 0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1682
    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
  1683
proof (cases "z \<in> \<real>\<^sub>\<le>\<^sub>0")
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1684
  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
  1685
    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
  1686
next
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1687
  case True
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1688
  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
  1689
    using assms
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1690
    apply (auto simp: complex_nonpos_Reals_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1691
    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
  1692
  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
  1693
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1694
  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
  1695
    using assms z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1696
    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
  1697
    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
  1698
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1699
  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
  1700
    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
  1701
    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
  1702
  also have "... = - (Ln z) + \<i> * 2 * complex_of_real pi"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  1703
    by (subst Ln_minus) (use assms z in auto)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  1704
  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
  1705
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1706
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1707
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
  1708
  assumes "z \<noteq> 0"
63589
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1709
    shows  "Ln(\<i> * z) = (if 0 \<le> Re(z) | Im(z) < 0
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1710
                          then Ln(z) + \<i> * of_real pi/2
58aab4745e85 more symbols;
wenzelm
parents: 63556
diff changeset
  1711
                          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
  1712
  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
  1713
        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
  1714
  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
  1715
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65585
diff changeset
  1716
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
  1717
  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
  1718
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
  1719
lemma Ln_of_nat_over_of_nat:
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1720
  assumes "m > 0" "n > 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1721
  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
  1722
proof -
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1723
  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
  1724
  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
  1725
    by (simp add: Ln_of_real[symmetric])
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1726
  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
  1727
    by (simp add: ln_div)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1728
  finally show ?thesis .
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1729
qed
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1730
67278
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1731
lemma Ln_measurable [measurable]: "Ln \<in> measurable borel borel"
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1732
proof -
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1733
  have *: "Ln (-of_real x) = of_real (ln x) + \<i> * pi" if "x > 0" for x
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1734
    using that by (subst Ln_minus) (auto simp: Ln_of_real)
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1735
  have **: "Ln (of_real x) = of_real (ln (-x)) + \<i> * pi" if "x < 0" for x
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1736
    using *[of "-x"] that by simp
67976
75b94eb58c3d Analysis builds using set_borel_measurable_def, etc.
paulson <lp15@cam.ac.uk>
parents: 67968
diff changeset
  1737
  have cont: "(\<lambda>x. indicat_real (- \<real>\<^sub>\<le>\<^sub>0) x *\<^sub>R Ln x) \<in> borel_measurable borel"
67278
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1738
    by (intro borel_measurable_continuous_on_indicator continuous_intros) auto
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1739
  have "(\<lambda>x. if x \<in> \<real>\<^sub>\<le>\<^sub>0 then ln (-Re x) + \<i> * pi else indicator (-\<real>\<^sub>\<le>\<^sub>0) x *\<^sub>R Ln x) \<in> borel \<rightarrow>\<^sub>M borel"
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1740
    (is "?f \<in> _") by (rule measurable_If_set[OF _ cont]) auto
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1741
  hence "(\<lambda>x. if x = 0 then Ln 0 else ?f x) \<in> borel \<rightarrow>\<^sub>M borel" by measurable
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1742
  also have "(\<lambda>x. if x = 0 then Ln 0 else ?f x) = Ln"
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1743
    by (auto simp: fun_eq_iff ** nonpos_Reals_def)
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1744
  finally show ?thesis .
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1745
qed
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1746
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1747
lemma powr_complex_measurable [measurable]:
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1748
  assumes [measurable]: "f \<in> measurable M borel" "g \<in> measurable M borel"
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1749
  shows   "(\<lambda>x. f x powr g x :: complex) \<in> measurable M borel"
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1750
  using assms by (simp add: powr_def)
c60e3d615b8c Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
eberlm <eberlm@in.tum.de>
parents: 67268
diff changeset
  1751
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1752
subsection\<open>The Argument of a Complex Number\<close>
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1753
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1754
text\<open>Finally: it's is defined for the same interval as the complex logarithm: $(-\pi,\pi]$.\<close>
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1755
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1756
definition%important Arg :: "complex \<Rightarrow> real" where "Arg z \<equiv> (if z = 0 then 0 else Im (Ln z))"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1757
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1758
lemma Arg_of_real: "Arg (of_real r) = (if r<0 then pi else 0)"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1759
  by (simp add: Im_Ln_eq_pi Arg_def)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1760
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1761
lemma mpi_less_Arg: "-pi < Arg z"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1762
    and Arg_le_pi: "Arg z \<le> pi"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1763
  by (auto simp: Arg_def mpi_less_Im_Ln Im_Ln_le_pi)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1764
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1765
lemma
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1766
  assumes "z \<noteq> 0"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1767
  shows Arg_eq: "z = of_real(norm z) * exp(\<i> * Arg z)"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1768
  using assms exp_Ln exp_eq_polar
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1769
  by (auto simp:  Arg_def)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1770
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1771
lemma is_Arg_Arg: "z \<noteq> 0 \<Longrightarrow> is_Arg z (Arg z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1772
  by (simp add: Arg_eq is_Arg_def)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1773
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1774
lemma Argument_exists:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1775
  assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1776
  obtains s where "is_Arg z s" "s\<in>R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1777
proof -
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1778
  let ?rp = "r - Arg z + pi"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1779
  define k where "k \<equiv> \<lfloor>?rp / (2 * pi)\<rfloor>"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1780
  have "(Arg z + of_int k * (2 * pi)) \<in> R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1781
    using floor_divide_lower [of "2*pi" ?rp] floor_divide_upper [of "2*pi" ?rp]
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1782
    by (auto simp: k_def algebra_simps R)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1783
  then show ?thesis
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1784
    using Arg_eq \<open>z \<noteq> 0\<close> is_Arg_2pi_iff is_Arg_def that by blast
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1785
qed
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1786
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1787
lemma Argument_exists_unique:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1788
  assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1789
  obtains s where "is_Arg z s" "s\<in>R" "\<And>t. \<lbrakk>is_Arg z t; t\<in>R\<rbrakk> \<Longrightarrow> s=t"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1790
proof -
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1791
  obtain s where s: "is_Arg z s" "s\<in>R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1792
    using Argument_exists [OF assms] .
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1793
  moreover have "\<And>t. \<lbrakk>is_Arg z t; t\<in>R\<rbrakk> \<Longrightarrow> s=t"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1794
    using assms s  by (auto simp: is_Arg_eqI)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1795
  ultimately show thesis
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1796
    using that by blast
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1797
qed
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1798
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1799
lemma Argument_Ex1:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1800
  assumes "z \<noteq> 0" and R: "R = {r-pi<..r+pi}"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1801
  shows "\<exists>!s. is_Arg z s \<and> s \<in> R"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1802
  using Argument_exists_unique [OF assms]  by metis
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1803
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1804
lemma Arg_divide:
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1805
  assumes "w \<noteq> 0" "z \<noteq> 0"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1806
  shows "is_Arg (z / w) (Arg z - Arg w)"
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1807
  using Arg_eq [of z] Arg_eq [of w] Arg_eq [of "norm(z / w)"] assms
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1808
  by (auto simp: is_Arg_def norm_divide field_simps exp_diff Arg_of_real)
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1809
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1810
lemma Arg_unique_lemma:
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1811
  assumes z:  "is_Arg z t"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1812
      and z': "is_Arg z t'"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1813
      and t:  "- pi < t"  "t \<le> pi"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1814
      and t': "- pi < t'" "t' \<le> pi"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1815
      and nz: "z \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1816
    shows "t' = t"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1817
  using Arg2pi_unique_lemma [of "- (inverse z)"]
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1818
proof -
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1819
  have "pi - t' = pi - t"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1820
  proof (rule Arg2pi_unique_lemma [of "- (inverse z)"])
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1821
    have "- (inverse z) = - (inverse (of_real(norm z) * exp(\<i> * t)))"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1822
      by (metis is_Arg_def z)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1823
    also have "... = (cmod (- inverse z)) * exp (\<i> * (pi - t))"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1824
      by (auto simp: field_simps exp_diff norm_divide)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1825
    finally show "is_Arg (- inverse z) (pi - t)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1826
      unfolding is_Arg_def .
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1827
    have "- (inverse z) = - (inverse (of_real(norm z) * exp(\<i> * t')))"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1828
      by (metis is_Arg_def z')
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1829
    also have "... = (cmod (- inverse z)) * exp (\<i> * (pi - t'))"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1830
      by (auto simp: field_simps exp_diff norm_divide)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1831
    finally show "is_Arg (- inverse z) (pi - t')"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1832
      unfolding is_Arg_def .
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1833
  qed (use assms in auto)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1834
  then show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1835
    by simp
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1836
qed
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1837
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1838
lemma complex_norm_eq_1_exp_eq: "norm z = 1 \<longleftrightarrow> exp(\<i> * (Arg z)) = z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1839
  by (metis Arg_eq exp_not_eq_zero exp_zero mult.left_neutral norm_zero of_real_1 norm_exp_i_times)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1840
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1841
lemma Arg_unique: "\<lbrakk>of_real r * exp(\<i> * a) = z; 0 < r; -pi < a; a \<le> pi\<rbrakk> \<Longrightarrow> Arg z = a"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1842
  by (rule Arg_unique_lemma [unfolded is_Arg_def, OF _ Arg_eq])
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1843
     (use mpi_less_Arg Arg_le_pi in \<open>auto simp: norm_mult\<close>)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1844
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1845
lemma Arg_minus:
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1846
  assumes "z \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1847
  shows "Arg (-z) = (if Arg z \<le> 0 then Arg z + pi else Arg z - pi)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1848
proof -
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1849
  have [simp]: "cmod z * cos (Arg z) = Re z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1850
    using assms Arg_eq [of z] by (metis Re_exp exp_Ln norm_exp_eq_Re Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1851
  have [simp]: "cmod z * sin (Arg z) = Im z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1852
    using assms Arg_eq [of z] by (metis Im_exp exp_Ln norm_exp_eq_Re Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1853
  show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1854
    apply (rule Arg_unique [of "norm z"])
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1855
       apply (rule complex_eqI)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1856
    using mpi_less_Arg [of z] Arg_le_pi [of z] assms
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1857
        apply (auto simp: Re_exp Im_exp cos_diff sin_diff cis_conv_exp [symmetric])
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1858
    done
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1859
qed
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1860
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1861
lemma Arg_times_of_real [simp]:
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1862
  assumes "0 < r" shows "Arg (of_real r * z) = Arg z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1863
proof (cases "z=0")
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1864
  case True
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1865
  then show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1866
    by (simp add: Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1867
next
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1868
  case False
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1869
  with Arg_eq assms show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1870
  by (auto simp: mpi_less_Arg Arg_le_pi intro!:  Arg_unique [of "r * norm z"])
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1871
qed
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1872
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1873
lemma Arg_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg (z * of_real r) = Arg z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1874
  by (metis Arg_times_of_real mult.commute)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1875
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1876
lemma Arg_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg (z / of_real r) = Arg z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1877
  by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1878
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1879
lemma Arg_less_0: "0 \<le> Arg z \<longleftrightarrow> 0 \<le> Im z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1880
  using Im_Ln_le_pi Im_Ln_pos_le
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1881
  by (simp add: Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1882
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1883
lemma Arg_eq_pi: "Arg z = pi \<longleftrightarrow> Re z < 0 \<and> Im z = 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1884
  by (auto simp: Arg_def Im_Ln_eq_pi)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1885
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1886
lemma Arg_lt_pi: "0 < Arg z \<and> Arg z < pi \<longleftrightarrow> 0 < Im z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1887
  using Arg_less_0 [of z] Im_Ln_pos_lt
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1888
  by (auto simp: order.order_iff_strict Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1889
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1890
lemma Arg_eq_0: "Arg z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1891
  using complex_is_Real_iff
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1892
  by (simp add: Arg_def Im_Ln_eq_0) (metis less_eq_real_def of_real_Re of_real_def scale_zero_left)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1893
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1894
corollary%unimportant Arg_ne_0: assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0" shows "Arg z \<noteq> 0"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1895
  using assms by (auto simp: nonneg_Reals_def Arg_eq_0)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1896
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1897
lemma Arg_eq_pi_iff: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1898
proof (cases "z=0")
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1899
  case False
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1900
  then show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1901
    using Arg_eq_0 [of "-z"] Arg_eq_pi complex_is_Real_iff by blast
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1902
qed (simp add: Arg_def)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1903
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1904
lemma Arg_eq_0_pi: "Arg z = 0 \<or> Arg z = pi \<longleftrightarrow> z \<in> \<real>"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1905
  using Arg_eq_pi_iff Arg_eq_0 by force
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1906
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1907
lemma Arg_real: "z \<in> \<real> \<Longrightarrow> Arg z = (if 0 \<le> Re z then 0 else pi)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1908
  using Arg_eq_0 Arg_eq_0_pi by auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1909
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1910
lemma Arg_inverse: "Arg(inverse z) = (if z \<in> \<real> then Arg z else - Arg z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1911
proof (cases "z \<in> \<real>")
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1912
  case True
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1913
  then show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1914
    by simp (metis Arg2pi_inverse Arg2pi_real Arg_real Reals_inverse)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1915
next
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1916
  case False
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1917
  then have "Arg z < pi" "z \<noteq> 0"
68527
2f4e2aab190a Generalising and renaming some basic results
paulson <lp15@cam.ac.uk>
parents: 68517
diff changeset
  1918
    using Arg_eq_0_pi Arg_le_pi by (auto simp: less_eq_real_def)
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1919
  then show ?thesis
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1920
    apply (simp add: False)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1921
    apply (rule Arg_unique [of "inverse (norm z)"])
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1922
    using False mpi_less_Arg [of z] Arg_eq [of z]
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1923
    apply (auto simp: exp_minus field_simps)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1924
    done
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1925
qed
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1926
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1927
lemma Arg_eq_iff:
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1928
  assumes "w \<noteq> 0" "z \<noteq> 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1929
     shows "Arg w = Arg z \<longleftrightarrow> (\<exists>x. 0 < x \<and> w = of_real x * z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1930
  using assms Arg_eq [of z] Arg_eq [of w]
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1931
  apply auto
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1932
  apply (rule_tac x="norm w / norm z" in exI)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1933
  apply (simp add: divide_simps)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1934
  by (metis mult.commute mult.left_commute)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1935
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1936
lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \<longleftrightarrow> Arg z = 0"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1937
  by (metis Arg_eq_0 Arg_inverse inverse_inverse_eq)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1938
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1939
lemma Arg_cnj_eq_inverse: "z\<noteq>0 \<Longrightarrow> Arg (cnj z) = Arg (inverse z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1940
  apply (simp add: Arg_eq_iff divide_simps complex_norm_square [symmetric] mult.commute)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1941
  by (metis of_real_power zero_less_norm_iff zero_less_power)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1942
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1943
lemma Arg_cnj: "Arg(cnj z) = (if z \<in> \<real> then Arg z else - Arg z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1944
  by (metis Arg_cnj_eq_inverse Arg_inverse Reals_0 complex_cnj_zero)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1945
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1946
lemma Arg_exp: "-pi < Im z \<Longrightarrow> Im z \<le> pi \<Longrightarrow> Arg(exp z) = Im z"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1947
  by (rule Arg_unique [of "exp(Re z)"]) (auto simp: exp_eq_polar)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1948
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1949
lemma Ln_Arg: "z\<noteq>0 \<Longrightarrow> Ln(z) = ln(norm z) + \<i> * Arg(z)"
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1950
  by (metis Arg_def Re_Ln complex_eq)
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1951
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1952
lemma continuous_at_Arg:
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1953
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1954
    shows "continuous (at z) Arg"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1955
proof -
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1956
  have [simp]: "(\<lambda>z. Im (Ln z)) \<midarrow>z\<rightarrow> Arg z"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1957
    using Arg_def assms continuous_at by fastforce
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1958
  show ?thesis
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1959
    unfolding continuous_at
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1960
  proof (rule Lim_transform_within_open)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1961
    show "\<And>w. \<lbrakk>w \<in> - \<real>\<^sub>\<le>\<^sub>0; w \<noteq> z\<rbrakk> \<Longrightarrow> Im (Ln w) = Arg w"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1962
      by (metis Arg_def Compl_iff nonpos_Reals_zero_I)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1963
  qed (use assms in auto)
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1964
qed
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1965
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1966
lemma continuous_within_Arg: "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> continuous (at z within S) Arg"
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  1967
  using continuous_at_Arg continuous_at_imp_continuous_within by blast
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  1968
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1969
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1970
subsection\<open>The Unwinding Number and the Ln product Formula\<close>
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1971
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1972
text\<open>Note that in this special case the unwinding number is -1, 0 or 1. But it's always an integer.\<close>
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1973
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1974
lemma is_Arg_exp_Im: "is_Arg (exp z) (Im z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1975
  using exp_eq_polar is_Arg_def norm_exp_eq_Re by auto
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1976
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1977
lemma is_Arg_exp_diff_2pi:
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1978
  assumes "is_Arg (exp z) \<theta>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1979
  shows "\<exists>k. Im z - of_int k * (2 * pi) = \<theta>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1980
proof (intro exI is_Arg_eqI)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1981
  let ?k = "\<lfloor>(Im z - \<theta>) / (2 * pi)\<rfloor>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1982
  show "is_Arg (exp z) (Im z - real_of_int ?k * (2 * pi))"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1983
    by (metis diff_add_cancel is_Arg_2pi_iff is_Arg_exp_Im)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1984
  show "\<bar>Im z - real_of_int ?k * (2 * pi) - \<theta>\<bar> < 2 * pi"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1985
    using floor_divide_upper [of "2*pi" "Im z - \<theta>"] floor_divide_lower [of "2*pi" "Im z - \<theta>"]
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1986
    by (auto simp: algebra_simps abs_if)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1987
qed (auto simp: is_Arg_exp_Im assms)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1988
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1989
lemma Arg_exp_diff_2pi: "\<exists>k. Im z - of_int k * (2 * pi) = Arg (exp z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1990
  using is_Arg_exp_diff_2pi [OF is_Arg_Arg] by auto
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1991
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1992
lemma unwinding_in_Ints: "(z - Ln(exp z)) / (of_real(2*pi) * \<i>) \<in> \<int>"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1993
  using Arg_exp_diff_2pi [of z]
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1994
  by (force simp: Ints_def image_def field_simps Arg_def intro!: complex_eqI)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1995
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  1996
definition%important unwinding :: "complex \<Rightarrow> int" where
68535
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1997
   "unwinding z \<equiv> THE k. of_int k = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1998
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  1999
lemma unwinding: "of_int (unwinding z) = (z - Ln(exp z)) / (of_real(2*pi) * \<i>)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2000
  using unwinding_in_Ints [of z]
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2001
  unfolding unwinding_def Ints_def by force
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2002
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2003
lemma unwinding_2pi: "(2*pi) * \<i> * unwinding(z) = z - Ln(exp z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2004
  by (simp add: unwinding)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2005
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2006
lemma Ln_times_unwinding:
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2007
    "w \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * \<i> * unwinding(Ln w + Ln z)"
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2008
  using unwinding_2pi by (simp add: exp_add)
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2009
4d09df93d1a2 The unwinding number is an integer.
paulson <lp15@cam.ac.uk>
parents: 68527
diff changeset
  2010
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2011
subsection%unimportant\<open>Relation between Ln and Arg2pi, and hence continuity of Arg2pi\<close>
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2012
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2013
lemma Arg2pi_Ln:
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2014
  assumes "0 < Arg2pi z" shows "Arg2pi z = Im(Ln(-z)) + pi"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2015
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
  2016
  case True
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2017
  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
  2018
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2019
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2020
  case False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2021
  then have "z / of_real(norm z) = exp(\<i> * of_real(Arg2pi z))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2022
    using Arg2pi [of z]
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2023
    by (metis is_Arg_def abs_norm_cancel nonzero_mult_div_cancel_left norm_of_real zero_less_norm_iff)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2024
  then have "- z / of_real(norm z) = exp (\<i> * (of_real (Arg2pi z) - pi))"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2025
    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
  2026
    by (auto simp: exp_diff algebra_simps)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2027
  then have "ln (- z / of_real(norm z)) = ln (exp (\<i> * (of_real (Arg2pi z) - pi)))"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2028
    by simp
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2029
  also have "... = \<i> * (of_real(Arg2pi z) - pi)"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2030
    using Arg2pi [of z] assms pi_not_less_zero
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2031
    by auto
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2032
  finally have "Arg2pi z =  Im (Ln (- z / of_real (cmod z))) + pi"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2033
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2034
  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
  2035
    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
  2036
  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
  2037
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2038
  finally show ?thesis .
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2039
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2040
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2041
lemma continuous_at_Arg2pi:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2042
  assumes "z \<notin> \<real>\<^sub>\<ge>\<^sub>0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2043
    shows "continuous (at z) Arg2pi"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2044
proof -
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2045
  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
  2046
    by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2047
  have [simp]: "Im x \<noteq> 0 \<Longrightarrow> Im (Ln (- x)) + pi = Arg2pi x" for x
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2048
    using Arg2pi_Ln  by (simp add: Arg2pi_gt_0 complex_nonneg_Reals_iff)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2049
  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
  2050
    using complex_nonneg_Reals_iff not_le by blast
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2051
  then have [simp]: "(\<lambda>z. Im (Ln (- z)) + pi) \<midarrow>z\<rightarrow> Arg2pi z"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  2052
    using "*" by (simp add: Arg2pi_Ln Arg2pi_gt_0 assms continuous_within) 
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2053
  show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2054
    unfolding continuous_at
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2055
  proof (rule Lim_transform_within_open)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2056
    show "\<And>x. \<lbrakk>x \<in> - \<real>\<^sub>\<ge>\<^sub>0; x \<noteq> z\<rbrakk> \<Longrightarrow> Im (Ln (- x)) + pi = Arg2pi x"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2057
      by (auto simp add: Arg2pi_Ln [OF Arg2pi_gt_0] complex_nonneg_Reals_iff)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2058
  qed (use assms in auto)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2059
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2060
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2061
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2062
text\<open>Relation between Arg2pi and arctangent in upper halfplane\<close>
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2063
lemma Arg2pi_arctan_upperhalf:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2064
  assumes "0 < Im z"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2065
    shows "Arg2pi z = pi/2 - arctan(Re z / Im z)"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2066
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
  2067
  case False
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2068
  show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2069
  proof (rule Arg2pi_unique [of "norm z"])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2070
    show "(cmod z) * exp (\<i> * (pi / 2 - arctan (Re z / Im z))) = z"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2071
      apply (auto simp: exp_Euler cos_diff sin_diff)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2072
      using assms norm_complex_def [of z, symmetric]
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2073
      apply (simp add: sin_of_real cos_of_real sin_arctan cos_arctan field_simps real_sqrt_divide)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2074
      apply (metis complex_eq)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2075
      done
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2076
  qed (use False arctan [of "Re z / Im z"] in auto)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2077
qed (use assms in auto)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2078
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2079
lemma Arg2pi_eq_Im_Ln:
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  2080
  assumes "0 \<le> Im z" "0 < Re z"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2081
    shows "Arg2pi z = Im (Ln z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2082
proof (cases "Im z = 0")
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2083
  case True then show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2084
    using Arg2pi_eq_0 Ln_in_Reals assms(2) complex_is_Real_iff by auto
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2085
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
  2086
  case False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2087
  then have *: "Arg2pi z > 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2088
    using Arg2pi_gt_0 complex_is_Real_iff by blast
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2089
  then have "z \<noteq> 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2090
    by auto
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2091
  with * assms False show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2092
    by (subst Arg2pi_Ln) (auto simp: Ln_minus)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2093
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2094
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2095
lemma continuous_within_upperhalf_Arg2pi:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2096
  assumes "z \<noteq> 0"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2097
    shows "continuous (at z within {z. 0 \<le> Im z}) Arg2pi"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2098
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
  2099
  case False then show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2100
    using continuous_at_Arg2pi continuous_at_imp_continuous_within by auto
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2101
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2102
  case True
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2103
  then have z: "z \<in> \<real>" "0 < Re z"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2104
    using assms  by (auto simp: complex_nonneg_Reals_iff complex_is_Real_iff complex_neq_0)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2105
  then have [simp]: "Arg2pi z = 0" "Im (Ln z) = 0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2106
    by (auto simp: Arg2pi_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  2107
  show ?thesis
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2108
  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
  2109
    fix e::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2110
    assume "0 < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2111
    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
  2112
      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
  2113
    ultimately
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2114
    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
  2115
      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
  2116
    { fix x
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2117
      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
  2118
      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
  2119
        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
  2120
      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
  2121
        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
  2122
    }
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2123
    then show "\<exists>d>0. \<forall>x. 0 \<le> Im x \<longrightarrow> x \<noteq> z \<and> cmod (x - z) < d \<longrightarrow> \<bar>Arg2pi x\<bar> < e"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2124
      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
  2125
      using z d
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2126
      apply (auto simp: Arg2pi_eq_Im_Ln)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2127
      done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2128
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2129
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2130
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2131
lemma continuous_on_upperhalf_Arg2pi: "continuous_on ({z. 0 \<le> Im z} - {0}) Arg2pi"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2132
  unfolding continuous_on_eq_continuous_within
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2133
  by (metis DiffE Diff_subset continuous_within_subset continuous_within_upperhalf_Arg2pi insertCI)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2134
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2135
lemma open_Arg2pi2pi_less_Int:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2136
  assumes "0 \<le> s" "t \<le> 2*pi"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2137
    shows "open ({y. s < Arg2pi y} \<inter> {y. Arg2pi y < t})"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2138
proof -
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2139
  have 1: "continuous_on (UNIV - \<real>\<^sub>\<ge>\<^sub>0) Arg2pi"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2140
    using continuous_at_Arg2pi continuous_at_imp_continuous_within
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2141
    by (auto simp: continuous_on_eq_continuous_within)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2142
  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
  2143
  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
  2144
    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
  2145
    by (metis greaterThan_def lessThan_def open_Int)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2146
  moreover have "{y. s < Arg2pi y} \<inter> {y. Arg2pi y < t} \<subseteq> - \<real>\<^sub>\<ge>\<^sub>0"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2147
    using assms by (auto simp: Arg2pi_real complex_nonneg_Reals_iff complex_is_Real_iff)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2148
  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
  2149
    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
  2150
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2151
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2152
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2153
lemma open_Arg2pi2pi_gt: "open {z. t < Arg2pi z}"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2154
proof (cases "t < 0")
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2155
  case True then have "{z. t < Arg2pi z} = UNIV"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2156
    using Arg2pi_ge_0 less_le_trans by auto
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2157
  then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2158
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2159
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2160
  case False then show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2161
    using open_Arg2pi2pi_less_Int [of t "2*pi"] Arg2pi_lt_2pi
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2162
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2163
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2164
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2165
lemma closed_Arg2pi2pi_le: "closed {z. Arg2pi z \<le> t}"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2166
  using open_Arg2pi2pi_gt [of t]
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2167
  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
  2168
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2169
subsection%unimportant\<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
  2170
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
  2171
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
  2172
  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
  2173
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
  2174
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
  2175
  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
  2176
  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
  2177
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
  2178
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
  2179
  apply (simp add: powr_def)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2180
  using Im_Ln_eq_0 complex_is_Real_iff norm_complex_def  by auto
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2181
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
  2182
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
  2183
  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
  2184
  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
  2185
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2186
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
  2187
  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
  2188
  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
  2189
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2190
lemma cnj_powr:
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2191
  assumes "Im a = 0 \<Longrightarrow> Re a \<ge> 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2192
  shows   "cnj (a powr b) = cnj a powr cnj b"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2193
proof (cases "a = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2194
  case False
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2195
  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
  2196
  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
  2197
qed simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2198
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
  2199
lemma powr_real_real:
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2200
  assumes "w \<in> \<real>" "z \<in> \<real>" "0 < Re w"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2201
  shows "w powr z = exp(Re z * ln(Re w))"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2202
proof -
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2203
  have "w \<noteq> 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2204
    using assms by auto
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2205
  with assms show ?thesis
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2206
    by (simp add: powr_def Ln_Reals_eq of_real_exp)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2207
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
  2208
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
  2209
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
  2210
  fixes x::real and y::real
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2211
  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
  2212
  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
  2213
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2214
lemma powr_of_int:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2215
  fixes z::complex and n::int
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2216
  assumes "z\<noteq>(0::complex)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2217
  shows "z powr of_int n = (if n\<ge>0 then z^nat n else inverse (z^nat (-n)))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2218
  by (metis assms not_le of_int_of_nat powr_complexpow powr_minus)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2219
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2220
lemma powr_Reals_eq: "\<lbrakk>x \<in> \<real>; y \<in> \<real>; Re x \<ge> 0\<rbrakk> \<Longrightarrow> x powr y = of_real (Re x powr Re y)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2221
  by (metis of_real_Re powr_of_real)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2222
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
  2223
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
  2224
    "\<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
  2225
     \<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
  2226
  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
  2227
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
  2228
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
  2229
    "\<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
  2230
           \<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
  2231
  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
  2232
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2233
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
  2234
  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
  2235
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2236
lemma
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2237
  fixes w::complex
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2238
  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
  2239
  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
  2240
  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
  2241
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2242
lemma powr_neg_real_complex:
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2243
  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
  2244
proof (cases "x = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2245
  assume x: "x \<noteq> 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2246
  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
  2247
  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
  2248
    by (simp add: Ln_minus Ln_of_real)
63092
a949b2a5f51d eliminated use of empty "assms";
wenzelm
parents: 63040
diff changeset
  2249
  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
  2250
    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
  2251
  also note cis_pi
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2252
  finally show ?thesis by simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2253
qed simp_all
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2254
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
  2255
lemma has_field_derivative_powr:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2256
  fixes z :: complex
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2257
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2258
  shows "((\<lambda>z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2259
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2260
  case False
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2261
  show ?thesis
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2262
    unfolding powr_def
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2263
  proof (rule DERIV_transform_at)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2264
    show "((\<lambda>z. exp (s * Ln z)) has_field_derivative s * (if z = 0 then 0 else exp ((s - 1) * Ln z)))
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2265
           (at z)"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2266
      apply (intro derivative_eq_intros | simp add: assms)+
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2267
      by (simp add: False divide_complex_def exp_diff left_diff_distrib')
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2268
  qed (use False in auto)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2269
qed (use assms in auto)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2270
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2271
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
  2272
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2273
lemma has_field_derivative_powr_of_int:
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2274
  fixes z :: complex
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2275
  assumes gderiv:"(g has_field_derivative gd) (at z within s)" and "g z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2276
  shows "((\<lambda>z. g z powr of_int n) has_field_derivative (n * g z powr (of_int n - 1) * gd)) (at z within s)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2277
proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2278
  define dd where "dd = of_int n * g z powr (of_int (n - 1)) * gd"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2279
  obtain e where "e>0" and e_dist:"\<forall>y\<in>s. dist z y < e \<longrightarrow> g y \<noteq> 0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2280
    using DERIV_continuous[OF gderiv,THEN continuous_within_avoid] \<open>g z\<noteq>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2281
  have ?thesis when "n\<ge>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2282
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2283
    define dd' where "dd' = of_int n * g z ^ (nat n - 1) * gd"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2284
    have "dd=dd'"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2285
    proof (cases "n=0")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2286
      case False
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2287
      then have "n-1 \<ge>0" using \<open>n\<ge>0\<close> by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2288
      then have "g z powr (of_int (n - 1)) = g z ^ (nat n - 1)"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2289
        using powr_of_int[OF \<open>g z\<noteq>0\<close>,of "n-1"] by (simp add: nat_diff_distrib')
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2290
      then show ?thesis unfolding dd_def dd'_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2291
    qed (simp add:dd_def dd'_def)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2292
    then have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2293
                \<longleftrightarrow> ((\<lambda>z. g z powr of_int n) has_field_derivative dd') (at z within s)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2294
      by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2295
    also have "... \<longleftrightarrow> ((\<lambda>z. g z ^ nat n) has_field_derivative dd') (at z within s)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2296
    proof (rule has_field_derivative_cong_eventually)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2297
      show "\<forall>\<^sub>F x in at z within s. g x powr of_int n = g x ^ nat n"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2298
        unfolding eventually_at
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2299
        apply (rule exI[where x=e])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2300
        using powr_of_int that \<open>e>0\<close> e_dist by (simp add: dist_commute)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2301
    qed (use powr_of_int \<open>g z\<noteq>0\<close> that in simp)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2302
    also have "..." unfolding dd'_def using gderiv that
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2303
      by (auto intro!: derivative_eq_intros)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2304
    finally have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)" .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2305
    then show ?thesis unfolding dd_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2306
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2307
  moreover have ?thesis when "n<0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2308
  proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2309
    define dd' where "dd' = of_int n / g z ^ (nat (1 - n)) * gd"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2310
    have "dd=dd'"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2311
    proof -
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2312
      have "g z powr of_int (n - 1) = inverse (g z ^ nat (1-n))"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2313
        using powr_of_int[OF \<open>g z\<noteq>0\<close>,of "n-1"] that by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2314
      then show ?thesis
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2315
        unfolding dd_def dd'_def by (simp add: divide_inverse)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2316
    qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2317
    then have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2318
                \<longleftrightarrow> ((\<lambda>z. g z powr of_int n) has_field_derivative dd') (at z within s)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2319
      by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2320
    also have "... \<longleftrightarrow> ((\<lambda>z. inverse (g z ^ nat (-n))) has_field_derivative dd') (at z within s)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2321
    proof (rule has_field_derivative_cong_eventually)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2322
      show "\<forall>\<^sub>F x in at z within s. g x powr of_int n = inverse (g x ^ nat (- n))"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2323
         unfolding eventually_at
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2324
        apply (rule exI[where x=e])
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2325
        using powr_of_int that \<open>e>0\<close> e_dist by (simp add: dist_commute)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2326
    qed (use powr_of_int \<open>g z\<noteq>0\<close> that in simp)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2327
    also have "..."
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2328
    proof -
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2329
      have "nat (- n) + nat (1 - n) - Suc 0 = nat (- n) + nat (- n)"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2330
        by auto
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2331
      then show ?thesis
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2332
        unfolding dd'_def using gderiv that \<open>g z\<noteq>0\<close>
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2333
        by (auto intro!: derivative_eq_intros simp add:divide_simps power_add[symmetric])
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2334
    qed
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2335
    finally have "((\<lambda>z. g z powr of_int n) has_field_derivative dd) (at z within s)" .
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2336
    then show ?thesis unfolding dd_def by simp
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2337
  qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2338
  ultimately show ?thesis by force
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2339
qed
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2340
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2341
lemma field_differentiable_powr_of_int:
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2342
  fixes z :: complex
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2343
  assumes gderiv:"g field_differentiable (at z within s)" and "g z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2344
  shows "(\<lambda>z. g z powr of_int n) field_differentiable (at z within s)"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2345
using has_field_derivative_powr_of_int assms(2) field_differentiable_def gderiv by blast
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2346
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2347
lemma holomorphic_on_powr_of_int [holomorphic_intros]:
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2348
  assumes "f holomorphic_on s" "\<forall>z\<in>s. f z\<noteq>0"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2349
  shows "(\<lambda>z. (f z) powr of_int n) holomorphic_on s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2350
proof (cases "n\<ge>0")
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2351
  case True
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2352
  then have "?thesis \<longleftrightarrow> (\<lambda>z. (f z) ^ nat n) holomorphic_on s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2353
    apply (rule_tac holomorphic_cong)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2354
    using assms(2) by (auto simp add:powr_of_int)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2355
  moreover have "(\<lambda>z. (f z) ^ nat n) holomorphic_on s"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2356
    using assms(1) by (auto intro:holomorphic_intros)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2357
  ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2358
next
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2359
  case False
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2360
  then have "?thesis \<longleftrightarrow> (\<lambda>z. inverse (f z) ^ nat (-n)) holomorphic_on s"
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2361
    apply (rule_tac holomorphic_cong)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2362
    using assms(2) by (auto simp add:powr_of_int power_inverse)
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2363
  moreover have "(\<lambda>z. inverse (f z) ^ nat (-n)) holomorphic_on s"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2364
    using assms by (auto intro!:holomorphic_intros)
67706
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2365
  ultimately show ?thesis by auto
4ddc49205f5d Unified the order of zeros and poles; improved reasoning around non-essential singularites
Wenda Li <wl302@cam.ac.uk>
parents: 67578
diff changeset
  2366
qed
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  2367
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
  2368
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
  2369
    "w \<noteq> 0 \<Longrightarrow> ((\<lambda>z. w powr z) has_field_derivative Ln w * w powr z) (at z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2370
  unfolding powr_def by (intro derivative_eq_intros | simp)+
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2371
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
  2372
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
  2373
  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
  2374
  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
  2375
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
  2376
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
  2377
lemma holomorphic_on_powr_right [holomorphic_intros]:
67268
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2378
  assumes "f holomorphic_on s"
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2379
  shows "(\<lambda>z. w powr (f z)) holomorphic_on s"
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2380
proof (cases "w = 0")
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2381
  case False
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2382
  with assms show ?thesis
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2383
    unfolding holomorphic_on_def field_differentiable_def
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2384
    by (metis (full_types) DERIV_chain' has_field_derivative_powr_right)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2385
qed simp
67268
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2386
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2387
lemma holomorphic_on_divide_gen [holomorphic_intros]:
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2388
  assumes f: "f holomorphic_on s" and g: "g holomorphic_on s" and 0: "\<And>z z'. \<lbrakk>z \<in> s; z' \<in> s\<rbrakk> \<Longrightarrow> g z = 0 \<longleftrightarrow> g z' = 0"
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2389
shows "(\<lambda>z. f z / g z) holomorphic_on s"
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2390
proof (cases "\<exists>z\<in>s. g z = 0")
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2391
  case True
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2392
  with 0 have "g z = 0" if "z \<in> s" for z
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2393
    using that by blast
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2394
  then show ?thesis
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2395
    using g holomorphic_transform by auto
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2396
next
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2397
  case False
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2398
  with 0 have "g z \<noteq> 0" if "z \<in> s" for z
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2399
    using that by blast
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2400
  with holomorphic_on_divide show ?thesis
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2401
    using f g by blast
bdf25939a550 new/improved theories involving convergence; better pretty-printing for bounded quantifiers and sum/product
paulson <lp15@cam.ac.uk>
parents: 67135
diff changeset
  2402
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
  2403
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
  2404
lemma norm_powr_real_powr:
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2405
  "w \<in> \<real> \<Longrightarrow> 0 \<le> Re w \<Longrightarrow> cmod (w powr z) = Re w powr Re z"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2406
  by (metis dual_order.order_iff_strict norm_powr_real norm_zero of_real_0 of_real_Re powr_def)
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2407
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2408
lemma tendsto_powr_complex:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2409
  fixes f g :: "_ \<Rightarrow> complex"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2410
  assumes a: "a \<notin> \<real>\<^sub>\<le>\<^sub>0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2411
  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
  2412
  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
  2413
proof -
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2414
  from a have [simp]: "a \<noteq> 0" by auto
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2415
  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
  2416
    by (auto intro!: tendsto_intros simp: powr_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2417
  also {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2418
    have "eventually (\<lambda>z. z \<noteq> 0) (nhds a)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2419
      by (intro t1_space_nhds) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2420
    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
  2421
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2422
  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
  2423
    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
  2424
  finally show ?thesis .
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2425
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2426
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2427
lemma tendsto_powr_complex_0:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2428
  fixes f g :: "'a \<Rightarrow> complex"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2429
  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
  2430
  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
  2431
proof (rule tendsto_norm_zero_cancel)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2432
  define h where
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2433
    "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
  2434
  {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2435
    fix z :: 'a assume z: "f z \<noteq> 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2436
    define c where "c = abs (Im (g z)) * pi"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2437
    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
  2438
      have "abs (Im (Ln (f z))) \<le> pi" by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2439
    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
  2440
      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
  2441
    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
  2442
    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
  2443
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2444
  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
  2445
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2446
  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
  2447
    by (rule tendsto_mono[OF _ g]) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2448
  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
  2449
    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
  2450
  moreover {
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2451
    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
  2452
      by (auto simp: filterlim_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2453
    hence "filterlim (\<lambda>x. norm (f x)) (principal {0<..})
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2454
             (inf F (principal {z. f z \<noteq> 0}))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2455
      by (rule filterlim_mono) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2456
  }
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2457
  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
  2458
    by (simp add: filterlim_inf at_within_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2459
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2460
  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
  2461
    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
  2462
          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
  2463
  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
  2464
          -\<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
  2465
    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
  2466
  have C: "(h \<longlongrightarrow> 0) F" unfolding h_def
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2467
    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
  2468
       (insert B, auto simp: filterlim_uminus_at_bot algebra_simps)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2469
  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
  2470
    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
  2471
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2472
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2473
lemma tendsto_powr_complex' [tendsto_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2474
  fixes f g :: "_ \<Rightarrow> complex"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2475
  assumes "a \<notin> \<real>\<^sub>\<le>\<^sub>0 \<or> (a = 0 \<and> Re b > 0)" and "(f \<longlongrightarrow> a) F" "(g \<longlongrightarrow> b) F"
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2476
  shows   "((\<lambda>z. f z powr g z) \<longlongrightarrow> a powr b) F"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2477
  using assms tendsto_powr_complex tendsto_powr_complex_0 by fastforce
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2478
67135
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2479
lemma tendsto_neg_powr_complex_of_real:
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2480
  assumes "filterlim f at_top F" and "Re s < 0"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2481
  shows   "((\<lambda>x. complex_of_real (f x) powr s) \<longlongrightarrow> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2482
proof -
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2483
  have "((\<lambda>x. norm (complex_of_real (f x) powr s)) \<longlongrightarrow> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2484
  proof (rule Lim_transform_eventually)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2485
    from assms(1) have "eventually (\<lambda>x. f x \<ge> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2486
      by (auto simp: filterlim_at_top)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2487
    thus "eventually (\<lambda>x. f x powr Re s = norm (of_real (f x) powr s)) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2488
      by eventually_elim (simp add: norm_powr_real_powr)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2489
    from assms show "((\<lambda>x. f x powr Re s) \<longlongrightarrow> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2490
      by (intro tendsto_neg_powr)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2491
  qed
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2492
  thus ?thesis by (simp add: tendsto_norm_zero_iff)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2493
qed
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2494
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2495
lemma tendsto_neg_powr_complex_of_nat:
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2496
  assumes "filterlim f at_top F" and "Re s < 0"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2497
  shows   "((\<lambda>x. of_nat (f x) powr s) \<longlongrightarrow> 0) F"
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2498
proof -
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2499
  have "((\<lambda>x. of_real (real (f x)) powr s) \<longlongrightarrow> 0) F" using assms(2)
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2500
    by (intro filterlim_compose[OF _ tendsto_neg_powr_complex_of_real]
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2501
              filterlim_compose[OF _ assms(1)] filterlim_real_sequentially filterlim_ident) auto
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2502
  thus ?thesis by simp
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2503
qed
1a94352812f4 Moved material from AFP to Analysis/Number_Theory
Manuel Eberl <eberlm@in.tum.de>
parents: 66827
diff changeset
  2504
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2505
lemma continuous_powr_complex:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2506
  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
  2507
  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
  2508
  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
  2509
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2510
lemma isCont_powr_complex [continuous_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2511
  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
  2512
  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
  2513
  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
  2514
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2515
lemma continuous_on_powr_complex [continuous_intros]:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2516
  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
  2517
  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
  2518
  assumes "continuous_on A f" "continuous_on A g"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2519
  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
  2520
  unfolding continuous_on_def
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2521
proof
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2522
  fix z assume z: "z \<in> A"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2523
  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
  2524
  proof (cases "f z = 0")
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2525
    case False
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2526
    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
  2527
    with assms(3,4) z show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2528
      by (intro tendsto_powr_complex')
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2529
         (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
  2530
  next
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2531
    case True
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2532
    with assms z show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2533
      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
  2534
  qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63092
diff changeset
  2535
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
  2536
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2537
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2538
subsection%unimportant\<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
  2539
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2540
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
  2541
  fixes s::complex
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2542
  assumes "0 < Re s"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2543
    shows "(\<lambda>n. Ln (of_nat n) / of_nat n powr s) \<longlonglongrightarrow> 0"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2544
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
  2545
  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
  2546
  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
  2547
  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
  2548
  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
  2549
    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
  2550
      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
  2551
  next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2552
    fix x::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2553
    assume x: "2 / (e * (Re s)\<^sup>2) \<le> x"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2554
    have "2 / (e * (Re s)\<^sup>2) > 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2555
      using e assms by simp
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2556
    with x have "x > 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2557
      by linarith
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2558
    then have "x * 2 \<le> e * (x\<^sup>2 * (Re s)\<^sup>2)"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2559
      using e assms x by (auto simp: power2_eq_square field_simps)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2560
    also have "... < e * (2 + (x * (Re s * 2) + x\<^sup>2 * (Re s)\<^sup>2))"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2561
      using e assms \<open>x > 0\<close>
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2562
      by (auto simp: power2_eq_square field_simps add_pos_pos)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2563
    finally show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2564
      by (auto simp: algebra_simps)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2565
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2566
  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
  2567
    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
  2568
  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
  2569
    using assms
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2570
    by (force intro: less_le_trans [OF _ exp_lower_taylor_quadratic])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2571
  then obtain xo where "xo > 0" and xo: "\<And>x. x \<ge> xo \<Longrightarrow> x < e * exp (Re s * x)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2572
    using e  by (auto simp: field_simps)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2573
  have "norm (Ln (of_nat n) / of_nat n powr s) < e" if "n \<ge> nat \<lceil>exp xo\<rceil>" for n
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2574
    using e xo [of "ln n"] that
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2575
    apply (auto simp: norm_divide norm_powr_real divide_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2576
    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
  2577
    done
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2578
  then show "\<exists>no. \<forall>n\<ge>no. norm (Ln (of_nat n) / of_nat n powr s) < e"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2579
    by blast
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2580
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2581
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2582
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
  2583
  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
  2584
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2585
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
  2586
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2587
  assumes "0 < s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2588
    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
  2589
  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
  2590
  apply (subst filterlim_sequentially_Suc [symmetric])
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2591
  apply (simp add: lim_sequentially dist_norm Ln_Reals_eq norm_powr_real_powr norm_divide)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2592
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2593
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2594
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
  2595
  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
  2596
  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
  2597
  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
  2598
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2599
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2600
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
  2601
  assumes "0 < Re s"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2602
  shows "(\<lambda>n. 1 / of_nat n powr s) \<longlonglongrightarrow> 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2603
proof (rule Lim_null_comparison)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2604
  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
  2605
    using ln_272_gt_1
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2606
    by (force intro: order_trans [of _ "ln (272/100)"])
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2607
  then show "\<forall>\<^sub>F x in sequentially. cmod (1 / of_nat x powr s) \<le> cmod (Ln (of_nat x) / of_nat x powr s)"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2608
    by (auto simp: norm_divide divide_simps eventually_sequentially)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2609
  show "(\<lambda>n. cmod (Ln (of_nat n) / of_nat n powr s)) \<longlonglongrightarrow> 0"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2610
    using lim_Ln_over_power [OF assms] by (metis tendsto_norm_zero_iff)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2611
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2612
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2613
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
  2614
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2615
  assumes "0 < s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2616
    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
  2617
  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
  2618
  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
  2619
  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
  2620
  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
  2621
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2622
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2623
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
  2624
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
  2625
  fix r::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2626
  assume "0 < r"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2627
  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
  2628
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2629
  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
  2630
    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
  2631
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2632
  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
  2633
    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
  2634
  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
  2635
    by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2636
  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
  2637
    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
  2638
  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
  2639
    using neq0_conv by fastforce
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2640
  ultimately show "\<exists>no. \<forall>k. Ln (of_nat k) \<noteq> 0 \<longrightarrow> no \<le> k \<longrightarrow> 1 < r * cmod (Ln (of_nat k))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2641
    using n \<open>0 < r\<close>
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2642
    by (rule_tac x=n in exI) (force simp: divide_simps intro: less_le_trans)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2643
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2644
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  2645
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
  2646
  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
  2647
  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
  2648
  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
  2649
  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
  2650
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  2651
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2652
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
  2653
proof (rule Lim_transform_eventually)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2654
  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
  2655
  proof (rule Lim_transform_bound)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2656
    show "(inverse o real) \<longlonglongrightarrow> 0"
66827
c94531b5007d Divided Topology_Euclidean_Space in two, creating new theory Connected. Also deleted some duplicate / variant theorems
paulson <lp15@cam.ac.uk>
parents: 66793
diff changeset
  2657
      by (metis comp_def lim_inverse_n tendsto_explicit)
65719
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2658
    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
  2659
    proof
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2660
      fix n::nat
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2661
      assume n: "3 \<le> n"
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2662
      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
  2663
        by auto
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2664
      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
  2665
        by (simp add: divide_simps)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2666
      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
  2667
        by (simp add: ln_add_one_self_le_self)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2668
      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
  2669
        by (intro mult_mono) (use n in auto)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2670
      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
  2671
        by (simp add: field_simps ln0)
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2672
      qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2673
  qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2674
  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
  2675
    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
  2676
  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
  2677
    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
  2678
qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2679
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2680
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
  2681
proof -
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2682
  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
  2683
    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
  2684
  then show ?thesis
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2685
    by simp
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2686
qed
7c57d79d61b7 A few more new lemmas
paulson <lp15@cam.ac.uk>
parents: 65587
diff changeset
  2687
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
  2688
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2689
subsection%unimportant\<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
  2690
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2691
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
  2692
  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
  2693
    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
  2694
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2695
  have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))"
64240
eabf80376aab more standardized names
haftmann
parents: 63918
diff changeset
  2696
    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
  2697
  also have "... = z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2698
    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
  2699
  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
  2700
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2701
  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
  2702
    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
  2703
    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
  2704
    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
  2705
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2706
  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
  2707
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2708
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2709
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2710
lemma csqrt_inverse:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2711
  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
  2712
    shows "csqrt (inverse z) = inverse (csqrt z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2713
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2714
  case False
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2715
  then show ?thesis
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2716
    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
  2717
    by (simp add: csqrt_exp_Ln Ln_inverse exp_minus)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2718
qed auto
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2719
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2720
lemma cnj_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2721
  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
  2722
    shows "cnj(csqrt z) = csqrt(cnj z)"
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2723
proof (cases "z=0")
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2724
  case False
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2725
  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
  2726
     by (simp add: assms cnj_Ln csqrt_exp_Ln exp_cnj)
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2727
qed auto
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2728
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2729
lemma has_field_derivative_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2730
  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
  2731
    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
  2732
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2733
  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
  2734
    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
  2735
  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
  2736
    by (simp add: divide_simps)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2737
  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
  2738
    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
  2739
  have "Im z = 0 \<Longrightarrow> 0 < Re z"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2740
    using assms complex_nonpos_Reals_iff not_less by blast
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2741
  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
  2742
    by (force intro: derivative_eq_intros * simp add: assms)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2743
  then show ?thesis
68257
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2744
  proof (rule DERIV_transform_at)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2745
    show "\<And>x. dist x z < cmod z \<Longrightarrow> exp (Ln x / 2) = csqrt x"
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2746
      by (metis csqrt_exp_Ln dist_0_norm less_irrefl)
e6e131577536 small tidy-up of Complex_Transcendental
paulson <lp15@cam.ac.uk>
parents: 68255
diff changeset
  2747
  qed (use z in auto)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2748
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2749
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2750
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
  2751
    "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
  2752
  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
  2753
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2754
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
  2755
    "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
  2756
  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
  2757
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2758
lemma continuous_at_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2759
    "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
  2760
  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
  2761
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2762
corollary%unimportant isCont_csqrt' [simp]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2763
   "\<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
  2764
  by (blast intro: isCont_o2 [OF _ continuous_at_csqrt])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  2765
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2766
lemma continuous_within_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2767
    "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
  2768
  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
  2769
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2770
lemma continuous_on_csqrt [continuous_intros]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2771
    "(\<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
  2772
  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
  2773
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2774
lemma holomorphic_on_csqrt:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2775
    "(\<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
  2776
  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
  2777
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2778
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
  2779
    "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
  2780
  using open_Compl
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2781
  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
  2782
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2783
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
  2784
    "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
  2785
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
  2786
  case True
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2787
  have *: "\<And>e. \<lbrakk>0 < e\<rbrakk>
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2788
         \<Longrightarrow> \<forall>x'\<in>\<real> \<inter> {w. 0 \<le> Re w}. cmod x' < e^2 \<longrightarrow> cmod (csqrt x') < e"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2789
    by (auto simp: Reals_def real_less_lsqrt)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2790
  have "Im z = 0" "Re z < 0 \<or> z = 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2791
    using True cnj.code complex_cnj_zero_iff  by (auto simp: Complex_eq complex_nonpos_Reals_iff) fastforce
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2792
  with * show ?thesis
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2793
    apply (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge])
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2794
    apply (auto simp: continuous_within_eps_delta)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2795
    using zero_less_power by blast
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2796
next
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2797
  case False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2798
    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
  2799
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  2800
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2801
subsection\<open>Complex arctangent\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2802
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2803
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
  2804
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2805
definition%important Arctan :: "complex \<Rightarrow> complex" where
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2806
    "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
  2807
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2808
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
  2809
  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
  2810
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2811
lemma Ln_conv_Arctan:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2812
  assumes "z \<noteq> -1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2813
  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
  2814
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2815
  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
  2816
             \<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
  2817
    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
  2818
  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
  2819
  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
  2820
  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
  2821
    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
  2822
  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
  2823
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2824
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2825
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
  2826
  by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2827
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2828
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
  2829
  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
  2830
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2831
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
  2832
  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
  2833
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2834
lemma tan_Arctan:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2835
  assumes "z\<^sup>2 \<noteq> -1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2836
    shows [simp]:"tan(Arctan z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2837
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2838
  have "1 + \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2839
    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
  2840
  moreover
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2841
  have "1 - \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2842
    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
  2843
  ultimately
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2844
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2845
    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
  2846
                  divide_simps power2_eq_square [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2847
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2848
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2849
lemma Arctan_tan [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2850
  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
  2851
    shows "Arctan(tan z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2852
proof -
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2853
  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
  2854
    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
  2855
  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
  2856
    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
  2857
  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
  2858
    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
  2859
  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
  2860
    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
  2861
  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
  2862
    by (simp add: exp_eq_1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2863
  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
  2864
    by (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2865
  also have "... \<longleftrightarrow> False"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2866
    using assms ge_pi2
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2867
    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
  2868
    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
  2869
  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
  2870
    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
  2871
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2872
    using assms *
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2873
    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
  2874
                     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
  2875
    apply (rule Ln_unique)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2876
    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
  2877
    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
  2878
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2879
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2880
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2881
lemma
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2882
  assumes "Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2883
  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
  2884
    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
  2885
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2886
  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
  2887
    using assms
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  2888
    by (metis abs_one add_diff_cancel_left' complex_i_mult_minus diff_0 i_squared imaginary_unit.simps
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2889
                less_asym neg_equal_iff_equal)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2890
  have "z \<noteq> -\<i>" using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2891
    by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2892
  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
  2893
    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
  2894
  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
  2895
    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
  2896
  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
  2897
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2898
    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
  2899
              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
  2900
  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
  2901
    using nz1 nz2 by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2902
  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
  2903
    apply (simp add: divide_complex_def)
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  2904
    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
  2905
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2906
    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
  2907
    done
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2908
  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
  2909
    by (auto simp add: complex_nonpos_Reals_iff)
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2910
  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
  2911
    unfolding Arctan_def divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2912
    using mpi_less_Im_Ln [OF nzi]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  2913
    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
  2914
    done
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2915
  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
  2916
    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
  2917
    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
  2918
    using nz0 nz1 zz
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2919
    apply (simp add: algebra_simps divide_simps power2_eq_square)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2920
    apply algebra
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2921
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2922
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2923
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2924
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
  2925
  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
  2926
  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
  2927
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  2928
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
  2929
    "(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
  2930
  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
  2931
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2932
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
  2933
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
  2934
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2935
lemma continuous_at_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2936
    "(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
  2937
  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
  2938
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2939
lemma continuous_within_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2940
    "(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
  2941
  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
  2942
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2943
lemma continuous_on_Arctan [continuous_intros]:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2944
    "(\<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
  2945
  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
  2946
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2947
lemma holomorphic_on_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2948
    "(\<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
  2949
  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
  2950
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  2951
theorem Arctan_series:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2952
  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
  2953
  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
  2954
  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
  2955
  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
  2956
  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
  2957
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2958
  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
  2959
  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
  2960
  proof (cases "u = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2961
    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
  2962
    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
  2963
              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
  2964
    proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2965
      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
  2966
      have "ereal (norm (h u n) / norm (h u (Suc n))) =
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2967
             ereal (inverse (norm u)^2) * ereal (((2*Suc n+1) / (Suc n)) /
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  2968
                 ((2*Suc n-1) / (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
  2969
      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
  2970
                    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
  2971
      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
  2972
        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
  2973
      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
  2974
        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
  2975
      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
  2976
              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
  2977
    qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2978
    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
  2979
      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
  2980
    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
  2981
      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
  2982
    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
  2983
      by (simp add: divide_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2984
    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
  2985
    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
  2986
      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
  2987
         (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
  2988
               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
  2989
    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
  2990
      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
  2991
         (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
  2992
  qed (simp add: h_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2993
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  2994
  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
  2995
  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
  2996
    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
  2997
    hence u: "norm u < 1" by (simp add: dist_0_norm)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  2998
    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
  2999
    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
  3000
    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
  3001
    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
  3002
      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
  3003
    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
  3004
      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
  3005
    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
  3006
      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
  3007
         (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
  3008
    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
  3009
    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
  3010
      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
  3011
    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
  3012
    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
  3013
      show "((\<lambda>u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3014
      by (simp_all add: at_within_open[OF _ open_ball])
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3015
  qed simp_all
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3016
  then obtain c where c: "\<And>u. norm u < 1 \<Longrightarrow> Arctan u - G u = c" by auto
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3017
  from this[of 0] have "c = 0" by (simp add: G_def g_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3018
  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
  3019
  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
  3020
  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
  3021
                              (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
  3022
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3023
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3024
text \<open>A quickly-converging series for the logarithm, based on the arctangent.\<close>
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3025
theorem ln_series_quadratic:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3026
  assumes x: "x > (0::real)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents: 61973
diff changeset
  3027
  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
  3028
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62843
diff changeset
  3029
  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
  3030
  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
  3031
  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
  3032
  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
  3033
    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
  3034
  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
  3035
  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
  3036
    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
  3037
  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
  3038
                 (\<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
  3039
    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
  3040
  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
  3041
    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
  3042
  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
  3043
    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
  3044
  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
  3045
    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
  3046
  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
  3047
    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
  3048
  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
  3049
  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
  3050
  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
  3051
  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
  3052
qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3053
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3054
subsection%unimportant \<open>Real arctangent\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3055
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3056
lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3057
proof -
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3058
  have ne: "1 + x\<^sup>2 \<noteq> 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3059
    by (metis power_one sum_power2_eq_zero_iff zero_neq_one)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3060
  have "Re (Ln ((1 - \<i> * x) * inverse (1 + \<i> * x))) = 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3061
    apply (rule norm_exp_imaginary)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3062
    apply (subst exp_Ln)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3063
    using ne apply (simp_all add: cmod_def complex_eq_iff)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3064
    apply (auto simp: divide_simps)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3065
    apply algebra
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3066
    done
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3067
  then show ?thesis
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3068
    unfolding Arctan_def divide_complex_def by (simp add: complex_eq_iff)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3069
qed
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3070
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3071
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
  3072
proof (rule arctan_unique)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3073
  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
  3074
    apply (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3075
    apply (rule Im_Ln_less_pi)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3076
    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
  3077
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3078
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3079
  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
  3080
    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
  3081
  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
  3082
    using mpi_less_Im_Ln [OF *]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3083
    by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3084
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3085
  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
  3086
    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
  3087
    apply (simp add: field_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3088
    by (simp add: power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3089
  also have "... = x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3090
    apply (subst tan_Arctan, auto)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3091
    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
  3092
  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
  3093
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3094
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3095
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
  3096
  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
  3097
  by (simp add: complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3098
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3099
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
  3100
  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
  3101
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3102
declare arctan_one [simp]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3103
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3104
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
  3105
  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
  3106
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3107
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
  3108
  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
  3109
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3110
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
  3111
  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
  3112
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3113
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
  3114
  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
  3115
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3116
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
  3117
  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
  3118
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3119
lemma arctan_add_raw:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3120
  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
  3121
    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
  3122
proof (rule arctan_unique [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3123
  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
  3124
    using assms by linarith+
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3125
  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
  3126
    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
  3127
    by (simp add: arctan tan_add)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3128
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3129
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3130
lemma arctan_inverse:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3131
  assumes "0 < x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3132
    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
  3133
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3134
  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
  3135
    by (simp add: arctan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3136
  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
  3137
    by (simp add: tan_cot)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3138
  also have "... = pi/2 - arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3139
  proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3140
    have "0 < pi - arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3141
    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
  3142
    with assms show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3143
      by (simp add: Transcendental.arctan_tan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3144
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3145
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3146
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3147
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3148
lemma arctan_add_small:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3149
  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
  3150
    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
  3151
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
  3152
  case True then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3153
    by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3154
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3155
  case False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3156
  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
  3157
    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
  3158
    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
  3159
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3160
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3161
    apply (rule arctan_add_raw)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3162
    using * by linarith
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3163
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3164
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3165
lemma abs_arctan_le:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3166
  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
  3167
proof -
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3168
  have 1: "\<And>x. x \<in> \<real> \<Longrightarrow> cmod (inverse (1 + x\<^sup>2)) \<le> 1"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3169
    by (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3170
  have "cmod (Arctan w - Arctan z) \<le> 1 * cmod (w-z)" if "w \<in> \<real>" "z \<in> \<real>" for w z
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3171
    apply (rule field_differentiable_bound [OF convex_Reals, of Arctan _ 1])
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3172
       apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan])
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3173
    using 1 that apply (auto simp: Reals_def)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3174
    done
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3175
  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
  3176
    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
  3177
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3178
    by (simp add: Arctan_of_real)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3179
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3180
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3181
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
  3182
  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
  3183
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3184
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
  3185
  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
  3186
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3187
lemma arctan_bounds:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3188
  assumes "0 \<le> x" "x < 1"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3189
  shows arctan_lower_bound:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3190
    "(\<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
  3191
    (is "(\<Sum>k<_. (- 1)^ k * ?a k) \<le> _")
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3192
    and arctan_upper_bound:
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3193
    "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
  3194
proof -
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3195
  have tendsto_zero: "?a \<longlonglongrightarrow> 0"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3196
  proof (rule tendsto_eq_rhs)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3197
    show "(\<lambda>k. 1 / real (k * 2 + 1) * x ^ (k * 2 + 1)) \<longlonglongrightarrow> 0 * 0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3198
      using assms
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3199
      by (intro tendsto_mult real_tendsto_divide_at_top)
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3200
        (auto simp: filterlim_real_sequentially filterlim_sequentially_iff_filterlim_real
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3201
          intro!: real_tendsto_divide_at_top tendsto_power_zero filterlim_real_sequentially
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3202
          tendsto_eq_intros filterlim_at_top_mult_tendsto_pos filterlim_tendsto_add_at_top)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3203
  qed simp
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3204
  have nonneg: "0 \<le> ?a n" for n
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3205
    by (force intro!: divide_nonneg_nonneg mult_nonneg_nonneg zero_le_power assms)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3206
  have le: "?a (Suc n) \<le> ?a n" for n
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3207
    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
  3208
  from summable_Leibniz'(4)[of ?a, OF tendsto_zero nonneg le, of n]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3209
    summable_Leibniz'(2)[of ?a, OF tendsto_zero nonneg le, of n]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3210
    assms
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3211
  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
  3212
    by (auto simp: arctan_series)
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3213
qed
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3214
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3215
subsection%unimportant \<open>Bounds on pi using real arctangent\<close>
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3216
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3217
lemma pi_machin: "pi = 16 * arctan (1 / 5) - 4 * arctan (1 / 239)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3218
  using machin by simp
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3219
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3220
lemma pi_approx: "3.141592653588 \<le> pi" "pi \<le> 3.1415926535899"
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3221
  unfolding pi_machin
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3222
  using arctan_bounds[of "1/5"   4]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3223
        arctan_bounds[of "1/239" 4]
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3224
  by (simp_all add: eval_nat_numeral)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  3225
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3226
lemma pi_gt3: "pi > 3"
65583
8d53b3bebab4 Further new material. The simprule status of some exp and ln identities was reverted.
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  3227
  using pi_approx by simp
63556
36e9732988ce numerical bounds on pi
immler
parents: 63492
diff changeset
  3228
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3229
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3230
subsection\<open>Inverse Sine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3231
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3232
definition%important Arcsin :: "complex \<Rightarrow> complex" where
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3233
   "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
  3234
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3235
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
  3236
  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
  3237
  apply auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3238
  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
  3239
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3240
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
  3241
  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
  3242
  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
  3243
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3244
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
  3245
  by (simp add: Arcsin_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3246
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3247
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
  3248
  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
  3249
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3250
lemma one_minus_z2_notin_nonpos_Reals:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3251
  assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3252
  shows "1 - z\<^sup>2 \<notin> \<real>\<^sub>\<le>\<^sub>0"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3253
  using assms
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3254
  apply (auto simp: complex_nonpos_Reals_iff Re_power2 Im_power2)
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3255
  using power2_less_0 [of "Im z"] apply force
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3256
  using abs_square_less_1 not_le by blast
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3257
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3258
lemma isCont_Arcsin_lemma:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3259
  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
  3260
    shows False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3261
proof (cases "Im z = 0")
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3262
  case True
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3263
  then show ?thesis
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3264
    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
  3265
next
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3266
  case False
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3267
  have leim: "(cmod (1 - z\<^sup>2) + (1 - Re (z\<^sup>2))) / 2 \<le> (Im z)\<^sup>2"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3268
    using le0 sqrt_le_D by fastforce
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3269
  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
  3270
  proof (clarsimp simp add: cmod_def)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3271
    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
  3272
    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
  3273
      by simp
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3274
    then show False using False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3275
      by (simp add: power2_eq_square algebra_simps)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3276
  qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3277
  moreover have 2: "(Im z)\<^sup>2 = (1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3278
    using leim cmod_power2 [of z] norm_triangle_ineq2 [of "z^2" 1]
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3279
    by (simp add: norm_power Re_power2 norm_minus_commute [of 1])
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3280
  ultimately show False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3281
    by (simp add: Re_power2 Im_power2 cmod_power2)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3282
qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3283
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3284
lemma isCont_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3285
  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
  3286
    shows "isCont Arcsin z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3287
proof -
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3288
  have 1: "\<i> * z + csqrt (1 - z\<^sup>2) \<notin> \<real>\<^sub>\<le>\<^sub>0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3289
    by (metis isCont_Arcsin_lemma assms complex_nonpos_Reals_iff)
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3290
  have 2: "1 - z\<^sup>2 \<notin> \<real>\<^sub>\<le>\<^sub>0"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3291
    by (simp add: one_minus_z2_notin_nonpos_Reals assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3292
  show ?thesis
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3293
    using assms unfolding Arcsin_def by (intro isCont_Ln' isCont_csqrt' continuous_intros 1 2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3294
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3295
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3296
lemma isCont_Arcsin' [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3297
  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
  3298
  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
  3299
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3300
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
  3301
proof -
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3302
  have "\<i>*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \<longleftrightarrow> (\<i>*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67371
diff changeset
  3303
    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
  3304
  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
  3305
    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
  3306
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3307
    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
  3308
    apply (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3309
    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
  3310
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3311
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3312
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3313
lemma Re_eq_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3314
    "\<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
  3315
      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
  3316
  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
  3317
  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
  3318
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3319
lemma Re_less_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3320
  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
  3321
    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
  3322
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3323
  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
  3324
    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
  3325
  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
  3326
    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
  3327
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3328
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3329
lemma Arcsin_sin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3330
    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
  3331
      shows "Arcsin(sin z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3332
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3333
  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
  3334
    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
  3335
  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
  3336
    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
  3337
  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
  3338
    apply (subst csqrt_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3339
    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
  3340
    apply auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3341
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3342
  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
  3343
    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
  3344
  also have "... = z"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3345
    using assms by (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: if_split_asm)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3346
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3347
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3348
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3349
lemma Arcsin_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3350
    "\<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
  3351
  by (metis Arcsin_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3352
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3353
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
  3354
  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
  3355
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3356
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
  3357
  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
  3358
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3359
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
  3360
  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
  3361
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3362
lemma has_field_derivative_Arcsin:
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3363
  assumes "Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3364
    shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  3365
proof -
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3366
  have "(sin (Arcsin z))\<^sup>2 \<noteq> 1"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3367
    using assms one_minus_z2_notin_nonpos_Reals by force
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3368
  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
  3369
    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
  3370
  then show ?thesis
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3371
    by (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) (auto intro: isCont_Arcsin assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3372
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3373
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3374
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
  3375
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
  3376
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3377
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
  3378
    "(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
  3379
  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
  3380
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3381
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
  3382
    "(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
  3383
  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
  3384
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3385
lemma continuous_within_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3386
    "(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
  3387
  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
  3388
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3389
lemma continuous_on_Arcsin [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3390
    "(\<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
  3391
  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
  3392
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3393
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
  3394
  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
  3395
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3396
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3397
subsection\<open>Inverse Cosine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3398
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3399
definition%important Arccos :: "complex \<Rightarrow> complex" where
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3400
   "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
  3401
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3402
lemma Arccos_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Im(z + \<i> * csqrt(1 - z\<^sup>2))"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3403
  using Arcsin_range_lemma [of "-z"]  by simp
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3404
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3405
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
  3406
  using Arcsin_body_lemma [of z]
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3407
  by (metis Arcsin_body_lemma complex_i_mult_minus diff_minus_eq_add power2_minus right_minus_eq)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3408
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3409
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
  3410
  by (simp add: Arccos_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3411
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3412
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
  3413
  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
  3414
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3415
text\<open>A very tricky argument to find!\<close>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3416
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
  3417
  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
  3418
    shows False
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3419
proof (cases "Im z = 0")
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3420
  case True
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3421
  then show ?thesis
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3422
    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
  3423
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3424
  case False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3425
  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
  3426
    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
  3427
    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
  3428
  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
  3429
  proof (clarsimp simp add: cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3430
    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
  3431
    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
  3432
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3433
    then show False using False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3434
      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
  3435
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3436
  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
  3437
    apply (subst Imz)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3438
    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
  3439
    apply (simp add: Re_power2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3440
    done
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3441
  ultimately show False
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3442
    by (simp add: cmod_power2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3443
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3444
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3445
lemma isCont_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3446
  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
  3447
    shows "isCont Arccos z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3448
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3449
  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
  3450
    by (metis complex_nonpos_Reals_iff isCont_Arccos_lemma assms)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3451
  with assms show ?thesis
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3452
    apply (simp add: Arccos_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3453
    apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62087
diff changeset
  3454
    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
  3455
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3456
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3457
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3458
lemma isCont_Arccos' [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3459
  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
  3460
  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
  3461
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3462
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
  3463
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3464
  have "z*2 + \<i> * (2 * csqrt (1 - z\<^sup>2)) = 0 \<longleftrightarrow> z*2 + \<i> * csqrt (1 - z\<^sup>2)*2 = 0"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67371
diff changeset
  3465
    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
  3466
  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
  3467
    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
  3468
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3469
    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
  3470
    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
  3471
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3472
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3473
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3474
lemma Arccos_cos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3475
    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
  3476
             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
  3477
             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
  3478
      shows "Arccos(cos z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3479
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
  3480
  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
  3481
    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
  3482
  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
  3483
    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
  3484
  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
  3485
                           \<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
  3486
    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
  3487
  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
  3488
                              \<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
  3489
    apply (subst csqrt_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3490
    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
  3491
    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
  3492
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3493
  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
  3494
    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
  3495
  also have "... = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3496
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3497
    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
  3498
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3499
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3500
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3501
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3502
lemma Arccos_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3503
    "\<lbrakk>cos z = w;
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3504
      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
  3505
      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
  3506
      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
  3507
  using Arccos_cos by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3508
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3509
lemma Arccos_0 [simp]: "Arccos 0 = pi/2"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3510
  by (rule Arccos_unique) auto
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3511
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3512
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
  3513
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3514
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3515
lemma Arccos_minus1: "Arccos(-1) = pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3516
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3517
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3518
lemma has_field_derivative_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3519
  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
  3520
    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
  3521
proof -
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3522
  have "x\<^sup>2 \<noteq> -1" for x::real
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3523
    by (sos "((R<1 + (([~1] * A=0) + (R<1 * (R<1 * [x__]^2)))))")
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3524
  with assms have "(cos (Arccos z))\<^sup>2 \<noteq> 1"
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3525
    by (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3526
  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
  3527
    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
  3528
  then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)"
68281
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3529
    by (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]])
faa4b49d1b34 more small tidying
paulson <lp15@cam.ac.uk>
parents: 68257
diff changeset
  3530
       (auto intro: isCont_Arccos assms)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3531
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3532
    by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3533
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3534
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3535
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
  3536
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
  3537
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3538
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
  3539
    "(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
  3540
  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
  3541
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  3542
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
  3543
    "(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
  3544
  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
  3545
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3546
lemma continuous_within_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3547
    "(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
  3548
  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
  3549
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3550
lemma continuous_on_Arccos [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3551
    "(\<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
  3552
  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
  3553
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3554
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
  3555
  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
  3556
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3557
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3558
subsection%unimportant\<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
  3559
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3560
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
  3561
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3562
  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
  3563
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3564
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
  3565
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3566
  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
  3567
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3568
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
  3569
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3570
  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
  3571
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3572
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
  3573
  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
  3574
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3575
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
  3576
proof -
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3577
  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
  3578
    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
  3579
    apply (simp only: abs_le_square_iff)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3580
    apply (simp add: divide_simps)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3581
    done
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3582
  also have "... \<le> (cmod w)\<^sup>2"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3583
    by (auto simp: cmod_power2)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3584
  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
  3585
    using abs_le_square_iff by force
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3586
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
  3587
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3588
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
  3589
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3590
  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
  3591
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3592
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
  3593
  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
  3594
64773
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3595
lemma norm_Arccos_bounded:
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3596
  fixes w :: complex
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3597
  shows "norm (Arccos w) \<le> pi + norm w"
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3598
proof -
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3599
  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
  3600
    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
  3601
  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
  3602
    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
  3603
  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
  3604
    apply (simp add: norm_le_square)
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3605
    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
  3606
  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
  3607
    by auto
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3608
qed
223b2ebdda79 Many new theorems, and more tidying
paulson <lp15@cam.ac.uk>
parents: 64593
diff changeset
  3609
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3610
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3611
subsection%unimportant\<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
  3612
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3613
lemma cos_Arcsin_nonzero:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3614
  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
  3615
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3616
  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
  3617
    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
  3618
  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
  3619
  proof
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3620
    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
  3621
    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
  3622
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3623
    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
  3624
      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
  3625
    then show False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3626
      using assms by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3627
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3628
  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
  3629
    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
  3630
  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
  3631
    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
  3632
  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
  3633
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3634
    apply (auto simp: power2_sum)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3635
    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
  3636
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3637
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3638
    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
  3639
    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
  3640
    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
  3641
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3642
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3643
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3644
lemma sin_Arccos_nonzero:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3645
  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
  3646
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3647
  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
  3648
    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
  3649
  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
  3650
  proof
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3651
    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
  3652
    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
  3653
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3654
    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
  3655
      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
  3656
    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
  3657
      using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3658
      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
  3659
    then show False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3660
      using assms by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3661
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3662
  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
  3663
    by (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3664
  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
  3665
    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
  3666
  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
  3667
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3668
    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
  3669
    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
  3670
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3671
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3672
    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
  3673
    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
  3674
    apply (simp add: power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3675
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3676
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3677
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3678
lemma cos_sin_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3679
  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
  3680
    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
  3681
  apply (rule csqrt_unique [THEN sym])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3682
  apply (simp add: cos_squared_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3683
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3684
  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
  3685
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3686
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3687
lemma sin_cos_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3688
  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
  3689
    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
  3690
  apply (rule csqrt_unique [THEN sym])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3691
  apply (simp add: sin_squared_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3692
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3693
  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
  3694
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3695
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3696
lemma Arcsin_Arccos_csqrt_pos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3697
    "(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
  3698
  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
  3699
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3700
lemma Arccos_Arcsin_csqrt_pos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3701
    "(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
  3702
  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
  3703
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3704
lemma sin_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3705
    "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
  3706
  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
  3707
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3708
lemma cos_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3709
    "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
  3710
  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
  3711
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3712
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3713
subsection%unimportant\<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
  3714
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3715
lemma Im_Arcsin_of_real:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3716
  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
  3717
    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
  3718
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3719
  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
  3720
    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
  3721
  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
  3722
    using assms abs_square_le_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3723
    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
  3724
  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
  3725
    by (simp add: norm_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3726
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3727
    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
  3728
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3729
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3730
corollary%unimportant Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3731
  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
  3732
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3733
lemma arcsin_eq_Re_Arcsin:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3734
  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
  3735
    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
  3736
unfolding arcsin_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3737
proof (rule the_equality, safe)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3738
  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
  3739
    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
  3740
    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
  3741
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3742
  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
  3743
    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
  3744
    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
  3745
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3746
  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
  3747
    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
  3748
    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
  3749
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3750
  fix x'
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3751
  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
  3752
  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
  3753
    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
  3754
    apply (subst Arcsin_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3755
    apply (auto simp: )
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3756
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3757
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3758
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3759
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
  3760
  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
  3761
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3762
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3763
subsection%unimportant\<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
  3764
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3765
lemma Im_Arccos_of_real:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3766
  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
  3767
    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
  3768
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3769
  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
  3770
    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
  3771
  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
  3772
    using assms abs_square_le_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3773
    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
  3774
  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
  3775
    by (simp add: norm_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3776
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3777
    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
  3778
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3779
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3780
corollary%unimportant Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3781
  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
  3782
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3783
lemma arccos_eq_Re_Arccos:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3784
  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
  3785
    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
  3786
unfolding arccos_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3787
proof (rule the_equality, safe)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3788
  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
  3789
    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
  3790
    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
  3791
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3792
  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
  3793
    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
  3794
    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
  3795
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3796
  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
  3797
    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
  3798
    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
  3799
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3800
  fix x'
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3801
  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
  3802
  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
  3803
    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
  3804
    apply (subst Arccos_cos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3805
    apply (auto simp: )
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3806
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3807
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  3808
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  3809
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
  3810
  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
  3811
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3812
subsection%unimportant\<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
  3813
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
  3814
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
  3815
  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
  3816
    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
  3817
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
  3818
  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
  3819
  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
  3820
    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
  3821
      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
  3822
      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
  3823
  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
  3824
    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
  3825
      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
  3826
      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
  3827
  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
  3828
    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
  3829
      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
  3830
      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
  3831
                    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
  3832
  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
  3833
  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
  3834
    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
  3835
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
  3836
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
  3837
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
  3838
  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
  3839
    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
  3840
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
  3841
  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
  3842
    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
  3843
    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
  3844
    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
  3845
    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
  3846
  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
  3847
    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
  3848
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
  3849
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
  3850
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
  3851
  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
  3852
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
  3853
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
  3854
  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
  3855
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
  3856
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
  3857
  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
  3858
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  3859
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
  3860
  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
  3861
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
  3862
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
  3863
  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
  3864
  apply (subst Arcsin_Arccos_csqrt_pos)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  3865
  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
  3866
  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
  3867
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
  3868
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
  3869
  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
  3870
  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
  3871
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
  3872
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
  3873
  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
  3874
  apply (subst Arccos_Arcsin_csqrt_pos)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  3875
  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
  3876
  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
  3877
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
  3878
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
  3879
  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
  3880
  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
  3881
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3882
subsection%unimportant\<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
  3883
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
  3884
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
  3885
    "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
  3886
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
  3887
  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
  3888
        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
  3889
    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
  3890
  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
  3891
    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
  3892
  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
  3893
    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
  3894
          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
  3895
    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
  3896
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
  3897
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
  3898
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
  3899
    "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
  3900
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
  3901
  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
  3902
    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
  3903
    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
  3904
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
  3905
  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
  3906
  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
  3907
    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
  3908
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
  3909
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
  3910
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
  3911
    "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
  3912
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
  3913
  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
  3914
        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
  3915
    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
  3916
  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
  3917
    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
  3918
  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
  3919
    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
  3920
          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
  3921
    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
  3922
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
  3923
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  3924
lemma continuous_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
  3925
    "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
  3926
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
  3927
  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
  3928
    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
  3929
    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
  3930
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
  3931
  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
  3932
  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
  3933
    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
  3934
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
  3935
67578
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3936
lemma sinh_ln_complex: "x \<noteq> 0 \<Longrightarrow> sinh (ln x :: complex) = (x - inverse x) / 2"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3937
  by (simp add: sinh_def exp_minus scaleR_conv_of_real exp_of_real)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3938
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3939
lemma cosh_ln_complex: "x \<noteq> 0 \<Longrightarrow> cosh (ln x :: complex) = (x + inverse x) / 2"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3940
  by (simp add: cosh_def exp_minus scaleR_conv_of_real)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3941
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3942
lemma tanh_ln_complex: "x \<noteq> 0 \<Longrightarrow> tanh (ln x :: complex) = (x ^ 2 - 1) / (x ^ 2 + 1)"
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3943
  by (simp add: tanh_def sinh_ln_complex cosh_ln_complex divide_simps power2_eq_square)
6a9a0f2bb9b4 Some lemmas about complex sinh/cosh/tanh
Manuel Eberl <eberlm@in.tum.de>
parents: 67443
diff changeset
  3944
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
  3945
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  3946
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
  3947
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  3948
theorem complex_root_unity:
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3949
  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
  3950
  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
  3951
    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
  3952
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
  3953
  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
  3954
    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
  3955
  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
  3956
    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
  3957
    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
  3958
    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
  3959
    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
  3960
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
  3961
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
  3962
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
  3963
  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
  3964
  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
  3965
    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
  3966
           \<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
  3967
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
  3968
    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
  3969
               \<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
  3970
          (\<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
  3971
              (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
  3972
      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
  3973
    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
  3974
      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
  3975
    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
  3976
      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
  3977
      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
  3978
      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
  3979
      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
  3980
    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
  3981
      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
  3982
    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
  3983
      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
  3984
    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
  3985
             \<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
  3986
   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
  3987
  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
  3988
    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
  3989
    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
  3990
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
  3991
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  3992
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
  3993
    "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
  3994
              {..<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
  3995
  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
  3996
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
  3997
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
  3998
  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
  3999
  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
  4000
    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
  4001
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
  4002
  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
  4003
    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
  4004
  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
  4005
     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
  4006
     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
  4007
  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
  4008
    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
  4009
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
  4010
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4011
lemma 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
  4012
     "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
  4013
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
  4014
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4015
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
  4016
     "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
  4017
  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
  4018
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4019
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
  4020
  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
  4021
    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
  4022
  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
  4023
  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
  4024
  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
  4025
  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
  4026
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
  4027
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
  4028
  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
  4029
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  4030
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
  4031
    "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
  4032
  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
  4033
  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
  4034
  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
  4035
69180
922833cc6839 Tagged some theories in HOL-Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 68721
diff changeset
  4036
subsection\<open>Formulation of loop homotopy in terms of maps out of type complex\<close>
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4037
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4038
lemma homotopic_circlemaps_imp_homotopic_loops:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4039
  assumes "homotopic_with (\<lambda>h. True) (sphere 0 1) S f g"
64508
874555896035 more symbols;
wenzelm
parents: 64394
diff changeset
  4040
   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
  4041
                            (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
  4042
proof -
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4043
  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
  4044
    using assms by (auto simp: sphere_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4045
  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
  4046
     by (intro continuous_intros)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4047
  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
  4048
    by (auto simp: norm_mult)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4049
  ultimately
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4050
  show ?thesis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4051
    apply (simp add: homotopic_loops_def comp_assoc)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4052
    apply (rule homotopic_with_compose_continuous_right)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4053
      apply (auto simp: pathstart_def pathfinish_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4054
    done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4055
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4056
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4057
lemma homotopic_loops_imp_homotopic_circlemaps:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4058
  assumes "homotopic_loops S p q"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4059
    shows "homotopic_with (\<lambda>h. True) (sphere 0 1) S
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4060
                          (p \<circ> (\<lambda>z. (Arg2pi z / (2 * pi))))
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4061
                          (q \<circ> (\<lambda>z. (Arg2pi z / (2 * pi))))"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4062
proof -
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4063
  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
  4064
             and him: "h ` ({0..1} \<times> {0..1}) \<subseteq> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4065
             and h0: "(\<forall>x. h (0, x) = p x)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4066
             and h1: "(\<forall>x. h (1, x) = q x)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4067
             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
  4068
    using assms
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4069
    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
  4070
  define j where "j \<equiv> \<lambda>z. if 0 \<le> Im (snd z)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4071
                          then h (fst z, Arg2pi (snd z) / (2 * pi))
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4072
                          else h (fst z, 1 - Arg2pi (cnj (snd z)) / (2 * pi))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4073
  have Arg2pi_eq: "1 - Arg2pi (cnj y) / (2 * pi) = Arg2pi y / (2 * pi) \<or> Arg2pi y = 0 \<and> Arg2pi (cnj y) = 0" if "cmod y = 1" for y
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4074
    using that Arg2pi_eq_0_pi Arg2pi_eq_pi by (force simp: Arg2pi_cnj divide_simps)
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4075
  show ?thesis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4076
  proof (simp add: homotopic_with; intro conjI ballI exI)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4077
    show "continuous_on ({0..1} \<times> sphere 0 1) (\<lambda>w. h (fst w, Arg2pi (snd w) / (2 * pi)))"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4078
    proof (rule continuous_on_eq)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4079
      show j: "j x = h (fst x, Arg2pi (snd x) / (2 * pi))" if "x \<in> {0..1} \<times> sphere 0 1" for x
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4080
        using Arg2pi_eq that h01 by (force simp: j_def)
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4081
      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
  4082
        by auto
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4083
      have c1: "continuous_on ({0..1} \<times> sphere 0 1 \<inter> UNIV \<times> {z. 0 \<le> Im z}) (\<lambda>x. h (fst x, Arg2pi (snd x) / (2 * pi)))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4084
        apply (intro continuous_intros continuous_on_compose2 [OF conth]  continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi])
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4085
            apply (auto simp: Arg2pi)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4086
        apply (meson Arg2pi_lt_2pi linear not_le)
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4087
        done
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4088
      have c2: "continuous_on ({0..1} \<times> sphere 0 1 \<inter> UNIV \<times> {z. Im z \<le> 0}) (\<lambda>x. h (fst x, 1 - Arg2pi (cnj (snd x)) / (2 * pi)))"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4089
        apply (intro continuous_intros continuous_on_compose2 [OF conth]  continuous_on_compose2 [OF continuous_on_upperhalf_Arg2pi])
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4090
            apply (auto simp: Arg2pi)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4091
        apply (meson Arg2pi_lt_2pi linear not_le)
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4092
        done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4093
      show "continuous_on ({0..1} \<times> sphere 0 1) j"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4094
        apply (simp add: j_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4095
        apply (subst eq)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4096
        apply (rule continuous_on_cases_local)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4097
            apply (simp_all add: eq [symmetric] closedin_closed_Int closed_Times closed_halfspace_Im_le closed_halfspace_Im_ge c1 c2)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4098
        using Arg2pi_eq h01
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4099
        by force
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4100
    qed
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4101
    have "(\<lambda>w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \<times> sphere 0 1) \<subseteq> h ` ({0..1} \<times> {0..1})"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4102
      by (auto simp: Arg2pi_ge_0 Arg2pi_lt_2pi less_imp_le)
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4103
    also have "... \<subseteq> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4104
      using him by blast
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4105
    finally show "(\<lambda>w. h (fst w, Arg2pi (snd w) / (2 * pi))) ` ({0..1} \<times> sphere 0 1) \<subseteq> S" .
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4106
  qed (auto simp: h0 h1)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4107
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4108
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4109
lemma simply_connected_homotopic_loops:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4110
  "simply_connected S \<longleftrightarrow>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4111
       (\<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
  4112
unfolding simply_connected_def using homotopic_loops_refl by metis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4113
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4114
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4115
lemma simply_connected_eq_homotopic_circlemaps1:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4116
  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
  4117
  assumes S: "simply_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4118
      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
  4119
      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
  4120
    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
  4121
proof -
64508
874555896035 more symbols;
wenzelm
parents: 64394
diff changeset
  4122
  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
  4123
    apply (rule S [unfolded simply_connected_homotopic_loops, rule_format])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4124
    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
  4125
    done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4126
  then show ?thesis
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4127
    apply (rule homotopic_with_eq [OF homotopic_loops_imp_homotopic_circlemaps])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4128
      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
  4129
    done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4130
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4131
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4132
lemma simply_connected_eq_homotopic_circlemaps2a:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4133
  fixes h :: "complex \<Rightarrow> 'a::topological_space"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4134
  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
  4135
      and hom: "\<And>f g::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4136
                \<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
  4137
                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
  4138
                \<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
  4139
            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
  4140
    apply (rule_tac x="h 1" in exI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4141
    apply (rule hom)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4142
    using assms
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4143
    by (auto simp: continuous_on_const)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4144
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4145
lemma simply_connected_eq_homotopic_circlemaps2b:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4146
  fixes S :: "'a::real_normed_vector set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4147
  assumes "\<And>f g::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4148
                \<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
  4149
                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
  4150
                \<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
  4151
  shows "path_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4152
proof (clarsimp simp add: path_connected_eq_homotopic_points)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4153
  fix a b
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4154
  assume "a \<in> S" "b \<in> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4155
  then show "homotopic_loops S (linepath a a) (linepath b b)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4156
    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
  4157
    by (auto simp: o_def continuous_on_const linepath_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4158
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4159
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4160
lemma simply_connected_eq_homotopic_circlemaps3:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4161
  fixes h :: "complex \<Rightarrow> 'a::real_normed_vector"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4162
  assumes "path_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4163
      and hom: "\<And>f::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4164
                  \<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
  4165
                  \<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
  4166
    shows "simply_connected S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4167
proof (clarsimp simp add: simply_connected_eq_contractible_loop_some assms)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4168
  fix p
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4169
  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
  4170
  then have "homotopic_loops S p p"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4171
    by (simp add: homotopic_loops_refl)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4172
  then obtain a where homp: "homotopic_with (\<lambda>h. True) (sphere 0 1) S (p \<circ> (\<lambda>z. Arg2pi z / (2 * pi))) (\<lambda>x. a)"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4173
    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
  4174
  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
  4175
  proof (intro exI conjI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4176
    show "a \<in> S"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4177
      using homotopic_with_imp_subset2 [OF homp]
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4178
      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
  4179
    have teq: "\<And>t. \<lbrakk>0 \<le> t; t \<le> 1\<rbrakk>
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4180
               \<Longrightarrow> t = Arg2pi (exp (2 * of_real pi * of_real t * \<i>)) / (2 * pi) \<or> t=1 \<and> Arg2pi (exp (2 * of_real pi * of_real t * \<i>)) = 0"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4181
      apply (rule disjCI)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4182
      using Arg2pi_of_real [of 1] apply (auto simp: Arg2pi_exp)
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4183
      done
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4184
    have "homotopic_loops S p (p \<circ> (\<lambda>z. Arg2pi z / (2 * pi)) \<circ> exp \<circ> (\<lambda>t. 2 * complex_of_real pi * complex_of_real t * \<i>))"
64394
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4185
      apply (rule homotopic_loops_eq [OF p])
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4186
      using p teq apply (fastforce simp: pathfinish_def pathstart_def)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4187
      done
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4188
    then
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4189
    show "homotopic_loops S p (linepath a a)"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4190
      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
  4191
  qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4192
qed
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4193
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4194
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4195
proposition simply_connected_eq_homotopic_circlemaps:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4196
  fixes S :: "'a::real_normed_vector set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4197
  shows "simply_connected S \<longleftrightarrow>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4198
         (\<forall>f g::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4199
              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
  4200
              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
  4201
              \<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
  4202
  apply (rule iffI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4203
   apply (blast elim: dest: simply_connected_eq_homotopic_circlemaps1)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4204
  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
  4205
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4206
proposition simply_connected_eq_contractible_circlemap:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4207
  fixes S :: "'a::real_normed_vector set"
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4208
  shows "simply_connected S \<longleftrightarrow>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4209
         path_connected S \<and>
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4210
         (\<forall>f::complex \<Rightarrow> 'a.
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4211
              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
  4212
              \<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
  4213
  apply (rule iffI)
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4214
   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
  4215
  using simply_connected_eq_homotopic_circlemaps3 by blast
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4216
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4217
corollary homotopy_eqv_simple_connectedness:
141e1ed8d5a0 more new material
paulson <lp15@cam.ac.uk>
parents: 64287
diff changeset
  4218
  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
  4219
  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
  4220
  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
  4221
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4222
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4223
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
  4224
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4225
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
  4226
  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
  4227
  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
  4228
  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
  4229
proof -
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4230
  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
  4231
    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
  4232
  then have hom: "homotopic_with (\<lambda>h. True) (sphere 0 1) (path_image \<gamma>)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4233
               (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi)))"
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4234
    by (rule homotopic_loops_imp_homotopic_circlemaps)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4235
  have "\<exists>g. homeomorphism (sphere 0 1) (path_image \<gamma>) (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) g"
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4236
  proof (rule homeomorphism_compact)
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4237
    show "continuous_on (sphere 0 1) (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi)))"
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4238
      using hom homotopic_with_imp_continuous by blast
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4239
    show "inj_on (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) (sphere 0 1)"
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4240
    proof
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4241
      fix x y
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4242
      assume xy: "x \<in> sphere 0 1" "y \<in> sphere 0 1"
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4243
         and eq: "(\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) x = (\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) y"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4244
      then have "(Arg2pi x / (2*pi)) = (Arg2pi y / (2*pi))"
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4245
      proof -
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4246
        have "(Arg2pi x / (2*pi)) \<in> {0..1}" "(Arg2pi y / (2*pi)) \<in> {0..1}"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4247
          using Arg2pi_ge_0 Arg2pi_lt_2pi dual_order.strict_iff_order by fastforce+
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4248
        with eq show ?thesis
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4249
          using \<open>simple_path \<gamma>\<close> Arg2pi_lt_2pi unfolding simple_path_def o_def
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4250
          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
  4251
      qed
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4252
      with xy show "x = y"
68499
d4312962161a Rationalisation of complex transcendentals, esp the Arg function
paulson <lp15@cam.ac.uk>
parents: 68493
diff changeset
  4253
        by (metis is_Arg_def Arg2pi Arg2pi_0 dist_0_norm divide_cancel_right dual_order.strict_iff_order mem_sphere)
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4254
    qed
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4255
    have "\<And>z. cmod z = 1 \<Longrightarrow> \<exists>x\<in>{0..1}. \<gamma> (Arg2pi z / (2*pi)) = \<gamma> x"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4256
       by (metis Arg2pi_ge_0 Arg2pi_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)
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4257
     moreover have "\<exists>z\<in>sphere 0 1. \<gamma> x = \<gamma> (Arg2pi z / (2*pi))" if "0 \<le> x" "x \<le> 1" for x
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4258
     proof (cases "x=1")
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4259
       case True
68517
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  4260
       with Arg2pi_of_real [of 1] loop show ?thesis
6b5f15387353 a few new lemmas
paulson <lp15@cam.ac.uk>
parents: 68499
diff changeset
  4261
         by (rule_tac x=1 in bexI) (auto simp: pathfinish_def pathstart_def \<open>0 \<le> x\<close>)
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4262
     next
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4263
       case False
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4264
       then have *: "(Arg2pi (exp (\<i>*(2* of_real pi* of_real x))) / (2*pi)) = x"
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4265
         using that by (auto simp: Arg2pi_exp divide_simps)
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4266
       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
  4267
         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
  4268
    qed
68493
143b4cc8fc74 Renaming Arg -> Arg2pi
paulson <lp15@cam.ac.uk>
parents: 68281
diff changeset
  4269
    ultimately show "(\<gamma> \<circ> (\<lambda>z. Arg2pi z / (2*pi))) ` sphere 0 1 = path_image \<gamma>"
64790
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4270
      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
  4271
    qed auto
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4272
    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
  4273
      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
  4274
  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
  4275
    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
  4276
  finally show ?thesis .
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4277
qed
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4278
ed38f9a834d8 New theory of arcwise connected sets and other new material
paulson <lp15@cam.ac.uk>
parents: 64773
diff changeset
  4279
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
  4280
  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
  4281
  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
  4282
  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
  4283
  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
  4284
  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
  4285
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  4286
end