src/HOL/Multivariate_Analysis/Complex_Transcendental.thy
author wenzelm
Wed, 30 Dec 2015 11:21:54 +0100
changeset 61973 0c7e865fa7cb
parent 61969 e01015e49041
child 62049 b0f941e207cf
permissions -rw-r--r--
more symbols;
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
61560
7c985fd653c5 tuned imports;
wenzelm
parents: 61524
diff changeset
     6
imports Complex_Analysis_Basics
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     7
begin
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
     8
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
     9
lemma cmod_add_real_less:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    10
  assumes "Im z \<noteq> 0" "r\<noteq>0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    11
    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
    12
proof (cases z)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    13
  case (Complex x y)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    14
  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
    15
    apply (rule real_less_rsqrt)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    16
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    17
    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
    18
    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
    19
  then show ?thesis using assms Complex
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    20
    apply (auto simp: cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    21
    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
    22
    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
    23
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    24
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    25
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
    26
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
    27
  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
    28
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    29
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    30
lemma cmod_square_less_1_plus:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    31
  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
    32
    shows "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    33
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    34
  apply (cases "Im z = 0 \<or> Re z = 0")
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    35
  using abs_square_less_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    36
    apply (force simp add: Re_power2 Im_power2 cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    37
  using cmod_diff_real_less [of "1 - z\<^sup>2" "1"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    38
  apply (simp add: norm_power Im_power2)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    39
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
    40
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    41
subsection\<open>The Exponential Function is Differentiable and Continuous\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    42
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    43
lemma complex_differentiable_within_exp: "exp complex_differentiable (at z within s)"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
    44
  using DERIV_exp complex_differentiable_at_within complex_differentiable_def by blast
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    45
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    46
lemma continuous_within_exp:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    47
  fixes z::"'a::{real_normed_field,banach}"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    48
  shows "continuous (at z within s) exp"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    49
by (simp add: continuous_at_imp_continuous_within)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    50
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    51
lemma continuous_on_exp:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    52
  fixes s::"'a::{real_normed_field,banach} set"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    53
  shows "continuous_on s exp"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    54
by (simp add: continuous_on_exp continuous_on_id)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    55
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    56
lemma holomorphic_on_exp: "exp holomorphic_on s"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    57
  by (simp add: complex_differentiable_within_exp holomorphic_on_def)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    58
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    59
subsection\<open>Euler and de Moivre formulas.\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    60
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
    61
text\<open>The sine series times @{term i}\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    62
lemma sin_ii_eq: "(\<lambda>n. (ii * sin_coeff n) * z^n) sums (ii * sin z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    63
proof -
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    64
  have "(\<lambda>n. ii * sin_coeff n *\<^sub>R z^n) sums (ii * sin z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    65
    using sin_converges sums_mult by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    66
  then show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    67
    by (simp add: scaleR_conv_of_real field_simps)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    68
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    69
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    70
theorem exp_Euler: "exp(ii * z) = cos(z) + ii * sin(z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    71
proof -
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
    72
  have "(\<lambda>n. (cos_coeff n + ii * sin_coeff n) * z^n)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    73
        = (\<lambda>n. (ii * z) ^ n /\<^sub>R (fact n))"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    74
  proof
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    75
    fix n
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    76
    show "(cos_coeff n + ii * sin_coeff n) * z^n = (ii * z) ^ n /\<^sub>R (fact n)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    77
      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
    78
  qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    79
  also have "... sums (exp (ii * z))"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    80
    by (rule exp_converges)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    81
  finally have "(\<lambda>n. (cos_coeff n + ii * sin_coeff n) * z^n) sums (exp (ii * z))" .
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    82
  moreover have "(\<lambda>n. (cos_coeff n + ii * sin_coeff n) * z^n) sums (cos z + ii * sin z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    83
    using sums_add [OF cos_converges [of z] sin_ii_eq [of z]]
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    84
    by (simp add: field_simps scaleR_conv_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    85
  ultimately show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    86
    using sums_unique2 by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    87
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    88
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    89
corollary exp_minus_Euler: "exp(-(ii * z)) = cos(z) - ii * sin(z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    90
  using exp_Euler [of "-z"]
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    91
  by simp
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    92
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    93
lemma sin_exp_eq: "sin z = (exp(ii * z) - exp(-(ii * z))) / (2*ii)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    94
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    95
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    96
lemma sin_exp_eq': "sin z = ii * (exp(-(ii * z)) - exp(ii * z)) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    97
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    98
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
    99
lemma cos_exp_eq:  "cos z = (exp(ii * z) + exp(-(ii * z))) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   100
  by (simp add: exp_Euler exp_minus_Euler)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   101
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   102
subsection\<open>Relationships between real and complex trig functions\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   103
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   104
lemma real_sin_eq [simp]:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   105
  fixes x::real
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   106
  shows "Re(sin(of_real x)) = sin x"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   107
  by (simp add: sin_of_real)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   108
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   109
lemma real_cos_eq [simp]:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   110
  fixes x::real
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   111
  shows "Re(cos(of_real x)) = cos x"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   112
  by (simp add: cos_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   113
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   114
lemma DeMoivre: "(cos z + ii * sin z) ^ n = cos(n * z) + ii * sin(n * z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   115
  apply (simp add: exp_Euler [symmetric])
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   116
  by (metis exp_of_nat_mult mult.left_commute)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   117
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   118
lemma exp_cnj:
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   119
  fixes z::complex
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   120
  shows "cnj (exp z) = exp (cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   121
proof -
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   122
  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
   123
    by auto
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   124
  also have "... sums (exp (cnj z))"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   125
    by (rule exp_converges)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   126
  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
   127
  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
   128
    by (metis exp_converges sums_cnj)
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   129
  ultimately show ?thesis
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   130
    using sums_unique2
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   131
    by blast
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   132
qed
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   133
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   134
lemma cnj_sin: "cnj(sin z) = sin(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   135
  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
   136
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   137
lemma cnj_cos: "cnj(cos z) = cos(cnj z)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   138
  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
   139
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   140
lemma complex_differentiable_at_sin: "sin complex_differentiable at z"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   141
  using DERIV_sin complex_differentiable_def by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   142
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   143
lemma complex_differentiable_within_sin: "sin complex_differentiable (at z within s)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   144
  by (simp add: complex_differentiable_at_sin complex_differentiable_at_within)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   145
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   146
lemma complex_differentiable_at_cos: "cos complex_differentiable at z"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   147
  using DERIV_cos complex_differentiable_def by blast
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   148
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   149
lemma complex_differentiable_within_cos: "cos complex_differentiable (at z within s)"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   150
  by (simp add: complex_differentiable_at_cos complex_differentiable_at_within)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   151
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   152
lemma holomorphic_on_sin: "sin holomorphic_on s"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   153
  by (simp add: complex_differentiable_within_sin holomorphic_on_def)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   154
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   155
lemma holomorphic_on_cos: "cos holomorphic_on s"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   156
  by (simp add: complex_differentiable_within_cos holomorphic_on_def)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   157
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   158
subsection\<open>Get a nice real/imaginary separation in Euler's formula.\<close>
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   159
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   160
lemma Euler: "exp(z) = of_real(exp(Re z)) *
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   161
              (of_real(cos(Im z)) + ii * of_real(sin(Im z)))"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   162
by (cases z) (simp add: exp_add exp_Euler cos_of_real exp_of_real sin_of_real)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   163
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   164
lemma Re_sin: "Re(sin z) = sin(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   165
  by (simp add: sin_exp_eq field_simps Re_divide Im_exp)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   166
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   167
lemma Im_sin: "Im(sin z) = cos(Re z) * (exp(Im z) - exp(-(Im z))) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   168
  by (simp add: sin_exp_eq field_simps Im_divide Re_exp)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   169
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   170
lemma Re_cos: "Re(cos z) = cos(Re z) * (exp(Im z) + exp(-(Im z))) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   171
  by (simp add: cos_exp_eq field_simps Re_divide Re_exp)
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   172
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   173
lemma Im_cos: "Im(cos z) = sin(Re z) * (exp(-(Im z)) - exp(Im z)) / 2"
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
   174
  by (simp add: cos_exp_eq field_simps Im_divide Im_exp)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   175
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   176
lemma Re_sin_pos: "0 < Re z \<Longrightarrow> Re z < pi \<Longrightarrow> Re (sin z) > 0"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   177
  by (auto simp: Re_sin Im_sin add_pos_pos sin_gt_zero)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   178
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   179
lemma Im_sin_nonneg: "Re z = 0 \<Longrightarrow> 0 \<le> Im z \<Longrightarrow> 0 \<le> Im (sin z)"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   180
  by (simp add: Re_sin Im_sin algebra_simps)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   181
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   182
lemma Im_sin_nonneg2: "Re z = pi \<Longrightarrow> Im z \<le> 0 \<Longrightarrow> 0 \<le> Im (sin z)"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   183
  by (simp add: Re_sin Im_sin algebra_simps)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   184
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   185
subsection\<open>More on the Polar Representation of Complex Numbers\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   186
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   187
lemma exp_Complex: "exp(Complex r t) = of_real(exp r) * Complex (cos t) (sin t)"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   188
  by (simp add: 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
   189
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   190
lemma exp_eq_1: "exp z = 1 \<longleftrightarrow> Re(z) = 0 \<and> (\<exists>n::int. 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
   191
apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   192
apply (metis exp_eq_one_iff norm_exp_eq_Re norm_one)
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
   193
apply (metis Re_exp cos_one_2pi_int mult.commute mult.left_neutral norm_exp_eq_Re norm_one one_complex.simps(1))
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
   194
by (metis Im_exp Re_exp complex_Re_Im_cancel_iff cos_one_2pi_int sin_double Re_complex_of_real complex_Re_numeral exp_zero mult.assoc mult.left_commute mult_eq_0_iff mult_numeral_1 numeral_One of_real_0 sin_zero_iff_int2)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   195
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   196
lemma exp_eq: "exp w = exp z \<longleftrightarrow> (\<exists>n::int. w = z + (of_int (2 * n) * pi) * ii)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   197
                (is "?lhs = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   198
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   199
  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
   200
    by (simp add: exp_diff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   201
  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
   202
    by (simp add: exp_eq_1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   203
  also have "... \<longleftrightarrow> ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   204
    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
   205
  finally show ?thesis .
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   206
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   207
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   208
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
   209
  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
   210
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   211
lemma exp_integer_2pi:
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   212
  assumes "n \<in> \<int>"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   213
  shows "exp((2 * n * pi) * ii) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   214
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   215
  have "exp((2 * n * pi) * ii) = exp 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   216
    using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   217
    by (simp only: Ints_def exp_eq) auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   218
  also have "... = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   219
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   220
  finally show ?thesis .
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   221
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   222
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   223
lemma sin_cos_eq_iff: "sin y = sin x \<and> cos y = cos x \<longleftrightarrow> (\<exists>n::int. y = x + 2 * n * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   224
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   225
  { 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
   226
    then have "cos (y-x) = 1"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   227
      using cos_add [of y "-x"] by simp
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
   228
    then have "\<exists>n::int. y-x = n * 2 * pi"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   229
      using cos_one_2pi_int by blast }
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   230
  then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   231
  apply (auto simp: sin_add cos_add)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   232
  apply (metis add.commute diff_add_cancel mult.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   233
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   234
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   235
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   236
lemma exp_i_ne_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   237
  assumes "0 < x" "x < 2*pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   238
  shows "exp(\<i> * of_real x) \<noteq> 1"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   239
proof
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   240
  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
   241
  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
   242
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   243
  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
   244
    by (simp only: Ints_def exp_eq) auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   245
  then have  "of_real x = (of_int (2 * n) * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   246
    by (metis complex_i_not_zero mult.commute mult_cancel_left of_real_eq_iff real_scaleR_def scaleR_conv_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   247
  then have  "x = (of_int (2 * n) * pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   248
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   249
  then show False using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   250
    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
   251
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   252
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   253
lemma sin_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   254
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   255
  shows "sin z = 0 \<longleftrightarrow> (\<exists>n::int. z = of_real(n * pi))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   256
  by (simp add: sin_exp_eq exp_eq of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   257
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   258
lemma cos_eq_0:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   259
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   260
  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
   261
  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
   262
  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
   263
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   264
lemma cos_eq_1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   265
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   266
  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
   267
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   268
  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
   269
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   270
  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
   271
    by (simp only: cos_double_sin)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   272
  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
   273
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   274
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   275
    by (auto simp: sin_eq_0 of_real_numeral)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   276
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   277
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   278
lemma csin_eq_1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   279
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   280
  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
   281
  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
   282
  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
   283
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   284
lemma csin_eq_minus1:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   285
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   286
  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
   287
        (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   288
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   289
  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
   290
    by (simp add: equation_minus_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   291
  also have "...  \<longleftrightarrow> (\<exists>n::int. -z = of_real(2 * n * pi) + of_real pi/2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   292
    by (simp only: csin_eq_1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   293
  also have "...  \<longleftrightarrow> (\<exists>n::int. z = - of_real(2 * n * pi) - of_real pi/2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   294
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   295
    by (metis (no_types)  is_num_normalize(8) minus_minus of_real_def real_vector.scale_minus_left uminus_add_conv_diff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   296
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   297
    apply (auto simp: of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   298
    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
   299
    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
   300
    apply (simp_all add: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   301
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   302
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   303
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   304
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   305
lemma ccos_eq_minus1:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   306
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   307
  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
   308
  using csin_eq_1 [of "z - of_real pi/2"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   309
  apply (simp add: sin_diff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   310
  apply (simp add: algebra_simps of_real_numeral equation_minus_iff)
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   311
  done
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   312
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   313
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
   314
                (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   315
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   316
  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
   317
    by (metis of_real_1 one_complex.simps(1) real_sin_eq sin_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   318
  also have "...  \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + of_real pi/2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   319
    by (simp only: csin_eq_1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   320
  also have "...  \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + of_real pi/2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   321
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   322
    apply (auto simp: algebra_simps of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   323
    apply (rule injD [OF inj_of_real [where 'a = complex]])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   324
    apply (auto simp: of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   325
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   326
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   327
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   328
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   329
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   330
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   331
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
   332
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   333
  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
   334
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left sin_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   335
  also have "...  \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + 3/2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   336
    by (simp only: csin_eq_minus1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   337
  also have "...  \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + 3/2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   338
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   339
    apply (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   340
    apply (rule injD [OF inj_of_real [where 'a = complex]], auto)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   341
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   342
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   343
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   344
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   345
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   346
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   347
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
   348
                      (is "_ = ?rhs")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   349
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   350
  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
   351
    by (metis Re_complex_of_real of_real_def scaleR_minus1_left cos_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   352
  also have "...  \<longleftrightarrow> (\<exists>n::int. complex_of_real x = of_real(2 * n * pi) + pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   353
    by (simp only: ccos_eq_minus1)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   354
  also have "...  \<longleftrightarrow> (\<exists>n::int. x = of_real(2 * n * pi) + pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   355
    apply (rule iff_exI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   356
    apply (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   357
    apply (rule injD [OF inj_of_real [where 'a = complex]], auto)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   358
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   359
  also have "... = ?rhs"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   360
    by (auto simp: algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   361
  finally show ?thesis .
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   362
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   363
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   364
lemma dist_exp_ii_1: "norm(exp(ii * 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
   365
  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
   366
  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
   367
  apply (simp add: real_sqrt_mult)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   368
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   369
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   370
lemma sinh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   371
  fixes z :: complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   372
  shows "(exp z - inverse (exp z)) / 2 = -ii * sin(ii * z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   373
  by (simp add: sin_exp_eq divide_simps exp_minus of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   374
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   375
lemma sin_ii_times:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   376
  fixes z :: complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   377
  shows "sin(ii * z) = ii * ((exp z - inverse (exp z)) / 2)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   378
  using sinh_complex by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   379
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   380
lemma sinh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   381
  fixes x :: real
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   382
  shows "of_real((exp x - inverse (exp x)) / 2) = -ii * sin(ii * of_real x)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   383
  by (simp add: exp_of_real sin_ii_times of_real_numeral)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   384
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   385
lemma cosh_complex:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   386
  fixes z :: complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   387
  shows "(exp z + inverse (exp z)) / 2 = cos(ii * z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   388
  by (simp add: cos_exp_eq divide_simps exp_minus of_real_numeral exp_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   389
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   390
lemma cosh_real:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   391
  fixes x :: real
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   392
  shows "of_real((exp x + inverse (exp x)) / 2) = cos(ii * of_real x)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   393
  by (simp add: cos_exp_eq divide_simps exp_minus of_real_numeral exp_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   394
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   395
lemmas cos_ii_times = cosh_complex [symmetric]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   396
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   397
lemma norm_cos_squared:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   398
    "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
   399
  apply (cases z)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   400
  apply (simp add: cos_add cmod_power2 cos_of_real sin_of_real)
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
   401
  apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   402
  apply (simp only: left_diff_distrib [symmetric] power_mult_distrib)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   403
  apply (simp add: sin_squared_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   404
  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
   405
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   406
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   407
lemma norm_sin_squared:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   408
    "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
   409
  apply (cases z)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   410
  apply (simp add: sin_add cmod_power2 cos_of_real sin_of_real cos_double_cos exp_double)
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
   411
  apply (simp add: cos_exp_eq sin_exp_eq exp_minus exp_of_real Re_divide Im_divide power_divide)
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   412
  apply (simp only: left_diff_distrib [symmetric] power_mult_distrib)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   413
  apply (simp add: cos_squared_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   414
  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
   415
  done
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   416
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   417
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
   418
  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
   419
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   420
lemma norm_cos_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   421
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   422
  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
   423
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   424
  have "Im z \<le> cmod z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   425
    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
   426
  with exp_uminus_Im show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   427
    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
   428
    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
   429
    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
   430
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   431
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   432
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   433
lemma norm_cos_plus1_le:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   434
  fixes z::complex
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   435
  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
   436
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   437
  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
   438
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   439
  have *: "Im z \<le> cmod z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   440
    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
   441
  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
   442
    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
   443
  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
   444
    by (simp add: cos_exp_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   445
  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
   446
    by (simp add: field_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   447
  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
   448
    by (simp add: norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   449
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   450
    apply (rule ssubst, simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   451
    apply (rule order_trans [OF triangle3], simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   452
    using exp_uminus_Im *
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   453
    apply (auto intro: mono)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   454
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   455
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   456
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   457
subsection\<open>Taylor series for complex exponential, sine and cosine.\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   458
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   459
declare power_Suc [simp del]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   460
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   461
lemma Taylor_exp:
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   462
  "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
   463
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
   464
  show "convex (closed_segment 0 z)"
61518
ff12606337e9 new lemmas about topology, etc., for Cauchy integral formula
paulson
parents: 61426
diff changeset
   465
    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
   466
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   467
  fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   468
  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
   469
  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
   470
    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
   471
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   472
  fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   473
  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
   474
  then show "Re x \<le> \<bar>Re z\<bar>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   475
    apply (auto simp: closed_segment_def scaleR_conv_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   476
    by (meson abs_ge_self abs_ge_zero linear mult_left_le_one_le mult_nonneg_nonpos order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   477
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   478
  show "0 \<in> closed_segment 0 z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   479
    by (auto simp: closed_segment_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   480
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   481
  show "z \<in> closed_segment 0 z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   482
    apply (simp add: closed_segment_def scaleR_conv_of_real)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   483
    using of_real_1 zero_le_one by blast
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   484
qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   485
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   486
lemma
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   487
  assumes "0 \<le> u" "u \<le> 1"
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   488
  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
   489
    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
   490
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   491
  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
   492
    by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   493
  show "cmod (sin (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   494
    apply (auto simp: scaleR_conv_of_real norm_mult norm_power sin_exp_eq norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   495
    apply (rule order_trans [OF norm_triangle_ineq4])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   496
    apply (rule mono)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   497
    apply (auto simp: abs_if mult_left_le_one_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   498
    apply (meson mult_nonneg_nonneg neg_le_0_iff_le not_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   499
    apply (meson less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   500
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   501
  show "cmod (cos (u *\<^sub>R z)) \<le> exp \<bar>Im z\<bar>" using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   502
    apply (auto simp: scaleR_conv_of_real norm_mult norm_power cos_exp_eq norm_divide)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   503
    apply (rule order_trans [OF norm_triangle_ineq])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   504
    apply (rule mono)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   505
    apply (auto simp: abs_if mult_left_le_one_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   506
    apply (meson mult_nonneg_nonneg neg_le_0_iff_le not_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   507
    apply (meson less_eq_real_def mult_nonneg_nonpos neg_0_le_iff_le order_trans)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   508
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   509
qed
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   510
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   511
lemma Taylor_sin:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   512
  "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
   513
   \<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
   514
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   515
  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
   516
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   517
  have *: "cmod (sin z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   518
                 (\<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
   519
           \<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
   520
  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
   521
                                 "\<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
   522
                                 "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
   523
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   524
    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
   525
            (- 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
   526
            (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
   527
      apply (auto simp: power_Suc)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   528
      apply (intro derivative_eq_intros | simp)+
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   529
      done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   530
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   531
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   532
    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
   533
    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
   534
      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
   535
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   536
  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
   537
            = (-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
   538
    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
   539
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   540
    apply (rule order_trans [OF _ *])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   541
    apply (simp add: **)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   542
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   543
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   544
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   545
lemma Taylor_cos:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   546
  "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
   547
   \<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
   548
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   549
  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
   550
      by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   551
  have *: "cmod (cos z -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   552
                 (\<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
   553
           \<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
   554
  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
   555
simplified])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   556
    fix k x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   557
    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
   558
    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
   559
            (- 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
   560
             (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
   561
      apply (auto simp: power_Suc)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   562
      apply (intro derivative_eq_intros | simp)+
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   563
      done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   564
  next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   565
    fix x
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   566
    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
   567
    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
   568
      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
   569
  qed
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   570
  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
   571
            = (-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
   572
    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
   573
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   574
    apply (rule order_trans [OF _ *])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   575
    apply (simp add: **)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   576
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   577
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   578
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
   579
declare power_Suc [simp]
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   580
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   581
text\<open>32-bit Approximation to e\<close>
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
   582
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
   583
  using Taylor_exp [of 1 14] exp_le
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   584
  apply (simp add: setsum_left_distrib in_Reals_norm Re_exp atMost_nat_numeral fact_numeral)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   585
  apply (simp only: pos_le_divide_eq [symmetric], linarith)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   586
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   587
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
   588
lemma e_less_3: "exp 1 < (3::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
   589
  using e_approx_32
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   590
  by (simp add: abs_if split: split_if_asm)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   591
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   592
lemma ln3_gt_1: "ln 3 > (1::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
   593
  by (metis e_less_3 exp_less_cancel_iff exp_ln_iff less_trans ln_exp)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   594
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
   595
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   596
subsection\<open>The argument of a complex number\<close>
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   597
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   598
definition Arg :: "complex \<Rightarrow> real" where
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   599
 "Arg z \<equiv> if z = 0 then 0
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   600
           else THE t. 0 \<le> t \<and> t < 2*pi \<and>
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   601
                    z = of_real(norm z) * exp(ii * of_real t)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   602
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   603
lemma Arg_0 [simp]: "Arg(0) = 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   604
  by (simp add: Arg_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   605
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   606
lemma Arg_unique_lemma:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   607
  assumes z:  "z = of_real(norm z) * exp(ii * of_real t)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   608
      and z': "z = of_real(norm z) * exp(ii * of_real t')"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   609
      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
   610
      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
   611
      and nz: "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   612
  shows "t' = t"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   613
proof -
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   614
  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
   615
    by arith
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   616
  have "of_real (cmod z) * exp (\<i> * of_real t') = of_real (cmod z) * exp (\<i> * of_real t)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   617
    by (metis z z')
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   618
  then have "exp (\<i> * of_real t') = exp (\<i> * of_real t)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   619
    by (metis nz mult_left_cancel mult_zero_left z)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   620
  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
   621
    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
   622
    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
   623
  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
   624
    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
   625
  then have "n=0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   626
    apply (rule_tac z=n in int_cases)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   627
    using t t'
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   628
    apply (auto simp: mult_less_0_iff algebra_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   629
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   630
  then show "t' = t"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   631
      by (simp add: n)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   632
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   633
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   634
lemma Arg: "0 \<le> Arg z & Arg z < 2*pi & z = of_real(norm z) * exp(ii * of_real(Arg z))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   635
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   636
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   637
    by (simp add: Arg_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   638
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   639
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   640
  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
   641
             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
   642
    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
   643
    by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   644
  have z: "z = of_real(norm z) * exp(ii * of_real t)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   645
    apply (rule complex_eqI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   646
    using t False ReIm
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   647
    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
   648
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   649
  show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   650
    apply (simp add: Arg_def False)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   651
    apply (rule theI [where a=t])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   652
    using t z False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   653
    apply (auto intro: Arg_unique_lemma)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   654
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   655
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   656
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   657
corollary
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   658
  shows Arg_ge_0: "0 \<le> Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   659
    and Arg_lt_2pi: "Arg z < 2*pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   660
    and Arg_eq: "z = of_real(norm z) * exp(ii * of_real(Arg z))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   661
  using Arg by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   662
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   663
lemma complex_norm_eq_1_exp: "norm z = 1 \<longleftrightarrow> (\<exists>t. z = exp(ii * of_real t))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   664
  using Arg [of z] by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   665
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   666
lemma Arg_unique: "\<lbrakk>of_real r * exp(ii * of_real a) = z; 0 < r; 0 \<le> a; a < 2*pi\<rbrakk> \<Longrightarrow> Arg z = a"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   667
  apply (rule Arg_unique_lemma [OF _ Arg_eq])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   668
  using Arg [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   669
  apply (auto simp: norm_mult)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   670
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   671
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   672
lemma Arg_minus: "z \<noteq> 0 \<Longrightarrow> Arg (-z) = (if Arg z < pi then Arg z + pi else Arg z - pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   673
  apply (rule Arg_unique [of "norm z"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   674
  apply (rule complex_eqI)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   675
  using Arg_ge_0 [of z] Arg_eq [of z] Arg_lt_2pi [of z] Arg_eq [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   676
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   677
  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
   678
  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
   679
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   680
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   681
lemma Arg_times_of_real [simp]: "0 < r \<Longrightarrow> Arg (of_real r * z) = Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   682
  apply (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   683
  apply (rule Arg_unique [of "r * norm z"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   684
  using Arg
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   685
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   686
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   687
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   688
lemma Arg_times_of_real2 [simp]: "0 < r \<Longrightarrow> Arg (z * of_real r) = Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   689
  by (metis Arg_times_of_real mult.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   690
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   691
lemma Arg_divide_of_real [simp]: "0 < r \<Longrightarrow> Arg (z / of_real r) = Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   692
  by (metis Arg_times_of_real2 less_numeral_extra(3) nonzero_eq_divide_eq of_real_eq_0_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   693
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   694
lemma Arg_le_pi: "Arg z \<le> pi \<longleftrightarrow> 0 \<le> Im z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   695
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   696
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   697
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   698
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   699
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   700
  have "0 \<le> Im z \<longleftrightarrow> 0 \<le> Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   701
    by (metis Arg_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   702
  also have "... = (0 \<le> Im (exp (\<i> * complex_of_real (Arg z))))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   703
    using False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   704
    by (simp add: zero_le_mult_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   705
  also have "... \<longleftrightarrow> Arg z \<le> pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   706
    by (simp add: Im_exp) (metis Arg_ge_0 Arg_lt_2pi sin_lt_zero sin_ge_zero not_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   707
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   708
    by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   709
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   710
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   711
lemma Arg_lt_pi: "0 < Arg z \<and> Arg z < pi \<longleftrightarrow> 0 < Im z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   712
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   713
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   714
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   715
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   716
  case False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   717
  have "0 < Im z \<longleftrightarrow> 0 < Im (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   718
    by (metis Arg_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   719
  also have "... = (0 < Im (exp (\<i> * complex_of_real (Arg z))))"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   720
    using False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   721
    by (simp add: zero_less_mult_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   722
  also have "... \<longleftrightarrow> 0 < Arg z \<and> Arg z < pi"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   723
    using Arg_ge_0  Arg_lt_2pi sin_le_zero sin_gt_zero
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   724
    apply (auto simp: Im_exp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   725
    using le_less apply fastforce
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   726
    using not_le by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   727
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   728
    by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   729
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   730
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   731
lemma Arg_eq_0: "Arg z = 0 \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re z"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   732
proof (cases "z=0")
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   733
  case True then show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   734
    by simp
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   735
next
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   736
  case False
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   737
  have "z \<in> \<real> \<and> 0 \<le> Re z \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (of_real (cmod z) * exp (\<i> * complex_of_real (Arg z)))"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   738
    by (metis Arg_eq)
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   739
  also have "... \<longleftrightarrow> z \<in> \<real> \<and> 0 \<le> Re (exp (\<i> * complex_of_real (Arg z)))"
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   740
    using False
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   741
    by (simp add: zero_le_mult_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   742
  also have "... \<longleftrightarrow> Arg z = 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   743
    apply (auto simp: Re_exp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   744
    apply (metis Arg_lt_pi Arg_ge_0 Arg_le_pi cos_pi complex_is_Real_iff leD less_linear less_minus_one_simps(2) minus_minus neg_less_eq_nonneg order_refl)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   745
    using Arg_eq [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   746
    apply (auto simp: Reals_def)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   747
    done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   748
  finally show ?thesis
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   749
    by blast
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   750
qed
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   751
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
   752
corollary Arg_gt_0:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   753
  assumes "z \<in> \<real> \<Longrightarrow> Re z < 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   754
    shows "Arg z > 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   755
  using Arg_eq_0 Arg_ge_0 assms dual_order.strict_iff_order by fastforce
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   756
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   757
lemma Arg_of_real: "Arg(of_real x) = 0 \<longleftrightarrow> 0 \<le> x"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   758
  by (simp add: Arg_eq_0)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   759
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   760
lemma Arg_eq_pi: "Arg z = pi \<longleftrightarrow> z \<in> \<real> \<and> Re z < 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   761
  apply  (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   762
  using Arg_eq_0 [of "-z"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   763
  apply (auto simp: complex_is_Real_iff Arg_minus)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   764
  apply (simp add: complex_Re_Im_cancel_iff)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   765
  apply (metis Arg_minus pi_gt_zero add.left_neutral minus_minus minus_zero)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   766
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   767
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   768
lemma Arg_eq_0_pi: "Arg z = 0 \<or> Arg z = pi \<longleftrightarrow> z \<in> \<real>"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   769
  using Arg_eq_0 Arg_eq_pi not_le by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   770
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   771
lemma Arg_inverse: "Arg(inverse z) = (if z \<in> \<real> \<and> 0 \<le> Re z then Arg z else 2*pi - Arg z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   772
  apply (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   773
  apply (rule Arg_unique [of "inverse (norm z)"])
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   774
  using Arg_ge_0 [of z] Arg_lt_2pi [of z] Arg_eq [of z] Arg_eq_0 [of z] exp_two_pi_i
59746
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   775
  apply (auto simp: of_real_numeral algebra_simps exp_diff divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   776
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   777
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   778
lemma Arg_eq_iff:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   779
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   780
     shows "Arg w = Arg z \<longleftrightarrow> (\<exists>x. 0 < x & w = of_real x * z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   781
  using assms Arg_eq [of z] Arg_eq [of w]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   782
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   783
  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
   784
  apply (simp add: divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   785
  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
   786
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   787
lemma Arg_inverse_eq_0: "Arg(inverse z) = 0 \<longleftrightarrow> Arg z = 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   788
  using complex_is_Real_iff
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   789
  apply (simp add: Arg_eq_0)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   790
  apply (auto simp: divide_simps not_sum_power2_lt_zero)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   791
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   792
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   793
lemma Arg_divide:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   794
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg w \<le> Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   795
    shows "Arg(z / w) = Arg z - Arg w"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   796
  apply (rule Arg_unique [of "norm(z / w)"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   797
  using assms Arg_eq [of z] Arg_eq [of w] Arg_ge_0 [of w] Arg_lt_2pi [of z]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   798
  apply (auto simp: exp_diff norm_divide algebra_simps divide_simps)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   799
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   800
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   801
lemma Arg_le_div_sum:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   802
  assumes "w \<noteq> 0" "z \<noteq> 0" "Arg w \<le> Arg z"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   803
    shows "Arg z = Arg w + Arg(z / w)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   804
  by (simp add: Arg_divide assms)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   805
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   806
lemma Arg_le_div_sum_eq:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   807
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   808
    shows "Arg w \<le> Arg z \<longleftrightarrow> Arg z = Arg w + Arg(z / w)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   809
  using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   810
  by (auto simp: Arg_ge_0 intro: Arg_le_div_sum)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   811
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   812
lemma Arg_diff:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   813
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   814
    shows "Arg w - Arg z = (if Arg z \<le> Arg w then Arg(w / z) else Arg(w/z) - 2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   815
  using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   816
  apply (auto simp: Arg_ge_0 Arg_divide not_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   817
  using Arg_divide [of w z] Arg_inverse [of "w/z"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   818
  apply auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   819
  by (metis Arg_eq_0 less_irrefl minus_diff_eq right_minus_eq)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   820
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   821
lemma Arg_add:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   822
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   823
    shows "Arg w + Arg z = (if Arg w + Arg z < 2*pi then Arg(w * z) else Arg(w * z) + 2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   824
  using assms
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   825
  using Arg_diff [of "w*z" z] Arg_le_div_sum_eq [of z "w*z"]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   826
  apply (auto simp: Arg_ge_0 Arg_divide not_le)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   827
  apply (metis Arg_lt_2pi add.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   828
  apply (metis (no_types) Arg add.commute diff_0 diff_add_cancel diff_less_eq diff_minus_eq_add not_less)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   829
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   830
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   831
lemma Arg_times:
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   832
  assumes "w \<noteq> 0" "z \<noteq> 0"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   833
    shows "Arg (w * z) = (if Arg w + Arg z < 2*pi then Arg w + Arg z
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   834
                            else (Arg w + Arg z) - 2*pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   835
  using Arg_add [OF assms]
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   836
  by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   837
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   838
lemma Arg_cnj: "Arg(cnj z) = (if z \<in> \<real> \<and> 0 \<le> Re z then Arg z else 2*pi - Arg z)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   839
  apply (cases "z=0", simp)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   840
  apply (rule trans [of _ "Arg(inverse z)"])
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   841
  apply (simp add: Arg_eq_iff divide_simps complex_norm_square [symmetric] mult.commute)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   842
  apply (metis norm_eq_zero of_real_power zero_less_power2)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   843
  apply (auto simp: of_real_numeral Arg_inverse)
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   844
  done
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   845
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   846
lemma Arg_real: "z \<in> \<real> \<Longrightarrow> Arg z = (if 0 \<le> Re z then 0 else pi)"
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   847
  using Arg_eq_0 Arg_eq_0_pi
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   848
  by auto
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   849
ddae5727c5a9 new HOL Light material about exp, sin, cos
paulson <lp15@cam.ac.uk>
parents: 59745
diff changeset
   850
lemma Arg_exp: "0 \<le> Im z \<Longrightarrow> Im z < 2*pi \<Longrightarrow> Arg(exp z) = Im z"
61762
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   851
  by (rule Arg_unique [of  "exp(Re z)"]) (auto simp: exp_eq_polar)
d50b993b4fb9 Removal of redundant lemmas (diff_less_iff, diff_le_iff) and of the abbreviation Exp. Addition of some new material.
paulson <lp15@cam.ac.uk>
parents: 61711
diff changeset
   852
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
   853
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
   854
  obtains r a::real where "z = complex_of_real r * (cos a + \<i> * sin a)" "0 \<le> r" "0 \<le> a" "a < 2*pi"
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
   855
  using Arg cis.ctr cis_conv_exp by fastforce
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   856
61806
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
   857
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
   858
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
   859
  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
   860
    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
   861
    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
   862
qed
d2e62ae01cd8 Cauchy's integral formula for circles. Starting to fix eventually_mono.
paulson <lp15@cam.ac.uk>
parents: 61762
diff changeset
   863
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   864
subsection\<open>Analytic properties of tangent function\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   865
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   866
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
   867
  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
   868
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   869
lemma complex_differentiable_at_tan: "~(cos z = 0) \<Longrightarrow> tan complex_differentiable at z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   870
  unfolding complex_differentiable_def
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   871
  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
   872
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   873
lemma complex_differentiable_within_tan: "~(cos z = 0)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   874
         \<Longrightarrow> tan complex_differentiable (at z within s)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   875
  using complex_differentiable_at_tan complex_differentiable_at_within by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   876
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   877
lemma continuous_within_tan: "~(cos z = 0) \<Longrightarrow> continuous (at z within s) tan"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   878
  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
   879
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   880
lemma continuous_on_tan [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 0)) \<Longrightarrow> continuous_on s tan"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   881
  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
   882
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   883
lemma holomorphic_on_tan: "(\<And>z. z \<in> s \<Longrightarrow> ~(cos z = 0)) \<Longrightarrow> tan holomorphic_on s"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   884
  by (simp add: complex_differentiable_within_tan holomorphic_on_def)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   885
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   886
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   887
subsection\<open>Complex logarithms (the conventional principal value)\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   888
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
   889
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
   890
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
   891
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
   892
definition ln_complex :: "complex \<Rightarrow> complex"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   893
  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
   894
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   895
lemma
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   896
  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
   897
    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
   898
      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
   899
      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
   900
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   901
  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
   902
    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
   903
    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
   904
  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
   905
    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
   906
    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
   907
  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
   908
    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
   909
    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
   910
    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
   911
    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
   912
  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
   913
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   914
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   915
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   916
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
   917
  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
   918
    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
   919
  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
   920
  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
   921
  apply auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   922
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   923
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   924
subsection\<open>Relation to Real Logarithm\<close>
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   925
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   926
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
   927
  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
   928
    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
   929
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
   930
  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
   931
    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
   932
  also have "... = of_real(ln z)"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   933
    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
   934
    by (subst Ln_exp) auto
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   935
  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
   936
    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
   937
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
   938
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   939
corollary Ln_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> Re z > 0 \<Longrightarrow> ln z \<in> \<real>"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   940
  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
   941
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   942
corollary Im_Ln_of_real [simp]: "r > 0 \<Longrightarrow> Im (ln (of_real r)) = 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   943
  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
   944
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
   945
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
   946
  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
   947
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
   948
lemma Ln_1: "ln 1 = (0::complex)"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   949
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
   950
  have "ln (exp 0) = (0::complex)"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   951
    by (metis (mono_tags, hide_lams) Ln_of_real exp_zero ln_one of_real_0 of_real_1 zero_less_one)
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   952
  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
   953
    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
   954
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
   955
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   956
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
   957
  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
   958
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   959
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
   960
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
   961
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
   962
  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
   963
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   964
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
   965
  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
   966
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   967
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
   968
  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
   969
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   970
lemma Re_Ln [simp]: "z \<noteq> 0 \<Longrightarrow> Re(Ln z) = ln(norm z)"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   971
  by (metis exp_Ln assms ln_exp norm_exp_eq_Re)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   972
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
   973
corollary ln_cmod_le:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
   974
  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
   975
    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
   976
  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
   977
  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
   978
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   979
lemma exists_complex_root:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   980
  fixes a :: complex
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   981
  shows "n \<noteq> 0 \<Longrightarrow> \<exists>z. z ^ n = a"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   982
  apply (cases "a=0", simp)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   983
  apply (rule_tac x= "exp(Ln(a) / n)" in exI)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   984
  apply (auto simp: exp_of_nat_mult [symmetric])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   985
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
   986
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   987
subsection\<open>The Unwinding Number and the Ln-product Formula\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   988
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
   989
text\<open>Note that in this special case the unwinding number is -1, 0 or 1.\<close>
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   990
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   991
definition unwinding :: "complex \<Rightarrow> complex" where
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   992
   "unwinding(z) = (z - Ln(exp z)) / (of_real(2*pi) * ii)"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   993
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   994
lemma unwinding_2pi: "(2*pi) * ii * unwinding(z) = z - Ln(exp z)"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   995
  by (simp add: unwinding_def)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   996
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   997
lemma Ln_times_unwinding:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   998
    "w \<noteq> 0 \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Ln(w * z) = Ln(w) + Ln(z) - (2*pi) * ii * unwinding(Ln w + Ln z)"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
   999
  using unwinding_2pi by (simp add: exp_add)
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1000
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1001
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1002
subsection\<open>Derivative of Ln away from the branch cut\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1003
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1004
lemma
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1005
  assumes "Im(z) = 0 \<Longrightarrow> 0 < Re(z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1006
    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
  1007
      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
  1008
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1009
  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
  1010
    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
  1011
  then show *: "Im (Ln z) < pi" using assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1012
    by (metis exp_Ln Im_Ln_le_pi Im_exp Re_exp abs_of_nonneg cmod_eq_Re cos_pi mult.right_neutral mult_minus_right mult_zero_right neg_less_0_iff_less norm_exp_eq_Re not_less not_less_iff_gr_or_eq sin_pi)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1013
  show "(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
  1014
    apply (rule has_complex_derivative_inverse_strong_x
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1015
              [where f = exp and s = "{w. -pi < Im(w) & Im(w) < pi}"])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1016
    using znz *
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1017
    apply (auto simp: continuous_on_exp open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1018
    apply (metis DERIV_exp exp_Ln)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1019
    apply (metis mpi_less_Im_Ln)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1020
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1021
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1022
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1023
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
  1024
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
  1025
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1026
lemma complex_differentiable_at_Ln: "(Im(z) = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> Ln complex_differentiable at z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1027
  using complex_differentiable_def has_field_derivative_Ln by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1028
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1029
lemma complex_differentiable_within_Ln: "(Im(z) = 0 \<Longrightarrow> 0 < Re(z))
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1030
         \<Longrightarrow> Ln complex_differentiable (at z within s)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1031
  using complex_differentiable_at_Ln complex_differentiable_within_subset by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1032
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1033
lemma continuous_at_Ln: "(Im(z) = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> continuous (at z) Ln"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1034
  by (simp add: complex_differentiable_imp_continuous_at complex_differentiable_within_Ln)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1035
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1036
lemma isCont_Ln' [simp]:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1037
   "\<lbrakk>isCont f z; Im(f z) = 0 \<Longrightarrow> 0 < Re(f z)\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. Ln (f x)) z"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1038
  by (blast intro: isCont_o2 [OF _ continuous_at_Ln])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1039
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1040
lemma continuous_within_Ln: "(Im(z) = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> continuous (at z within s) Ln"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1041
  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
  1042
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1043
lemma continuous_on_Ln [continuous_intros]: "(\<And>z. z \<in> s \<Longrightarrow> Im(z) = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> continuous_on s Ln"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1044
  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
  1045
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1046
lemma holomorphic_on_Ln: "(\<And>z. z \<in> s \<Longrightarrow> Im(z) = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> Ln holomorphic_on s"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1047
  by (simp add: complex_differentiable_within_Ln holomorphic_on_def)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1048
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1049
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1050
subsection\<open>Quadrant-type results for Ln\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1051
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1052
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
  1053
  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
  1054
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1055
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1056
lemma 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
  1057
  assumes "z \<noteq> 0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1058
    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
  1059
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1060
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1061
    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
  1062
    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
  1063
      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
  1064
      by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1065
    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
  1066
      apply (auto simp: Re_exp zero_less_mult_iff cos_gt_zero_pi)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1067
      using cos_lt_zero_pi [of "-(Im w)"] cos_lt_zero_pi [of "(Im w)"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1068
      apply (simp add: abs_if split: split_if_asm)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1069
      apply (metis (no_types) cos_minus cos_pi_half eq_divide_eq_numeral1(1) eq_numeral_simps(4)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1070
               less_numeral_extra(3) linorder_neqE_linordered_idom minus_mult_minus minus_mult_right
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1071
               mult_numeral_1_right)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1072
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1073
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1074
  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
  1075
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1076
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1077
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1078
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
  1079
  assumes "z \<noteq> 0"
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1080
    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
  1081
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1082
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1083
    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
  1084
    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
  1085
      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
  1086
      by auto
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1087
    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
  1088
      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
  1089
      using cos_lt_zero_pi [of "- (Im w)"] cos_lt_zero_pi [of "(Im w)"] not_le
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1090
      apply (auto simp: abs_if split: split_if_asm)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1091
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1092
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1093
  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
  1094
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1095
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1096
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1097
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
  1098
  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
  1099
    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
  1100
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1101
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1102
    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
  1103
    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
  1104
      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
  1105
      by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1106
    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
  1107
      using sin_gt_zero [of "- (Im w)"] sin_gt_zero [of "(Im w)"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1108
      apply (auto simp: Im_exp zero_less_mult_iff)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1109
      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
  1110
      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
  1111
      done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1112
  }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1113
  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
  1114
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1115
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1116
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1117
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
  1118
  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
  1119
    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
  1120
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1121
  { fix w
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1122
    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
  1123
    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
  1124
      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
  1125
      by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1126
    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
  1127
      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
  1128
      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
  1129
      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
  1130
      done }
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1131
  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
  1132
    by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1133
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1134
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1135
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
  1136
  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
  1137
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1138
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
  1139
  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
  1140
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1141
lemma Im_Ln_eq_0: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = 0 \<longleftrightarrow> 0 < Re(z) \<and> Im(z) = 0)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1142
  by (metis exp_Ln Im_Ln_less_pi Im_Ln_pos_le Im_Ln_pos_lt Re_complex_of_real add.commute add.left_neutral
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1143
       complex_eq exp_of_real le_less mult_zero_right norm_exp_eq_Re norm_le_zero_iff not_le of_real_0)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1144
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1145
lemma Im_Ln_eq_pi: "z \<noteq> 0 \<Longrightarrow> (Im(Ln z) = pi \<longleftrightarrow> Re(z) < 0 \<and> Im(z) = 0)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1146
  by (metis Im_Ln_eq_0 Im_Ln_less_pi Im_Ln_pos_le Im_Ln_pos_lt add.right_neutral complex_eq mult_zero_right not_less not_less_iff_gr_or_eq of_real_0)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1147
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1148
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1149
subsection\<open>More Properties of Ln\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1150
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1151
lemma cnj_Ln: "(Im z = 0 \<Longrightarrow> 0 < Re z) \<Longrightarrow> cnj(Ln z) = Ln(cnj z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1152
  apply (cases "z=0", auto)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1153
  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
  1154
  apply (auto simp: abs_if split: split_if_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
  1155
  apply (metis Im_Ln_less_pi add_mono_thms_linordered_field(5) cnj.simps mult_2 neg_equal_0_iff_equal)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1156
  apply (metis complex_cnj_zero_iff diff_minus_eq_add diff_strict_mono minus_less_iff mpi_less_Im_Ln mult.commute mult_2_right)
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1157
  by (metis exp_Ln exp_cnj)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1158
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1159
lemma Ln_inverse: "(Im(z) = 0 \<Longrightarrow> 0 < Re z) \<Longrightarrow> Ln(inverse z) = -(Ln z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1160
  apply (cases "z=0", auto)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1161
  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
  1162
  using mpi_less_Im_Ln [of z] mpi_less_Im_Ln [of "inverse z"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1163
  apply (auto simp: abs_if exp_minus split: split_if_asm)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1164
  apply (metis Im_Ln_less_pi Im_Ln_pos_le add_less_cancel_left add_strict_mono
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1165
               inverse_inverse_eq inverse_zero le_less mult.commute mult_2_right)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1166
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1167
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1168
lemma Ln_minus1 [simp]: "Ln(-1) = ii * pi"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1169
  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
  1170
  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
  1171
  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
  1172
  done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1173
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1174
lemma Ln_ii [simp]: "Ln ii = ii * of_real pi/2"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1175
  using Ln_exp [of "ii * (of_real pi/2)"]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1176
  unfolding exp_Euler
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1177
  by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1178
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1179
lemma Ln_minus_ii [simp]: "Ln(-ii) = - (ii * pi/2)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1180
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1181
  have  "Ln(-ii) = Ln(1/ii)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1182
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1183
  also have "... = - (Ln ii)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1184
    by (metis Ln_inverse ii.sel(2) inverse_eq_divide zero_neq_one)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1185
  also have "... = - (ii * pi/2)"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1186
    by simp
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1187
  finally show ?thesis .
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1188
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1189
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1190
lemma Ln_times:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1191
  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
  1192
    shows "Ln(w * z) =
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1193
                (if Im(Ln w + Ln z) \<le> -pi then
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1194
                  (Ln(w) + Ln(z)) + ii * of_real(2*pi)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1195
                else if Im(Ln w + Ln z) > pi then
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1196
                  (Ln(w) + Ln(z)) - ii * of_real(2*pi)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1197
                else Ln(w) + Ln(z))"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1198
  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
  1199
  using assms mpi_less_Im_Ln [of w] mpi_less_Im_Ln [of z]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1200
  by (auto simp: of_real_numeral exp_add exp_diff sin_double cos_double exp_Euler intro!: Ln_unique)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1201
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1202
corollary Ln_times_simple:
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1203
    "\<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
  1204
         \<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
  1205
  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
  1206
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1207
corollary Ln_times_of_real:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1208
    "\<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
  1209
  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
  1210
  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
  1211
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1212
corollary Ln_divide_of_real:
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1213
    "\<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
  1214
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
  1215
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
  1216
         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
  1217
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1218
lemma Ln_minus:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1219
  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
  1220
    shows "Ln(-z) = (if Im(z) \<le> 0 \<and> ~(Re(z) < 0 \<and> Im(z) = 0)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1221
                     then Ln(z) + ii * pi
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1222
                     else Ln(z) - ii * pi)" (is "_ = ?rhs")
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1223
  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
  1224
        Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1225
    by (auto simp: of_real_numeral exp_add exp_diff exp_Euler intro!: Ln_unique)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1226
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1227
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
  1228
  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
  1229
    shows "Ln (inverse z) =
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1230
            (if (Im(z) = 0 \<longrightarrow> 0 < Re z)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1231
             then -(Ln z)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1232
             else -(Ln z) + \<i> * 2 * complex_of_real pi)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1233
proof (cases "(Im(z) = 0 \<longrightarrow> 0 < Re z)")
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1234
  case True then show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1235
    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
  1236
next
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1237
  case False
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1238
  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
  1239
    using assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1240
    apply auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1241
    by (metis cnj.code complex_cnj_cnj not_less_iff_gr_or_eq zero_complex.simps(1) 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
  1242
  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
  1243
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1244
  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
  1245
    using assms z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1246
    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
  1247
    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
  1248
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1249
  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
  1250
    apply (subst Ln_inverse)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1251
    using z assms by auto
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1252
  also have "... = - (Ln z) + \<i> * 2 * complex_of_real pi"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1253
    apply (subst Ln_minus [OF assms])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1254
    using assms z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1255
    apply simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1256
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1257
  finally show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1258
    using assms z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1259
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1260
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1261
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1262
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
  1263
  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
  1264
    shows  "Ln(ii * z) = (if 0 \<le> Re(z) | Im(z) < 0
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1265
                          then Ln(z) + ii * of_real pi/2
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1266
                          else Ln(z) - ii * of_real(3 * pi/2))"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1267
  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
  1268
        Im_Ln_eq_pi [of z] Im_Ln_pos_lt [of z] Re_Ln_pos_le [of z]
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1269
  by (auto simp: of_real_numeral Ln_times)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1270
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1271
lemma Ln_of_nat: "0 < n \<Longrightarrow> Ln (of_nat n) = of_real (ln (of_nat n))"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1272
  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
  1273
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
  1274
lemma Ln_of_nat_over_of_nat:
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1275
  assumes "m > 0" "n > 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1276
  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
  1277
proof -
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1278
  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
  1279
  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
  1280
    by (simp add: Ln_of_real[symmetric])
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1281
  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
  1282
    by (simp add: ln_div)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1283
  finally show ?thesis .
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1284
qed
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1285
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1286
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1287
subsection\<open>Relation between Ln and Arg, and hence continuity of Arg\<close>
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1288
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
  1289
lemma Arg_Ln:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1290
  assumes "0 < Arg z" shows "Arg z = Im(Ln(-z)) + pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1291
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
  1292
  case True
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1293
  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
  1294
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1295
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1296
  case False
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1297
  then have "z / of_real(norm z) = exp(ii * of_real(Arg z))"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1298
    using Arg [of z]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1299
    by (metis abs_norm_cancel nonzero_mult_divide_cancel_left norm_of_real zero_less_norm_iff)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1300
  then have "- z / of_real(norm z) = exp (\<i> * (of_real (Arg z) - pi))"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1301
    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
  1302
    by (auto simp: exp_diff algebra_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1303
  then have "ln (- z / of_real(norm z)) = ln (exp (\<i> * (of_real (Arg z) - pi)))"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1304
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1305
  also have "... = \<i> * (of_real(Arg z) - pi)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1306
    using Arg [of z] assms pi_not_less_zero
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1307
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1308
  finally have "Arg z =  Im (Ln (- z / of_real (cmod z))) + pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1309
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1310
  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
  1311
    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
  1312
  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
  1313
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1314
  finally show ?thesis .
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1315
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1316
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
  1317
lemma continuous_at_Arg:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1318
  assumes "z \<in> \<real> \<Longrightarrow> Re z < 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1319
    shows "continuous (at z) Arg"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1320
proof -
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1321
  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
  1322
    by (rule Complex.isCont_Im isCont_Ln' continuous_intros | simp add: assms complex_is_Real_iff)+
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1323
  then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1324
    apply (simp add: continuous_at)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1325
    apply (rule Lim_transform_within_open [of "-{z. z \<in> \<real> & 0 \<le> Re z}" _ "\<lambda>z. Im(Ln(-z)) + pi"])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1326
    apply (simp add: closed_def [symmetric] closed_Collect_conj closed_complex_Reals closed_halfspace_Re_ge)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1327
    apply (simp_all add: assms not_le Arg_Ln [OF Arg_gt_0])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1328
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1329
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1330
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1331
text\<open>Relation between Arg and arctangent in upper halfplane\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1332
lemma Arg_arctan_upperhalf:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1333
  assumes "0 < Im z"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1334
    shows "Arg z = pi/2 - arctan(Re z / Im z)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1335
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
  1336
  case True with assms show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1337
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1338
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1339
  case False
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1340
  show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1341
    apply (rule Arg_unique [of "norm z"])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1342
    using False assms arctan [of "Re z / Im z"] pi_ge_two pi_half_less_two
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1343
    apply (auto simp: exp_Euler cos_diff sin_diff)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1344
    using norm_complex_def [of z, symmetric]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1345
    apply (simp add: of_real_numeral sin_of_real cos_of_real sin_arctan cos_arctan field_simps real_sqrt_divide)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1346
    apply (metis complex_eq mult.assoc ring_class.ring_distribs(2))
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1347
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1348
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1349
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
  1350
lemma Arg_eq_Im_Ln:
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1351
  assumes "0 \<le> Im z" "0 < Re z"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1352
    shows "Arg z = Im (Ln z)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1353
proof (cases "z = 0 \<or> Im z = 0")
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1354
  case True then show ?thesis
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1355
    using assms Arg_eq_0 complex_is_Real_iff
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1356
    apply auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1357
    by (metis Arg_eq_0_pi Arg_eq_pi Im_Ln_eq_0 Im_Ln_eq_pi less_numeral_extra(3) zero_complex.simps(1))
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1358
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
  1359
  case False
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1360
  then have "Arg z > 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1361
    using Arg_gt_0 complex_is_Real_iff by blast
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1362
  then show ?thesis
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1363
    using assms False
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1364
    by (subst Arg_Ln) (auto simp: Ln_minus)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1365
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1366
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
  1367
lemma continuous_within_upperhalf_Arg:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1368
  assumes "z \<noteq> 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1369
    shows "continuous (at z within {z. 0 \<le> Im z}) Arg"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1370
proof (cases "z \<in> \<real> & 0 \<le> Re z")
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1371
  case False then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1372
    using continuous_at_Arg continuous_at_imp_continuous_within by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1373
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1374
  case True
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1375
  then have z: "z \<in> \<real>" "0 < Re z"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1376
    using assms  by (auto simp: complex_is_Real_iff complex_neq_0)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1377
  then have [simp]: "Arg z = 0" "Im (Ln z) = 0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1378
    by (auto simp: Arg_eq_0 Im_Ln_eq_0 assms complex_is_Real_iff)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1379
  show ?thesis
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1380
  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
  1381
    fix e::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1382
    assume "0 < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1383
    moreover have "continuous (at z) (\<lambda>x. Im (Ln x))"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1384
      using z  by (rule continuous_intros | simp)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1385
    ultimately
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1386
    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
  1387
      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
  1388
    { fix x
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1389
      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
  1390
      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
  1391
        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
  1392
      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
  1393
        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
  1394
    }
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1395
    then show "\<exists>d>0. \<forall>x. 0 \<le> Im x \<longrightarrow> x \<noteq> z \<and> cmod (x - z) < d \<longrightarrow> \<bar>Arg x\<bar> < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1396
      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
  1397
      using z d
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1398
      apply (auto simp: Arg_eq_Im_Ln)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1399
      done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1400
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1401
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1402
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1403
lemma continuous_on_upperhalf_Arg: "continuous_on ({z. 0 \<le> Im z} - {0}) Arg"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1404
  apply (auto simp: continuous_on_eq_continuous_within)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1405
  by (metis Diff_subset continuous_within_subset continuous_within_upperhalf_Arg)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1406
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
  1407
lemma open_Arg_less_Int:
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1408
  assumes "0 \<le> s" "t \<le> 2*pi"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1409
    shows "open ({y. s < Arg y} \<inter> {y. Arg y < t})"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1410
proof -
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1411
  have 1: "continuous_on (UNIV - {z \<in> \<real>. 0 \<le> Re z}) Arg"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1412
    using continuous_at_Arg continuous_at_imp_continuous_within
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1413
    by (auto simp: continuous_on_eq_continuous_within set_diff_eq)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1414
  have 2: "open (UNIV - {z \<in> \<real>. 0 \<le> Re z})"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1415
    by (simp add: closed_Collect_conj closed_complex_Reals closed_halfspace_Re_ge open_Diff)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1416
  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
  1417
    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
  1418
    by (metis greaterThan_def lessThan_def open_Int)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1419
  moreover have "{y. s < Arg y} \<inter> {y. Arg y < t} \<subseteq> - {z \<in> \<real>. 0 \<le> Re z}"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1420
    using assms
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1421
    by (auto simp: Arg_real)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1422
  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
  1423
    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
  1424
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1425
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1426
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1427
lemma open_Arg_gt: "open {z. t < Arg z}"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1428
proof (cases "t < 0")
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1429
  case True then have "{z. t < Arg z} = UNIV"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1430
    using Arg_ge_0 less_le_trans by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1431
  then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1432
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1433
next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1434
  case False then show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1435
    using open_Arg_less_Int [of t "2*pi"] Arg_lt_2pi
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1436
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1437
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1438
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1439
lemma closed_Arg_le: "closed {z. Arg z \<le> t}"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1440
  using open_Arg_gt [of t]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1441
  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
  1442
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1443
subsection\<open>Complex Powers\<close>
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1444
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1445
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
  1446
  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
  1447
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1448
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
  1449
  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
  1450
  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
  1451
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60420
diff changeset
  1452
lemma powr_add_complex:
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
  1453
  fixes w::complex shows "w powr (z1 + z2) = w powr z1 * w powr z2"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1454
  by (simp add: powr_def algebra_simps exp_add)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1455
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60420
diff changeset
  1456
lemma powr_minus_complex:
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
  1457
  fixes w::complex shows  "w powr (-z) = inverse(w powr z)"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1458
  by (simp add: powr_def exp_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
  1459
60809
457abb82fb9e the Cauchy integral theorem and related material
paulson <lp15@cam.ac.uk>
parents: 60420
diff changeset
  1460
lemma powr_diff_complex:
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
  1461
  fixes w::complex shows  "w powr (z1 - z2) = w powr z1 / w powr z2"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1462
  by (simp add: powr_def algebra_simps exp_diff)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1463
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1464
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
  1465
  apply (simp add: powr_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1466
  using Im_Ln_eq_0 complex_is_Real_iff norm_complex_def
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1467
  by auto
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1468
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1469
lemma cnj_powr:
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1470
  assumes "Im a = 0 \<Longrightarrow> Re a \<ge> 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1471
  shows   "cnj (a powr b) = cnj a powr cnj b"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1472
proof (cases "a = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1473
  case False
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1474
  with assms have "Im a = 0 \<Longrightarrow> Re a > 0" by (auto simp: complex_eq_iff)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1475
  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
  1476
qed simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1477
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
  1478
lemma powr_real_real:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1479
    "\<lbrakk>w \<in> \<real>; z \<in> \<real>; 0 < Re w\<rbrakk> \<Longrightarrow> w powr z = exp(Re z * ln(Re w))"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1480
  apply (simp add: powr_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1481
  by (metis complex_eq complex_is_Real_iff diff_0 diff_0_right diff_minus_eq_add exp_ln exp_not_eq_zero
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1482
       exp_of_real Ln_of_real mult_zero_right of_real_0 of_real_mult)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1483
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1484
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
  1485
  fixes x::real and y::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
  1486
  shows "0 < x \<Longrightarrow> of_real x powr (of_real y::complex) = of_real (x powr y)"
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1487
  by (simp add: powr_def) (metis exp_of_real of_real_mult 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
  1488
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1489
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
  1490
    "\<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
  1491
     \<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
  1492
  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
  1493
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1494
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
  1495
    "\<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
  1496
           \<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
  1497
  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
  1498
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1499
lemma powr_neg_real_complex:
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1500
  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
  1501
proof (cases "x = 0")
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1502
  assume x: "x \<noteq> 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1503
  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
  1504
  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
  1505
    by (simp add: Ln_minus Ln_of_real)
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1506
  also from x assms have "exp (a * ...) = cis pi powr (of_real (sgn x) * a) * of_real x powr a"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1507
    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
  1508
  also note cis_pi
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1509
  finally show ?thesis by simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1510
qed simp_all
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1511
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
  1512
lemma has_field_derivative_powr:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1513
    "(Im z = 0 \<Longrightarrow> 0 < Re z)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1514
     \<Longrightarrow> ((\<lambda>z. z powr s) has_field_derivative (s * z powr (s - 1))) (at z)"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1515
  apply (cases "z=0", auto)
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1516
  apply (simp add: powr_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1517
  apply (rule DERIV_transform_at [where d = "norm z" and f = "\<lambda>z. exp (s * Ln z)"])
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1518
  apply (auto simp: dist_complex_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1519
  apply (intro derivative_eq_intros | simp add: 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
  1520
  apply (simp add: field_simps exp_diff)
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1521
  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
  1522
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1523
lemma has_field_derivative_powr_complex':
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1524
  assumes "Im z \<noteq> 0 \<or> Re z > 0"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1525
  shows "((\<lambda>z. z powr r :: complex) has_field_derivative r * z powr (r - 1)) (at z)"
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1526
proof (subst DERIV_cong_ev[OF refl _ refl])
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1527
  from assms have "eventually (\<lambda>z. z \<noteq> 0) (nhds z)" by (intro t1_space_nhds) auto
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
  1528
  thus "eventually (\<lambda>z. z powr r = exp (r * Ln z)) (nhds z)"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1529
    unfolding powr_def by eventually_elim simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1530
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
  1531
  have "((\<lambda>z. exp (r * Ln z)) has_field_derivative exp (r * Ln z) * (inverse z * r)) (at z)"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1532
    using assms by (auto intro!: derivative_eq_intros has_field_derivative_powr)
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
  1533
  also have "exp (r * Ln z) * (inverse z * r) = r * z powr (r - 1)"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1534
    unfolding powr_def by (simp add: assms exp_diff field_simps)
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
  1535
  finally show "((\<lambda>z. exp (r * Ln z)) has_field_derivative r * z powr (r - 1)) (at z)"
61524
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1536
    by simp
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1537
qed
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1538
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1539
declare has_field_derivative_powr_complex'[THEN DERIV_chain2, derivative_intros]
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1540
f2e51e704a96 added many small lemmas about setsum/setprod/powr/...
eberlm
parents: 61518
diff changeset
  1541
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
  1542
lemma has_field_derivative_powr_right:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1543
    "w \<noteq> 0 \<Longrightarrow> ((\<lambda>z. w powr z) has_field_derivative Ln w * w powr z) (at z)"
60020
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  1544
  apply (simp add: powr_def)
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1545
  apply (intro derivative_eq_intros | simp add: 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
  1546
  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
  1547
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1548
lemma complex_differentiable_powr_right:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1549
    "w \<noteq> 0 \<Longrightarrow> (\<lambda>z. w powr z) complex_differentiable (at z)"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1550
using complex_differentiable_def has_field_derivative_powr_right 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
  1551
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1552
lemma holomorphic_on_powr_right:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1553
    "f holomorphic_on s \<Longrightarrow> w \<noteq> 0 \<Longrightarrow> (\<lambda>z. w powr (f z)) holomorphic_on s"
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1554
    unfolding holomorphic_on_def
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1555
    using DERIV_chain' complex_differentiable_def has_field_derivative_powr_right 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
  1556
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1557
lemma norm_powr_real_powr:
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  1558
  "w \<in> \<real> \<Longrightarrow> 0 < Re w \<Longrightarrow> norm(w powr z) = Re w powr Re 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
  1559
  by (auto simp add: norm_powr_real powr_def Im_Ln_eq_0 complex_is_Real_iff in_Reals_norm)
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
  1560
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1561
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1562
subsection\<open>Some Limits involving Logarithms\<close>
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1563
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1564
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
  1565
  fixes s::complex
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1566
  assumes "0 < Re s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1567
    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
  1568
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
  1569
  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
  1570
  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
  1571
  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
  1572
  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
  1573
    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
  1574
      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
  1575
  next
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1576
    fix x::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1577
    assume x: "2 / (e * (Re s)\<^sup>2) \<le> x"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1578
    then have "x>0"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1579
    using e assms
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1580
      by (metis less_le_trans mult_eq_0_iff mult_pos_pos pos_less_divide_eq power2_eq_square
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1581
                zero_less_numeral)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1582
    then show "0 < e * 2 + (e * Re s * 2 - 2) * x + e * (Re s)\<^sup>2 * x\<^sup>2"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1583
      using e assms x
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1584
      apply (auto simp: field_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1585
      apply (rule_tac y = "e * (x\<^sup>2 * (Re s)\<^sup>2)" in le_less_trans)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1586
      apply (auto simp: power2_eq_square field_simps add_pos_pos)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1587
      done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1588
  qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1589
  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
  1590
    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
  1591
  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
  1592
    using assms
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1593
    by (force intro: less_le_trans [OF _ exp_lower_taylor_quadratic])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1594
  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
  1595
    using e   by (auto simp: field_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1596
  with e show "\<exists>no. \<forall>n\<ge>no. norm (Ln (of_nat n) / of_nat n powr s) < e"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1597
    apply (auto simp: norm_divide norm_powr_real divide_simps)
61942
f02b26f7d39d prefer symbols for "floor", "ceiling";
wenzelm
parents: 61808
diff changeset
  1598
    apply (rule_tac x="nat \<lceil>exp xo\<rceil>" in exI)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1599
    apply clarify
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1600
    apply (drule_tac x="ln n" in spec)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1601
    apply (auto simp: exp_less_mono nat_ceiling_le_eq not_le)
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1602
    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
  1603
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1604
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1605
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1606
lemma lim_Ln_over_n: "((\<lambda>n. Ln(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
  1607
  using lim_Ln_over_power [of 1]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1608
  by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1609
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
  1610
lemma Ln_Reals_eq: "x \<in> \<real> \<Longrightarrow> Re x > 0 \<Longrightarrow> Ln x = of_real (ln (Re x))"
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
  1611
  using Ln_of_real 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
  1612
61070
b72a990adfe2 prefer symbols;
wenzelm
parents: 60809
diff changeset
  1613
lemma powr_Reals_eq: "x \<in> \<real> \<Longrightarrow> Re x > 0 \<Longrightarrow> x powr complex_of_real y = of_real (x powr y)"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1614
  by (simp add: powr_of_real)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1615
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1616
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
  1617
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1618
  assumes "0 < s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1619
    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
  1620
  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
  1621
  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
  1622
  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
  1623
          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
  1624
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1625
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1626
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
  1627
  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
  1628
  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
  1629
  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
  1630
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1631
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1632
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
  1633
  assumes "0 < Re s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1634
    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
  1635
proof -
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1636
  have "\<forall>n>0. 3 \<le> n \<longrightarrow> 1 \<le> ln (real_of_nat n)"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1637
    using ln3_gt_1
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1638
    by (force intro: order_trans [of _ "ln 3"] ln3_gt_1)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
  1639
  moreover have "(\<lambda>n. cmod (Ln (of_nat n) / of_nat n powr s)) \<longlonglongrightarrow> 0"
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1640
    using lim_Ln_over_power [OF assms]
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1641
    by (metis tendsto_norm_zero_iff)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1642
  ultimately show ?thesis
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1643
    apply (auto intro!: Lim_null_comparison [where g = "\<lambda>n. norm (Ln(of_nat n) / of_nat n powr s)"])
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1644
    apply (auto simp: norm_divide divide_simps eventually_sequentially)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1645
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1646
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1647
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1648
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
  1649
  fixes s :: real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1650
  assumes "0 < s"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1651
    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
  1652
  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
  1653
  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
  1654
  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
  1655
  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
  1656
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1657
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1658
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
  1659
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
  1660
  fix r::real
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1661
  assume "0 < r"
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1662
  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
  1663
    by simp
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1664
  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
  1665
    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
  1666
    by auto
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1667
  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
  1668
    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
  1669
  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
  1670
    by (metis exp_gt_zero less_trans ln_exp ln_less_cancel_iff)
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1671
  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
  1672
    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
  1673
  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
  1674
    using neq0_conv by fastforce
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1675
  ultimately show "\<exists>no. \<forall>n. Ln (of_nat n) \<noteq> 0 \<longrightarrow> no \<le> n \<longrightarrow> 1 < r * cmod (Ln (of_nat n))"
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1676
    using n \<open>0 < r\<close>
60150
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1677
    apply (rule_tac x=n in exI)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1678
    apply (auto simp: divide_simps)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1679
    apply (erule less_le_trans, auto)
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1680
    done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1681
qed
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1682
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
  1683
lemma lim_1_over_ln: "((\<lambda>n. 1 / ln(real_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
  1684
  using lim_1_over_Ln [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
  1685
  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
  1686
  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
  1687
  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
  1688
  done
bd773c47ad0b New material about complex transcendental functions (especially Ln, Arg) and polynomials
paulson <lp15@cam.ac.uk>
parents: 60141
diff changeset
  1689
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
  1690
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1691
subsection\<open>Relation between Square Root and exp/ln, hence its derivative\<close>
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1692
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1693
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
  1694
  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
  1695
    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
  1696
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1697
  have "(exp (Ln z / 2))\<^sup>2 = (exp (Ln z))"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1698
    by (metis exp_double nonzero_mult_divide_cancel_left times_divide_eq_right zero_neq_numeral)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1699
  also have "... = z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1700
    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
  1701
  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
  1702
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1703
  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
  1704
    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
  1705
    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
  1706
    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
  1707
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1708
  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
  1709
    by simp
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1710
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1711
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1712
lemma csqrt_inverse:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1713
  assumes "Im(z) = 0 \<Longrightarrow> 0 < Re z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1714
    shows "csqrt (inverse z) = inverse (csqrt z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1715
proof (cases "z=0", simp)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1716
  assume "z \<noteq> 0 "
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1717
  then show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1718
    using assms
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1719
    by (simp add: csqrt_exp_Ln Ln_inverse exp_minus)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1720
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1721
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1722
lemma cnj_csqrt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1723
  assumes "Im z = 0 \<Longrightarrow> 0 \<le> Re(z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1724
    shows "cnj(csqrt z) = csqrt(cnj z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1725
proof (cases "z=0", simp)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1726
  assume 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
  1727
  then have "Im z = 0 \<Longrightarrow> 0 < Re(z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1728
    using assms cnj.code complex_cnj_zero_iff by fastforce
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1729
  then show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1730
   using z by (simp add: csqrt_exp_Ln cnj_Ln exp_cnj)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1731
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1732
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1733
lemma has_field_derivative_csqrt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1734
  assumes "Im z = 0 \<Longrightarrow> 0 < Re(z)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1735
    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
  1736
proof -
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1737
  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
  1738
    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
  1739
  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
  1740
    by (simp add: divide_simps)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1741
  show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1742
    apply (rule DERIV_transform_at [where f = "\<lambda>z. exp(Ln(z) / 2)" and d = "norm z"])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1743
    apply (intro derivative_eq_intros | simp add: assms)+
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1744
    apply (rule *)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1745
    using z
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1746
    apply (auto simp: field_simps csqrt_exp_Ln [symmetric])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1747
    apply (metis power2_csqrt power2_eq_square)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1748
    apply (metis csqrt_exp_Ln dist_0_norm less_irrefl)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1749
    done
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1750
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1752
lemma complex_differentiable_at_csqrt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1753
    "(Im z = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> csqrt complex_differentiable at z"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1754
  using complex_differentiable_def has_field_derivative_csqrt by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1755
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1756
lemma complex_differentiable_within_csqrt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1757
    "(Im z = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> csqrt complex_differentiable (at z within s)"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1758
  using complex_differentiable_at_csqrt complex_differentiable_within_subset by blast
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1759
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1760
lemma continuous_at_csqrt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1761
    "(Im z = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> continuous (at z) csqrt"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1762
  by (simp add: complex_differentiable_within_csqrt complex_differentiable_imp_continuous_at)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1763
59862
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1764
corollary isCont_csqrt' [simp]:
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1765
   "\<lbrakk>isCont f z; Im(f z) = 0 \<Longrightarrow> 0 < Re(f z)\<rbrakk> \<Longrightarrow> isCont (\<lambda>x. csqrt (f x)) z"
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1766
  by (blast intro: isCont_o2 [OF _ continuous_at_csqrt])
44b3f4fa33ca New material and binomial fix
paulson <lp15@cam.ac.uk>
parents: 59751
diff changeset
  1767
59751
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1768
lemma continuous_within_csqrt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1769
    "(Im z = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> continuous (at z within s) csqrt"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1770
  by (simp add: complex_differentiable_imp_continuous_at complex_differentiable_within_csqrt)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1771
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1772
lemma continuous_on_csqrt [continuous_intros]:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1773
    "(\<And>z. z \<in> s \<and> Im z = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> continuous_on s csqrt"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1774
  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
  1775
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1776
lemma holomorphic_on_csqrt:
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1777
    "(\<And>z. z \<in> s \<and> Im z = 0 \<Longrightarrow> 0 < Re(z)) \<Longrightarrow> csqrt holomorphic_on s"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1778
  by (simp add: complex_differentiable_within_csqrt holomorphic_on_def)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1779
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1780
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
  1781
    "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
  1782
  using open_Compl
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1783
  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
  1784
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1785
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
  1786
    "continuous (at z within (\<real> \<inter> {w. 0 \<le> Re(w)})) csqrt"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1787
proof (cases "Im z = 0 --> 0 < Re(z)")
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1788
  case True then show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1789
    by (blast intro: continuous_within_csqrt)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1790
next
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1791
  case False
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1792
  then have "Im z = 0" "Re z < 0 \<or> z = 0"
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1793
    using False cnj.code complex_cnj_zero_iff by auto force
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1794
  then show ?thesis
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1795
    apply (auto simp: continuous_within_closed_nontrivial [OF closed_Real_halfspace_Re_ge])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1796
    apply (auto simp: continuous_within_eps_delta norm_conv_dist [symmetric])
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1797
    apply (rule_tac x="e^2" in exI)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1798
    apply (auto simp: Reals_def)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1799
by (metis linear not_less real_sqrt_less_iff real_sqrt_pow2_iff real_sqrt_power)
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1800
qed
916c0f6c83e3 New material for complex sin, cos, tan, Ln, also some reorganisation
paulson <lp15@cam.ac.uk>
parents: 59746
diff changeset
  1801
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1802
subsection\<open>Complex arctangent\<close>
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1803
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1804
text\<open>branch cut gives standard bounds in real case.\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1805
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1806
definition Arctan :: "complex \<Rightarrow> complex" where
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1807
    "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
  1808
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1809
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
  1810
  by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1811
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1812
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
  1813
  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
  1814
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1815
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
  1816
  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
  1817
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1818
lemma tan_Arctan:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1819
  assumes "z\<^sup>2 \<noteq> -1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1820
    shows [simp]:"tan(Arctan z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1821
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1822
  have "1 + \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1823
    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
  1824
  moreover
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1825
  have "1 - \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1826
    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
  1827
  ultimately
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1828
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1829
    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
  1830
                  divide_simps power2_eq_square [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1831
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1832
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1833
lemma Arctan_tan [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1834
  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
  1835
    shows "Arctan(tan z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1836
proof -
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1837
  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
  1838
    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
  1839
  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
  1840
    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
  1841
  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
  1842
    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
  1843
  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
  1844
    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
  1845
  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
  1846
    by (simp add: exp_eq_1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1847
  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
  1848
    by (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1849
  also have "... \<longleftrightarrow> False"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1850
    using assms ge_pi2
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1851
    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
  1852
    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
  1853
  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
  1854
    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
  1855
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1856
    using assms *
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1857
    apply (simp add: Arctan_def tan_def sin_exp_eq cos_exp_eq exp_minus divide_simps
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1858
                     ii_times_eq_iff power2_eq_square [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1859
    apply (rule Ln_unique)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1860
    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
  1861
    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
  1862
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1863
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1864
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1865
lemma
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1866
  assumes "Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1"
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1867
  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
  1868
    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
  1869
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1870
  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
  1871
    using assms
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60020
diff changeset
  1872
    by (metis abs_one complex_i_mult_minus diff_0_right diff_minus_eq_add ii.simps(1) ii.simps(2)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1873
              less_irrefl minus_diff_eq mult.right_neutral right_minus_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1874
  have "z \<noteq> -\<i>" using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1875
    by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1876
  then have zz: "1 + z * z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1877
    by (metis abs_one assms i_squared ii.simps less_irrefl minus_unique square_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1878
  have nz1: "1 - \<i>*z \<noteq> 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1879
    using assms by (force simp add: ii_times_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1880
  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
  1881
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1882
    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
  1883
              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
  1884
  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
  1885
    using nz1 nz2 by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1886
  have *: "Im ((1 - \<i>*z) / (1 + \<i>*z)) = 0 \<Longrightarrow> 0 < Re ((1 - \<i>*z) / (1 + \<i>*z))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1887
    apply (simp add: divide_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1888
    apply (simp add: divide_simps split: split_if_asm)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1889
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1890
    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
  1891
    done
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1892
  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
  1893
    unfolding Arctan_def divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1894
    using mpi_less_Im_Ln [OF nzi]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1895
    by (auto simp: abs_if intro: Im_Ln_less_pi * [unfolded divide_complex_def])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1896
  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
  1897
    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
  1898
    apply (rule DERIV_cong)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1899
    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
  1900
    using nz0 nz1 zz
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1901
    apply (simp add: divide_simps power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1902
    apply (auto simp: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1903
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1904
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1905
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1906
lemma complex_differentiable_at_Arctan: "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan complex_differentiable at z"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1907
  using has_field_derivative_Arctan
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1908
  by (auto simp: complex_differentiable_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1909
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1910
lemma complex_differentiable_within_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1911
    "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan complex_differentiable (at z within s)"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1912
  using complex_differentiable_at_Arctan complex_differentiable_at_within by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1913
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1914
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
  1915
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
  1916
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1917
lemma continuous_at_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1918
    "(Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> continuous (at z) Arctan"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1919
  by (simp add: complex_differentiable_imp_continuous_at complex_differentiable_within_Arctan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1920
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1921
lemma continuous_within_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1922
    "(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
  1923
  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
  1924
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1925
lemma continuous_on_Arctan [continuous_intros]:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1926
    "(\<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
  1927
  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
  1928
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1929
lemma holomorphic_on_Arctan:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1930
    "(\<And>z. z \<in> s \<Longrightarrow> Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan holomorphic_on s"
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1931
  by (simp add: complex_differentiable_within_Arctan holomorphic_on_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1932
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1933
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  1934
subsection \<open>Real arctangent\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1935
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1936
lemma norm_exp_ii_times [simp]: "norm (exp(\<i> * of_real y)) = 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1937
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1938
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1939
lemma norm_exp_imaginary: "norm(exp z) = 1 \<Longrightarrow> Re z = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1940
  by (simp add: complex_norm_eq_1_exp)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1941
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1942
lemma Im_Arctan_of_real [simp]: "Im (Arctan (of_real x)) = 0"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1943
  unfolding Arctan_def divide_complex_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1944
  apply (simp add: complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1945
  apply (rule norm_exp_imaginary)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1946
  apply (subst exp_Ln, auto)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1947
  apply (simp_all add: cmod_def complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1948
  apply (auto simp: divide_simps)
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 61524
diff changeset
  1949
  apply (metis power_one sum_power2_eq_zero_iff zero_neq_one, algebra)
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1950
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1951
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1952
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
  1953
proof (rule arctan_unique)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1954
  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
  1955
    apply (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1956
    apply (rule Im_Ln_less_pi)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1957
    apply (auto simp: Im_complex_div_lemma)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1958
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1959
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1960
  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
  1961
    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
  1962
  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
  1963
    using mpi_less_Im_Ln [OF *]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1964
    by (simp add: Arctan_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1965
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1966
  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
  1967
    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
  1968
    apply (simp add: field_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1969
    by (simp add: power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1970
  also have "... = x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1971
    apply (subst tan_Arctan, auto)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1972
    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
  1973
  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
  1974
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1975
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1976
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
  1977
  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
  1978
  by (simp add: complex_eq_iff)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1979
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1980
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
  1981
  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
  1982
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1983
declare arctan_one [simp]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1984
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1985
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
  1986
  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
  1987
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  1988
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
  1989
  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
  1990
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1991
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
  1992
  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
  1993
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1994
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
  1995
  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
  1996
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  1997
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
  1998
  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
  1999
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2000
lemma arctan_add_raw:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2001
  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
  2002
    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
  2003
proof (rule arctan_unique [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2004
  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
  2005
    using assms by linarith+
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2006
  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
  2007
    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
  2008
    by (simp add: arctan tan_add)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2009
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2010
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2011
lemma arctan_inverse:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2012
  assumes "0 < x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2013
    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
  2014
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2015
  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
  2016
    by (simp add: arctan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2017
  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
  2018
    by (simp add: tan_cot)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2019
  also have "... = pi/2 - arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2020
  proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2021
    have "0 < pi - arctan x"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2022
    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
  2023
    with assms show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2024
      by (simp add: Transcendental.arctan_tan)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2025
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2026
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2027
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2028
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2029
lemma arctan_add_small:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2030
  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
  2031
    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
  2032
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
  2033
  case True then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2034
    by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2035
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2036
  case False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2037
  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
  2038
    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
  2039
    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
  2040
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2041
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2042
    apply (rule arctan_add_raw)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2043
    using * by linarith
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2044
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2045
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2046
lemma abs_arctan_le:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2047
  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
  2048
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2049
  { fix w::complex and z::complex
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2050
    assume *: "w \<in> \<real>" "z \<in> \<real>"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2051
    have "cmod (Arctan w - Arctan z) \<le> 1 * cmod (w-z)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2052
      apply (rule complex_differentiable_bound [OF convex_Reals, of Arctan _ 1])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2053
      apply (rule has_field_derivative_at_within [OF has_field_derivative_Arctan])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2054
      apply (force simp add: Reals_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2055
      apply (simp add: norm_divide divide_simps in_Reals_norm complex_is_Real_iff power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2056
      using * by auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2057
  }
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2058
  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
  2059
    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
  2060
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2061
    by (simp add: Arctan_of_real)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2062
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2063
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2064
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
  2065
  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
  2066
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2067
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
  2068
  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
  2069
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2070
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2071
subsection\<open>Inverse Sine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2072
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2073
definition Arcsin :: "complex \<Rightarrow> complex" where
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2074
   "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
  2075
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2076
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
  2077
  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
  2078
  apply auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2079
  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
  2080
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2081
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
  2082
  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
  2083
  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
  2084
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2085
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
  2086
  by (simp add: Arcsin_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2087
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2088
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
  2089
  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
  2090
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2091
lemma isCont_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2092
  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
  2093
    shows "isCont Arcsin z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2094
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2095
  have rez: "Im (1 - z\<^sup>2) = 0 \<Longrightarrow> 0 < Re (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2096
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2097
    by (auto simp: Re_power2 Im_power2 abs_square_less_1 add_pos_nonneg algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2098
  have cmz: "(cmod z)\<^sup>2 < 1 + cmod (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2099
    by (blast intro: assms cmod_square_less_1_plus)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2100
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2101
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2102
    apply (simp add: Arcsin_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2103
    apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2104
    apply (erule rez)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2105
    apply (auto simp: Re_power2 Im_power2 abs_square_less_1 [symmetric] real_less_rsqrt algebra_simps split: split_if_asm)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2106
    apply (simp add: norm_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2107
    using cmod_power2 [of z, symmetric] cmz
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2108
    apply (simp add: real_less_rsqrt)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2109
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2110
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2111
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2112
lemma isCont_Arcsin' [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2113
  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
  2114
  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
  2115
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2116
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
  2117
proof -
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2118
  have "\<i>*z*2 + csqrt (1 - z\<^sup>2)*2 = 0 \<longleftrightarrow> (\<i>*z)*2 + csqrt (1 - z\<^sup>2)*2 = 0"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61806
diff changeset
  2119
    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
  2120
  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
  2121
    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
  2122
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2123
    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
  2124
    apply (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2125
    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
  2126
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2127
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2128
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2129
lemma Re_eq_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2130
    "\<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
  2131
      Re ((exp (\<i>*z) + inverse (exp (\<i>*z))) / 2) = 0 \<and> 0 \<le> Im ((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
  2132
  apply (simp add: cos_ii_times [symmetric] Re_cos Im_cos abs_if del: eq_divide_eq_numeral1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2133
  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
  2134
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2135
lemma Re_less_pihalf_lemma:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2136
  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
  2137
    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
  2138
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2139
  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
  2140
    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
  2141
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2142
    by (simp add: cos_ii_times [symmetric] Re_cos Im_cos add_pos_pos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2143
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2144
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2145
lemma Arcsin_sin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2146
    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
  2147
      shows "Arcsin(sin z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2148
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2149
  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
  2150
    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
  2151
  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
  2152
    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
  2153
  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
  2154
    apply (subst csqrt_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2155
    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
  2156
    apply auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2157
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2158
  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
  2159
    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
  2160
  also have "... = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2161
    apply (subst Complex_Transcendental.Ln_exp)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2162
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2163
    apply (auto simp: abs_if simp del: eq_divide_eq_numeral1 split: split_if_asm)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2164
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2165
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2166
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2167
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2168
lemma Arcsin_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2169
    "\<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
  2170
  by (metis Arcsin_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2171
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2172
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
  2173
  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
  2174
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2175
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
  2176
  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
  2177
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2178
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
  2179
  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
  2180
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2181
lemma has_field_derivative_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2182
  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
  2183
    shows "(Arcsin has_field_derivative inverse(cos(Arcsin z))) (at z)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2184
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2185
  have "(sin (Arcsin z))\<^sup>2 \<noteq> 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2186
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2187
    apply atomize
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2188
    apply (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2189
    apply (metis abs_minus_cancel abs_one abs_power2 numeral_One numeral_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2190
    by (metis abs_minus_cancel abs_one abs_power2 one_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2191
  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
  2192
    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
  2193
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2194
    apply (rule has_complex_derivative_inverse_basic [OF DERIV_sin])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2195
    apply (auto intro: isCont_Arcsin open_ball [of z 1] assms)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2196
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2197
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2198
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2199
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
  2200
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
  2201
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2202
lemma complex_differentiable_at_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2203
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin complex_differentiable at z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2204
  using complex_differentiable_def has_field_derivative_Arcsin by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2205
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2206
lemma complex_differentiable_within_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2207
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin complex_differentiable (at z within s)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2208
  using complex_differentiable_at_Arcsin complex_differentiable_within_subset by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2209
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2210
lemma continuous_within_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2211
    "(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
  2212
  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
  2213
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2214
lemma continuous_on_Arcsin [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2215
    "(\<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
  2216
  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
  2217
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2218
lemma holomorphic_on_Arcsin: "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arcsin holomorphic_on s"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2219
  by (simp add: complex_differentiable_within_Arcsin holomorphic_on_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2220
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2221
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2222
subsection\<open>Inverse Cosine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2223
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2224
definition Arccos :: "complex \<Rightarrow> complex" where
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2225
   "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
  2226
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2227
lemma Arccos_range_lemma: "\<bar>Re z\<bar> < 1 \<Longrightarrow> 0 < Im(z + \<i> * csqrt(1 - z\<^sup>2))"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2228
  using Arcsin_range_lemma [of "-z"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2229
  by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2230
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2231
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
  2232
  using Arcsin_body_lemma [of z]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2233
  by (metis complex_i_mult_minus diff_add_cancel minus_diff_eq minus_unique mult.assoc mult.left_commute
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2234
           power2_csqrt power2_eq_square zero_neq_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2235
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2236
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
  2237
  by (simp add: Arccos_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2238
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2239
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
  2240
  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
  2241
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2242
text\<open>A very tricky argument to find!\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2243
lemma abs_Re_less_1_preserve:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2244
  assumes "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1)"  "Im (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
  2245
    shows "0 < Re (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
  2246
proof (cases "Im z = 0")
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2247
  case True
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2248
  then show ?thesis
60141
833adf7db7d8 New material, mostly about limits. Consolidation.
paulson <lp15@cam.ac.uk>
parents: 60020
diff changeset
  2249
    using assms
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2250
    by (fastforce simp add: cmod_def Re_power2 Im_power2 algebra_simps abs_square_less_1 [symmetric])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2251
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2252
  case False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2253
  have Imz: "Im z = - sqrt ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2254
    using assms abs_Re_le_cmod [of "1-z\<^sup>2"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2255
    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
  2256
  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
  2257
  proof (clarsimp simp add: cmod_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2258
    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
  2259
    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
  2260
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2261
    then show False using False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2262
      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
  2263
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2264
  moreover have "(Im z)\<^sup>2 = ((1 + ((Im z)\<^sup>2 + cmod (1 - z\<^sup>2)) - (Re z)\<^sup>2) / 2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2265
    apply (subst Imz, simp)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2266
    apply (subst real_sqrt_pow2)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2267
    using abs_Re_le_cmod [of "1-z\<^sup>2"]
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2268
    apply (auto simp: Re_power2 field_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2269
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2270
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2271
    by (simp add: Re_power2 Im_power2 cmod_power2)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2272
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2273
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2274
lemma isCont_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2275
  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
  2276
    shows "isCont Arccos z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2277
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2278
  have rez: "Im (1 - z\<^sup>2) = 0 \<Longrightarrow> 0 < Re (1 - z\<^sup>2)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2279
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2280
    by (auto simp: Re_power2 Im_power2 abs_square_less_1 add_pos_nonneg algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2281
  show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2282
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2283
    apply (simp add: Arccos_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2284
    apply (rule isCont_Ln' isCont_csqrt' continuous_intros)+
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2285
    apply (erule rez)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2286
    apply (blast intro: abs_Re_less_1_preserve)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2287
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2288
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2289
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2290
lemma isCont_Arccos' [simp]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2291
  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
  2292
  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
  2293
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2294
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
  2295
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2296
  have "z*2 + \<i> * (2 * csqrt (1 - z\<^sup>2)) = 0 \<longleftrightarrow> z*2 + \<i> * csqrt (1 - z\<^sup>2)*2 = 0"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61806
diff changeset
  2297
    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
  2298
  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
  2299
    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
  2300
  ultimately show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2301
    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
  2302
    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
  2303
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2304
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2305
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2306
lemma Arccos_cos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2307
    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
  2308
             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
  2309
             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
  2310
      shows "Arccos(cos z) = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2311
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
  2312
  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
  2313
    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
  2314
  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
  2315
    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
  2316
  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
  2317
                           \<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
  2318
    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
  2319
  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
  2320
                              \<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
  2321
    apply (subst csqrt_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2322
    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
  2323
    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
  2324
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2325
  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
  2326
    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
  2327
  also have "... = z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2328
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2329
    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
  2330
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2331
  finally show ?thesis .
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2332
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2333
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2334
lemma Arccos_unique:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2335
    "\<lbrakk>cos z = w;
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2336
      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
  2337
      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
  2338
      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
  2339
  using Arccos_cos by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2340
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2341
lemma Arccos_0 [simp]: "Arccos 0 = pi/2"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2342
  by (rule Arccos_unique) (auto simp: of_real_numeral)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2343
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2344
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
  2345
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2346
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2347
lemma Arccos_minus1: "Arccos(-1) = pi"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2348
  by (rule Arccos_unique) auto
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2349
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2350
lemma has_field_derivative_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2351
  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
  2352
    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
  2353
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2354
  have "(cos (Arccos z))\<^sup>2 \<noteq> 1"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2355
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2356
    apply atomize
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2357
    apply (auto simp: complex_eq_iff Re_power2 Im_power2 abs_square_eq_1)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2358
    apply (metis abs_minus_cancel abs_one abs_power2 numeral_One numeral_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2359
    apply (metis left_minus less_irrefl power_one sum_power2_gt_zero_iff zero_neq_neg_one)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2360
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2361
  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
  2362
    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
  2363
  then have "(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
  2364
    apply (rule has_complex_derivative_inverse_basic [OF DERIV_cos])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2365
    apply (auto intro: isCont_Arccos open_ball [of z 1] assms)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2366
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2367
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2368
    by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2369
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2370
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2371
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
  2372
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
  2373
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2374
lemma complex_differentiable_at_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2375
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos complex_differentiable at z"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2376
  using complex_differentiable_def has_field_derivative_Arccos by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2377
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2378
lemma complex_differentiable_within_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2379
    "(Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos complex_differentiable (at z within s)"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2380
  using complex_differentiable_at_Arccos complex_differentiable_within_subset by blast
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2381
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2382
lemma continuous_within_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2383
    "(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
  2384
  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
  2385
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2386
lemma continuous_on_Arccos [continuous_intros]:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2387
    "(\<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
  2388
  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
  2389
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2390
lemma holomorphic_on_Arccos: "(\<And>z. z \<in> s \<Longrightarrow> Im z = 0 \<Longrightarrow> \<bar>Re z\<bar> < 1) \<Longrightarrow> Arccos holomorphic_on s"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2391
  by (simp add: complex_differentiable_within_Arccos holomorphic_on_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2392
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2393
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2394
subsection\<open>Upper and Lower Bounds for Inverse Sine and Cosine\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2395
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2396
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
  2397
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2398
  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
  2399
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2400
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
  2401
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2402
  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
  2403
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2404
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
  2405
  unfolding Re_Arccos
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2406
  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
  2407
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2408
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
  2409
  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
  2410
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2411
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
  2412
  unfolding Re_Arcsin
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2413
  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
  2414
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2415
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
  2416
  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
  2417
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2418
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2419
subsection\<open>Interrelations between Arcsin and Arccos\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2420
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2421
lemma cos_Arcsin_nonzero:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2422
  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
  2423
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2424
  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
  2425
    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
  2426
  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
  2427
  proof
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2428
    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
  2429
    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
  2430
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2431
    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
  2432
      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
  2433
    then show False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2434
      using assms by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2435
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2436
  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
  2437
    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
  2438
  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
  2439
    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
  2440
  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
  2441
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2442
    apply (auto simp: power2_sum)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2443
    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
  2444
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2445
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2446
    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
  2447
    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
  2448
    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
  2449
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2450
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2451
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2452
lemma sin_Arccos_nonzero:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2453
  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
  2454
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2455
  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
  2456
    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
  2457
  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
  2458
  proof
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2459
    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
  2460
    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
  2461
      by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2462
    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
  2463
      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
  2464
    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
  2465
      using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2466
      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
  2467
    then show False
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2468
      using assms by simp
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2469
  qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2470
  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
  2471
    by (simp add: algebra_simps)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2472
  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
  2473
    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
  2474
  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
  2475
    using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2476
    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
  2477
    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
  2478
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2479
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2480
    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
  2481
    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
  2482
    apply (simp add: power2_eq_square)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2483
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2484
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2485
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2486
lemma cos_sin_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2487
  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
  2488
    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
  2489
  apply (rule csqrt_unique [THEN sym])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2490
  apply (simp add: cos_squared_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2491
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2492
  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
  2493
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2494
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2495
lemma sin_cos_csqrt:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2496
  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
  2497
    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
  2498
  apply (rule csqrt_unique [THEN sym])
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2499
  apply (simp add: sin_squared_eq)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2500
  using assms
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2501
  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
  2502
  done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2503
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2504
lemma Arcsin_Arccos_csqrt_pos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2505
    "(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
  2506
  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
  2507
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2508
lemma Arccos_Arcsin_csqrt_pos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2509
    "(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
  2510
  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
  2511
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2512
lemma sin_Arccos:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2513
    "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
  2514
  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
  2515
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2516
lemma cos_Arcsin:
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2517
    "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
  2518
  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
  2519
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2520
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2521
subsection\<open>Relationship with Arcsin on the Real Numbers\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2522
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2523
lemma Im_Arcsin_of_real:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2524
  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
  2525
    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
  2526
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2527
  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
  2528
    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
  2529
  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
  2530
    using assms abs_square_le_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2531
    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
  2532
  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
  2533
    by (simp add: norm_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2534
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2535
    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
  2536
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2537
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2538
corollary Arcsin_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arcsin z \<in> \<real>"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2539
  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
  2540
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2541
lemma arcsin_eq_Re_Arcsin:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2542
  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
  2543
    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
  2544
unfolding arcsin_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2545
proof (rule the_equality, safe)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2546
  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
  2547
    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
  2548
    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
  2549
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2550
  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
  2551
    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
  2552
    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
  2553
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2554
  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
  2555
    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
  2556
    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
  2557
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2558
  fix x'
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2559
  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
  2560
  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
  2561
    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
  2562
    apply (subst Arcsin_sin)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2563
    apply (auto simp: )
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2564
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2565
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2566
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2567
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
  2568
  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
  2569
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2570
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2571
subsection\<open>Relationship with Arccos on the Real Numbers\<close>
59870
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2572
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2573
lemma Im_Arccos_of_real:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2574
  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
  2575
    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
  2576
proof -
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2577
  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
  2578
    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
  2579
  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
  2580
    using assms abs_square_le_1
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2581
    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
  2582
  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
  2583
    by (simp add: norm_complex_def)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2584
  then show ?thesis
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2585
    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
  2586
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2587
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2588
corollary Arccos_in_Reals [simp]: "z \<in> \<real> \<Longrightarrow> \<bar>Re z\<bar> \<le> 1 \<Longrightarrow> Arccos z \<in> \<real>"
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2589
  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
  2590
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2591
lemma arccos_eq_Re_Arccos:
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2592
  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
  2593
    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
  2594
unfolding arccos_def
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2595
proof (rule the_equality, safe)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2596
  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
  2597
    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
  2598
    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
  2599
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2600
  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
  2601
    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
  2602
    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
  2603
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2604
  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
  2605
    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
  2606
    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
  2607
next
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2608
  fix x'
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2609
  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
  2610
  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
  2611
    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
  2612
    apply (subst Arccos_cos)
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2613
    apply (auto simp: )
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2614
    done
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2615
qed
68d6b6aa4450 HOL Light Libraries for complex Arctan, Arcsin, Arccos
paulson <lp15@cam.ac.uk>
parents: 59862
diff changeset
  2616
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61942
diff changeset
  2617
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
  2618
  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
  2619
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2620
subsection\<open>Some interrelationships among the real inverse trig functions.\<close>
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2621
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2622
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
  2623
  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
  2624
    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
  2625
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
  2626
  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
  2627
  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
  2628
    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
  2629
      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
  2630
      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
  2631
  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
  2632
    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
  2633
      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
  2634
      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
  2635
  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
  2636
    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
  2637
      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
  2638
      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
  2639
                    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
  2640
  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
  2641
  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
  2642
    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
  2643
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
  2644
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2645
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
  2646
  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
  2647
    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
  2648
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
  2649
  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
  2650
    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
  2651
    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
  2652
    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
  2653
    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
  2654
  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
  2655
    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
  2656
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
  2657
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2658
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
  2659
  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
  2660
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2661
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
  2662
  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
  2663
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2664
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
  2665
  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
  2666
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  2667
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
  2668
  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
  2669
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2670
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
  2671
  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
  2672
  apply (subst Arcsin_Arccos_csqrt_pos)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  2673
  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
  2674
  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
  2675
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2676
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
  2677
  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
  2678
  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
  2679
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2680
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
  2681
  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
  2682
  apply (subst Arccos_Arcsin_csqrt_pos)
60162
645058aa9d6f tidying some messy proofs
paulson <lp15@cam.ac.uk>
parents: 60150
diff changeset
  2683
  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
  2684
  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
  2685
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2686
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
  2687
  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
  2688
  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
  2689
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2690
subsection\<open>continuity results for arcsin and arccos.\<close>
60017
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2691
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2692
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
  2693
    "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
  2694
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
  2695
  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
  2696
        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
  2697
    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
  2698
  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
  2699
    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
  2700
  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
  2701
    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
  2702
          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
  2703
    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
  2704
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
  2705
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2706
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
  2707
    "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
  2708
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
  2709
  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
  2710
    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
  2711
    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
  2712
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
  2713
  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
  2714
  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
  2715
    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
  2716
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
  2717
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2718
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
  2719
    "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
  2720
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
  2721
  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
  2722
        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
  2723
    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
  2724
  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
  2725
    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
  2726
  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
  2727
    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
  2728
          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
  2729
    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
  2730
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
  2731
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2732
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
  2733
    "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
  2734
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
  2735
  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
  2736
    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
  2737
    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
  2738
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
  2739
  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
  2740
  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
  2741
    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
  2742
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
  2743
b785d6d06430 Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
paulson <lp15@cam.ac.uk>
parents: 59870
diff changeset
  2744
60420
884f54e01427 isabelle update_cartouches;
wenzelm
parents: 60162
diff changeset
  2745
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
  2746
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2747
lemma complex_root_unity:
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2748
  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
  2749
  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
  2750
    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
  2751
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
  2752
  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
  2753
    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
  2754
  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
  2755
    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
  2756
    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
  2757
    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
  2758
    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
  2759
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
  2760
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2761
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
  2762
  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
  2763
  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
  2764
    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
  2765
           \<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
  2766
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
  2767
    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
  2768
               \<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
  2769
          (\<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
  2770
              (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
  2771
      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
  2772
    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
  2773
      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
  2774
    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
  2775
      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
  2776
      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
  2777
      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
  2778
      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
  2779
    also have "... \<longleftrightarrow> int j mod int n = int k mod int 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
  2780
      by (auto simp: zmod_eq_dvd_iff dvd_def 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
  2781
    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
  2782
      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
  2783
    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
  2784
             \<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
  2785
   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
  2786
  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
  2787
    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
  2788
    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
  2789
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
  2790
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2791
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
  2792
    "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
  2793
              {..<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
  2794
  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
  2795
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2796
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
  2797
  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
  2798
  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
  2799
    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
  2800
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
  2801
  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
  2802
    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
  2803
  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
  2804
     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
  2805
     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
  2806
  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
  2807
    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
  2808
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
  2809
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2810
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
  2811
     "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
  2812
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
  2813
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2814
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
  2815
     "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
  2816
  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
  2817
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2818
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
  2819
  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
  2820
    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
  2821
  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
  2822
  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
  2823
  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
  2824
  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
  2825
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2826
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
  2827
  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
  2828
065ecea354d0 Complex roots of unity. Better definition of ln for complex numbers. Used [code del] to stop code generation for powr.
paulson <lp15@cam.ac.uk>
parents: 60017
diff changeset
  2829
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
  2830
    "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
  2831
  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
  2832
  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
  2833
  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
  2834
59745
390476a0ef13 new file for complex transcendental functions
paulson <lp15@cam.ac.uk>
parents:
diff changeset
  2835
end