src/HOL/Analysis/Gamma_Function.thy
author wenzelm
Fri, 22 Dec 2017 13:51:20 +0100
changeset 67252 c7f859868b7c
parent 66936 cf8d8fc23891
child 67278 c60e3d615b8c
permissions -rw-r--r--
proper HTML title;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63992
3aa9837d05c7 updated headers;
wenzelm
parents: 63952
diff changeset
     1
(*  Title:    HOL/Analysis/Gamma_Function.thy
62055
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
     2
    Author:   Manuel Eberl, TU München
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
     3
*)
62055
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
     4
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
     5
section \<open>The Gamma Function\<close>
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
     6
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
     7
theory Gamma_Function
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
     8
imports
66286
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
     9
  Conformal_Mappings
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
    10
  Summation_Tests
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    11
  Harmonic_Numbers
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66447
diff changeset
    12
  "HOL-Library.Nonpos_Ints"
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 66447
diff changeset
    13
  "HOL-Library.Periodic_Fun"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    14
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    15
62055
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
    16
text \<open>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
    17
  Several equivalent definitions of the Gamma function and its
62055
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
    18
  most important properties. Also contains the definition and some properties
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
    19
  of the log-Gamma function and the Digamma function and the other Polygamma functions.
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
    20
62055
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
    21
  Based on the Gamma function, we also prove the Weierstraß product form of the
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
    22
  sin function and, based on this, the solution of the Basel problem (the
62055
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
    23
  sum over all @{term "1 / (n::nat)^2"}.
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
    24
\<close>
62055
755fda743c49 Multivariate-Analysis: fixed headers and a LaTex error (c.f. Isabelle b0f941e207cf)
hoelzl
parents: 62049
diff changeset
    25
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
    26
lemma pochhammer_eq_0_imp_nonpos_Int:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    27
  "pochhammer (x::'a::field_char_0) n = 0 \<Longrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    28
  by (auto simp: pochhammer_eq_0_iff)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
    29
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    30
lemma closed_nonpos_Ints [simp]: "closed (\<int>\<^sub>\<le>\<^sub>0 :: 'a :: real_normed_algebra_1 set)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    31
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
    32
  have "\<int>\<^sub>\<le>\<^sub>0 = (of_int ` {n. n \<le> 0} :: 'a set)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    33
    by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    34
  also have "closed \<dots>" by (rule closed_of_int_image)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    35
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    36
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    37
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    38
lemma plus_one_in_nonpos_Ints_imp: "z + 1 \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    39
  using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    40
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    41
lemma of_int_in_nonpos_Ints_iff:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    42
  "(of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    43
  by (auto simp: nonpos_Ints_def)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    44
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    45
lemma one_plus_of_int_in_nonpos_Ints_iff:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    46
  "(1 + of_int n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n \<le> -1"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    47
proof -
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    48
  have "1 + of_int n = (of_int (n + 1) :: 'a)" by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    49
  also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n + 1 \<le> 0" by (subst of_int_in_nonpos_Ints_iff) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    50
  also have "\<dots> \<longleftrightarrow> n \<le> -1" by presburger
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    51
  finally show ?thesis .
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    52
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    53
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    54
lemma one_minus_of_nat_in_nonpos_Ints_iff:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    55
  "(1 - of_nat n :: 'a :: ring_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    56
proof -
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    57
  have "(1 - of_nat n :: 'a) = of_int (1 - int n)" by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    58
  also have "\<dots> \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n > 0" by (subst of_int_in_nonpos_Ints_iff) presburger
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    59
  finally show ?thesis .
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    60
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
    61
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    62
lemma fraction_not_in_ints:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    63
  assumes "\<not>(n dvd m)" "n \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    64
  shows   "of_int m / of_int n \<notin> (\<int> :: 'a :: {division_ring,ring_char_0} set)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    65
proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    66
  assume "of_int m / (of_int n :: 'a) \<in> \<int>"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    67
  then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    68
  with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: divide_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    69
  hence "m = k * n" by (subst (asm) of_int_eq_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    70
  hence "n dvd m" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    71
  with assms(1) show False by contradiction
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    72
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    73
63317
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    74
lemma fraction_not_in_nats:
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    75
  assumes "\<not>n dvd m" "n \<noteq> 0"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    76
  shows   "of_int m / of_int n \<notin> (\<nat> :: 'a :: {division_ring,ring_char_0} set)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    77
proof
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    78
  assume "of_int m / of_int n \<in> (\<nat> :: 'a set)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    79
  also note Nats_subset_Ints
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    80
  finally have "of_int m / of_int n \<in> (\<int> :: 'a set)" .
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    81
  moreover have "of_int m / of_int n \<notin> (\<int> :: 'a set)"
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    82
    using assms by (intro fraction_not_in_ints)
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    83
  ultimately show False by contradiction
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    84
qed
ca187a9f66da Various additions to polynomials, FPSs, Gamma function
eberlm
parents: 63296
diff changeset
    85
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    86
lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \<notin> \<int> \<Longrightarrow> z \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    87
  by (auto simp: Ints_def nonpos_Ints_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    88
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    89
lemma double_in_nonpos_Ints_imp:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    90
  assumes "2 * (z :: 'a :: field_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    91
  shows   "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<or> z + 1/2 \<in> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    92
proof-
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    93
  from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    94
  thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    95
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    96
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    97
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    98
lemma sin_series: "(\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
    99
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   100
  from sin_converges[of z] have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z" .
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   101
  also have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z \<longleftrightarrow>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   102
                 (\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   103
    by (subst sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66286
diff changeset
   104
       (auto simp: sin_coeff_def strict_mono_def ac_simps elim!: oddE)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   105
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   106
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   107
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   108
lemma cos_series: "(\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   109
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   110
  from cos_converges[of z] have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z" .
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   111
  also have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z \<longleftrightarrow>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   112
                 (\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   113
    by (subst sums_mono_reindex[of "\<lambda>n. 2*n", symmetric])
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66286
diff changeset
   114
       (auto simp: cos_coeff_def strict_mono_def ac_simps elim!: evenE)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   115
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   116
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   117
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   118
lemma sin_z_over_z_series:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   119
  fixes z :: "'a :: {real_normed_field,banach}"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   120
  assumes "z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   121
  shows   "(\<lambda>n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   122
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   123
  from sin_series[of z] have "(\<lambda>n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   124
    by (simp add: field_simps scaleR_conv_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   125
  from sums_mult[OF this, of "inverse z"] and assms show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   126
    by (simp add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   127
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   128
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   129
lemma sin_z_over_z_series':
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   130
  fixes z :: "'a :: {real_normed_field,banach}"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   131
  assumes "z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   132
  shows   "(\<lambda>n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   133
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   134
  from sums_split_initial_segment[OF sin_converges[of z], of 1]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   135
    have "(\<lambda>n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   136
  from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   137
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   138
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   139
lemma has_field_derivative_sin_z_over_z:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   140
  fixes A :: "'a :: {real_normed_field,banach} set"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   141
  shows "((\<lambda>z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   142
      (is "(?f has_field_derivative ?f') _")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   143
proof (rule has_field_derivative_at_within)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   144
  have "((\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   145
            has_field_derivative (\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   146
  proof (rule termdiffs_strong)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   147
    from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   148
      show "summable (\<lambda>n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   149
  qed simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   150
  also have "(\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n) = ?f"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   151
  proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   152
    fix z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   153
    show "(\<Sum>n. of_real (sin_coeff (n+1)) * z^n)  = ?f z"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   154
      by (cases "z = 0") (insert sin_z_over_z_series'[of z],
66936
cf8d8fc23891 tuned some proofs and added some lemmas
haftmann
parents: 66526
diff changeset
   155
            simp_all add: scaleR_conv_of_real sums_iff sin_coeff_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   156
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   157
  also have "(\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) =
66936
cf8d8fc23891 tuned some proofs and added some lemmas
haftmann
parents: 66526
diff changeset
   158
                 diffs (\<lambda>n. of_real (sin_coeff (Suc n))) 0" by simp
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   159
  also have "\<dots> = 0" by (simp add: sin_coeff_def diffs_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   160
  finally show "((\<lambda>z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   161
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   162
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   163
lemma round_Re_minimises_norm:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   164
  "norm ((z::complex) - of_int m) \<ge> norm (z - of_int (round (Re z)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   165
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   166
  let ?n = "round (Re z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   167
  have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   168
    by (simp add: cmod_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   169
  also have "\<bar>Re z - of_int ?n\<bar> \<le> \<bar>Re z - of_int m\<bar>" by (rule round_diff_minimal)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   170
  hence "sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2) \<le> sqrt ((Re z - of_int m)\<^sup>2 + (Im z)\<^sup>2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   171
    by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   172
  also have "\<dots> = norm (z - of_int m)" by (simp add: cmod_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   173
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   174
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   175
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   176
lemma Re_pos_in_ball:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   177
  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   178
  shows   "Re t > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   179
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   180
  have "Re (z - t) \<le> norm (z - t)" by (rule complex_Re_le_cmod)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   181
  also from assms have "\<dots> < Re z / 2" by (simp add: dist_complex_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   182
  finally show "Re t > 0" using assms by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   183
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   184
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   185
lemma no_nonpos_Int_in_ball_complex:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   186
  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   187
  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   188
  using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   189
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   190
lemma no_nonpos_Int_in_ball:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   191
  assumes "t \<in> ball z (dist z (round (Re z)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   192
  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   193
proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   194
  assume "t \<in> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   195
  then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   196
  have "dist z (of_int n) \<le> dist z t + dist t (of_int n)" by (rule dist_triangle)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   197
  also from assms have "dist z t < dist z (round (Re z))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   198
  also have "\<dots> \<le> dist z (of_int n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   199
    using round_Re_minimises_norm[of z] by (simp add: dist_complex_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   200
  finally have "dist t (of_int n) > 0" by simp
62072
bf3d9f113474 isabelle update_cartouches -c -t;
wenzelm
parents: 62055
diff changeset
   201
  with \<open>t = of_int n\<close> show False by simp
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   202
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   203
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   204
lemma no_nonpos_Int_in_ball':
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   205
  assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   206
  obtains d where "d > 0" "\<And>t. t \<in> ball z d \<Longrightarrow> t \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   207
proof (rule that)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   208
  from assms show "setdist {z} \<int>\<^sub>\<le>\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   209
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   210
  fix t assume "t \<in> ball z (setdist {z} \<int>\<^sub>\<le>\<^sub>0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   211
  thus "t \<notin> \<int>\<^sub>\<le>\<^sub>0" using setdist_le_dist[of z "{z}" t "\<int>\<^sub>\<le>\<^sub>0"] by force
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   212
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   213
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   214
lemma no_nonpos_Real_in_ball:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   215
  assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0" and t: "t \<in> ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   216
  shows   "t \<notin> \<real>\<^sub>\<le>\<^sub>0"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   217
using z
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   218
proof (cases "Im z = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   219
  assume A: "Im z = 0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   220
  with z have "Re z > 0" by (force simp add: complex_nonpos_Reals_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   221
  with t A Re_pos_in_ball[of z t] show ?thesis by (force simp add: complex_nonpos_Reals_iff)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   222
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   223
  assume A: "Im z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   224
  have "abs (Im z) - abs (Im t) \<le> abs (Im z - Im t)" by linarith
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   225
  also have "\<dots> = abs (Im (z - t))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   226
  also have "\<dots> \<le> norm (z - t)" by (rule abs_Im_le_cmod)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   227
  also from A t have "\<dots> \<le> abs (Im z) / 2" by (simp add: dist_complex_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   228
  finally have "abs (Im t) > 0" using A by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   229
  thus ?thesis by (force simp add: complex_nonpos_Reals_iff)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   230
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   231
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   232
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   233
subsection \<open>Definitions\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   234
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   235
text \<open>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   236
  We define the Gamma function by first defining its multiplicative inverse @{term "Gamma_inv"}.
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   237
  This is more convenient because @{term "Gamma_inv"} is entire, which makes proofs of its
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   238
  properties more convenient because one does not have to watch out for discontinuities.
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   239
  (e.g. @{term "Gamma_inv"} fulfils @{term "rGamma z = z * rGamma (z + 1)"} everywhere,
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   240
  whereas @{term "Gamma"} does not fulfil the analogous equation on the non-positive integers)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   241
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   242
  We define the Gamma function (resp. its inverse) in the Euler form. This form has the advantage
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   243
  that it is a relatively simple limit that converges everywhere. The limit at the poles is 0
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   244
  (due to division by 0). The functional equation @{term "Gamma (z + 1) = z * Gamma z"} follows
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   245
  immediately from the definition.
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   246
\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   247
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   248
definition Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   249
  "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   250
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   251
definition Gamma_series' :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   252
  "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   253
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   254
definition rGamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   255
  "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   256
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   257
lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   258
  and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   259
  unfolding Gamma_series_def rGamma_series_def by simp_all
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   260
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   261
lemma rGamma_series_minus_of_nat:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   262
  "eventually (\<lambda>n. rGamma_series (- of_nat k) n = 0) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   263
  using eventually_ge_at_top[of k]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   264
  by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   265
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   266
lemma Gamma_series_minus_of_nat:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   267
  "eventually (\<lambda>n. Gamma_series (- of_nat k) n = 0) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   268
  using eventually_ge_at_top[of k]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   269
  by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   270
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   271
lemma Gamma_series'_minus_of_nat:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   272
  "eventually (\<lambda>n. Gamma_series' (- of_nat k) n = 0) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   273
  using eventually_gt_at_top[of k]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   274
  by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   275
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   276
lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma_series z \<longlonglongrightarrow> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   277
  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   278
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   279
lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series z \<longlonglongrightarrow> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   280
  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   281
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   282
lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series' z \<longlonglongrightarrow> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   283
  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   284
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   285
lemma Gamma_series_Gamma_series':
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   286
  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   287
  shows   "(\<lambda>n. Gamma_series' z n / Gamma_series z n) \<longlonglongrightarrow> 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   288
proof (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   289
  from eventually_gt_at_top[of "0::nat"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   290
    show "eventually (\<lambda>n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   291
  proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   292
    fix n :: nat assume n: "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   293
    from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   294
      by (cases n, simp)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   295
         (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec'
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   296
               dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   297
    also from n have "\<dots> = z / of_nat n + 1" by (simp add: divide_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   298
    finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" ..
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   299
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   300
  have "(\<lambda>x. z / of_nat x) \<longlonglongrightarrow> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   301
    by (rule tendsto_norm_zero_cancel)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   302
       (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n],
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   303
        simp add:  norm_divide inverse_eq_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   304
  from tendsto_add[OF this tendsto_const[of 1]] show "(\<lambda>n. z / of_nat n + 1) \<longlonglongrightarrow> 1" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   305
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   306
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   307
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   308
subsection \<open>Convergence of the Euler series form\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   309
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   310
text \<open>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   311
  We now show that the series that defines the Gamma function in the Euler form converges
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   312
  and that the function defined by it is continuous on the complex halfspace with positive
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   313
  real part.
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   314
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   315
  We do this by showing that the logarithm of the Euler series is continuous and converges
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   316
  locally uniformly, which means that the log-Gamma function defined by its limit is also
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   317
  continuous.
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   318
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   319
  This will later allow us to lift holomorphicity and continuity from the log-Gamma
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
   320
  function to the inverse of the Gamma function, and from that to the Gamma function itself.
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   321
\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   322
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   323
definition ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   324
  "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\<Sum>k=1..n. ln (z / of_nat k + 1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   325
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   326
definition ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   327
  "ln_Gamma_series' z n =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   328
     - euler_mascheroni*z - ln z + (\<Sum>k=1..n. z / of_nat n - ln (z / of_nat k + 1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   329
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   330
definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> 'a" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   331
  "ln_Gamma z = lim (ln_Gamma_series z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   332
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   333
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   334
text \<open>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   335
  We now show that the log-Gamma series converges locally uniformly for all complex numbers except
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   336
  the non-positive integers. We do this by proving that the series is locally Cauchy, adapting this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   337
  proof:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   338
  http://math.stackexchange.com/questions/887158/convergence-of-gammaz-lim-n-to-infty-fracnzn-prod-m-0nzm
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   339
\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   340
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   341
context
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   342
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   343
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   344
private lemma ln_Gamma_series_complex_converges_aux:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   345
  fixes z :: complex and k :: nat
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   346
  assumes z: "z \<noteq> 0" and k: "of_nat k \<ge> 2*norm z" "k \<ge> 2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   347
  shows "norm (z * ln (1 - 1/of_nat k) + ln (z/of_nat k + 1)) \<le> 2*(norm z + norm z^2) / of_nat k^2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   348
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   349
  let ?k = "of_nat k :: complex" and ?z = "norm z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   350
  have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   351
    by (simp add: algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   352
  also have "norm ... \<le> ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   353
    by (subst norm_mult [symmetric], rule norm_triangle_ineq)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   354
  also have "norm (Ln (1 + -1/?k) - (-1/?k)) \<le> (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   355
    using k by (intro Ln_approx_linear) (simp add: norm_divide)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   356
  hence "?z * norm (ln (1-1/?k) + 1/?k) \<le> ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   357
    by (intro mult_left_mono) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   358
  also have "... \<le> (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   359
    by (simp add: field_simps power2_eq_square norm_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   360
  also have "... \<le> (?z * 2) / of_nat k^2" using k
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   361
    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   362
  also have "norm (ln (1+z/?k) - z/?k) \<le> norm (z/?k)^2 / (1 - norm (z/?k))" using k
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   363
    by (intro Ln_approx_linear) (simp add: norm_divide)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   364
  hence "norm (ln (1+z/?k) - z/?k) \<le> ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   365
    by (simp add: field_simps norm_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   366
  also have "... \<le> (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   367
    by (simp add: field_simps power2_eq_square)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   368
  also have "... \<le> (?z^2 * 2) / of_nat k^2" using k
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   369
    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   370
  also note add_divide_distrib [symmetric]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   371
  finally show ?thesis by (simp only: distrib_left mult.commute)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   372
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   373
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   374
lemma ln_Gamma_series_complex_converges:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   375
  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   376
  assumes d: "d > 0" "\<And>n. n \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> norm (z - of_int n) > d"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   377
  shows "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n :: complex)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   378
proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   379
  fix e :: real assume e: "e > 0"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   380
  define e'' where "e'' = (SUP t:ball z d. norm t + norm t^2)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   381
  define e' where "e' = e / (2*e'')"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   382
  have "bounded ((\<lambda>t. norm t + norm t^2) ` cball z d)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   383
    by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   384
  hence "bounded ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   385
  hence bdd: "bdd_above ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   386
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   387
  with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   388
    by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   389
  have e'': "norm t + norm t^2 \<le> e''" if "t \<in> ball z d" for t unfolding e''_def using that
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   390
    by (rule cSUP_upper[OF _ bdd])
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   391
  from e z e''_pos have e': "e' > 0" unfolding e'_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   392
    by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   393
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   394
  have "summable (\<lambda>k. inverse ((real_of_nat k)^2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   395
    by (rule inverse_power_summable) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   396
  from summable_partial_sum_bound[OF this e'] guess M . note M = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   397
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   398
  define N where "N = max 2 (max (nat \<lceil>2 * (norm z + d)\<rceil>) M)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   399
  {
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   400
    from d have "\<lceil>2 * (cmod z + d)\<rceil> \<ge> \<lceil>0::real\<rceil>"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   401
      by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   402
    hence "2 * (norm z + d) \<le> of_nat (nat \<lceil>2 * (norm z + d)\<rceil>)" unfolding N_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   403
      by (simp_all add: le_of_int_ceiling)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   404
    also have "... \<le> of_nat N" unfolding N_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   405
      by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   406
    finally have "of_nat N \<ge> 2 * (norm z + d)" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   407
    moreover have "N \<ge> 2" "N \<ge> M" unfolding N_def by simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   408
    moreover have "(\<Sum>k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \<ge> N" for m n
62072
bf3d9f113474 isabelle update_cartouches -c -t;
wenzelm
parents: 62055
diff changeset
   409
      using M[OF order.trans[OF \<open>N \<ge> M\<close> that]] unfolding real_norm_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   410
      by (subst (asm) abs_of_nonneg) (auto intro: sum_nonneg simp: divide_simps)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   411
    moreover note calculation
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   412
  } note N = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   413
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   414
  show "\<exists>M. \<forall>t\<in>ball z d. \<forall>m\<ge>M. \<forall>n>m. dist (ln_Gamma_series t m) (ln_Gamma_series t n) < e"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   415
    unfolding dist_complex_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   416
  proof (intro exI[of _ N] ballI allI impI)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   417
    fix t m n assume t: "t \<in> ball z d" and mn: "m \<ge> N" "n > m"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   418
    from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   419
    also have "dist z 0 - dist z t \<le> dist 0 t" using dist_triangle[of 0 z t]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   420
      by (simp add: dist_commute)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   421
    finally have t_nz: "t \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   422
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   423
    have "norm t \<le> norm z + norm (t - z)" by (rule norm_triangle_sub)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   424
    also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   425
    also have "2 * (norm z + d) \<le> of_nat N" by (rule N)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   426
    also have "N \<le> m" by (rule mn)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   427
    finally have norm_t: "2 * norm t < of_nat m" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   428
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   429
    have "ln_Gamma_series t m - ln_Gamma_series t n =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   430
              (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) +
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   431
              ((\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   432
      by (simp add: ln_Gamma_series_def algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   433
    also have "(\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   434
                 (\<Sum>k\<in>{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   435
      by (simp add: sum_diff)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   436
    also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   437
    also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   438
                   (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   439
      by (subst sum_telescope'' [symmetric]) simp_all
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   440
    also have "... = (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   441
      by (intro sum_cong_Suc)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   442
         (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   443
    also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \<in> {Suc m..n}" for k
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   444
      using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   445
    hence "(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   446
             (\<Sum>k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   447
      by (intro sum.cong) simp_all
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   448
    also note sum.distrib [symmetric]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   449
    also have "norm (\<Sum>k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \<le>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   450
      (\<Sum>k=Suc m..n. 2 * (norm t + (norm t)\<^sup>2) / (real_of_nat k)\<^sup>2)" using t_nz N(2) mn norm_t
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   451
      by (intro order.trans[OF norm_sum sum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   452
    also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   453
      by (simp add: sum_distrib_left)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   454
    also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   455
      by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   456
    also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   457
      by (simp add: e'_def field_simps power2_eq_square)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   458
    also from e''[OF t] e''_pos e
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   459
      have "\<dots> \<le> e * 1" by (intro mult_left_mono) (simp_all add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   460
    finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   461
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   462
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   463
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   464
end
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   465
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   466
lemma ln_Gamma_series_complex_converges':
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   467
  assumes z: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   468
  shows "\<exists>d>0. uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   469
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   470
  define d' where "d' = Re z"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   471
  define d where "d = (if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   472
  have "of_int (round d') \<in> \<int>\<^sub>\<le>\<^sub>0" if "d' \<le> 0" using that
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   473
    by (intro nonpos_Ints_of_int) (simp_all add: round_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   474
  with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   475
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   476
  have "d < cmod (z - of_int n)" if "n \<in> \<int>\<^sub>\<le>\<^sub>0" for n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   477
  proof (cases "Re z > 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   478
    case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   479
    from nonpos_Ints_nonpos[OF that] have n: "n \<le> 0" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   480
    from True have "d = Re z/2" by (simp add: d_def d'_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   481
    also from n True have "\<dots> < Re (z - of_int n)" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   482
    also have "\<dots> \<le> norm (z - of_int n)" by (rule complex_Re_le_cmod)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   483
    finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   484
  next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   485
    case False
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   486
    with assms nonpos_Ints_of_int[of "round (Re z)"]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   487
      have "z \<noteq> of_int (round d')" by (auto simp: not_less)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   488
    with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   489
    also have "\<dots> \<le> norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   490
    finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   491
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   492
  hence conv: "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   493
    by (intro ln_Gamma_series_complex_converges d_pos z) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   494
  from d_pos conv show ?thesis by blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   495
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   496
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   497
lemma ln_Gamma_series_complex_converges'': "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> convergent (ln_Gamma_series z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   498
  by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   499
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   500
lemma ln_Gamma_complex_LIMSEQ: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma_series z \<longlonglongrightarrow> ln_Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   501
  using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   502
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   503
lemma exp_ln_Gamma_series_complex:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   504
  assumes "n > 0" "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   505
  shows   "exp (ln_Gamma_series z n :: complex) = Gamma_series z n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   506
proof -
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
   507
  from assms obtain m where m: "n = Suc m" by (cases n) blast
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   508
  from assms have "z \<noteq> 0" by (intro notI) auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   509
  with assms have "exp (ln_Gamma_series z n) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   510
          (of_nat n) powr z / (z * (\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   511
    unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_sum)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   512
  also from assms have "(\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))) = (\<Prod>k=1..n. z / of_nat k + 1)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   513
    by (intro prod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
   514
  also have "... = (\<Prod>k=1..n. z + k) / fact n"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   515
    by (simp add: fact_prod)
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   516
    (subst prod_dividef [symmetric], simp_all add: field_simps)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
   517
  also from m have "z * ... = (\<Prod>k=0..n. z + k) / fact n"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   518
    by (simp add: prod.atLeast0_atMost_Suc_shift prod.atLeast_Suc_atMost_Suc_shift)
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
   519
  also have "(\<Prod>k=0..n. z + k) = pochhammer z (Suc n)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   520
    unfolding pochhammer_prod
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
   521
    by (simp add: prod.atLeast0_atMost_Suc atLeastLessThanSuc_atLeastAtMost)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   522
  also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n"
66936
cf8d8fc23891 tuned some proofs and added some lemmas
haftmann
parents: 66526
diff changeset
   523
    unfolding Gamma_series_def using assms by (simp add: divide_simps powr_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   524
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   525
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   526
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   527
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   528
lemma ln_Gamma_series'_aux:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   529
  assumes "(z::complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   530
  shows   "(\<lambda>k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   531
              (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   532
unfolding sums_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   533
proof (rule Lim_transform)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   534
  show "(\<lambda>n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \<longlonglongrightarrow> ?s"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   535
    (is "?g \<longlonglongrightarrow> _")
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   536
    by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   537
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   538
  have A: "eventually (\<lambda>n. (\<Sum>k<n. ?f k) - ?g n = 0) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   539
    using eventually_gt_at_top[of "0::nat"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   540
  proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   541
    fix n :: nat assume n: "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   542
    have "(\<Sum>k<n. ?f k) = (\<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   543
      by (subst atLeast0LessThan [symmetric], subst sum_shift_bounds_Suc_ivl [symmetric],
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   544
          subst atLeastLessThanSuc_atLeastAtMost) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   545
    also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   546
      by (simp add: harm_def sum_subtractf sum_distrib_left divide_inverse)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   547
    also from n have "\<dots> - ?g n = 0"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   548
      by (simp add: ln_Gamma_series_def sum_subtractf algebra_simps Ln_of_nat)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   549
    finally show "(\<Sum>k<n. ?f k) - ?g n = 0" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   550
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   551
  show "(\<lambda>n. (\<Sum>k<n. ?f k) - ?g n) \<longlonglongrightarrow> 0" by (subst tendsto_cong[OF A]) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   552
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   553
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   554
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   555
lemma uniformly_summable_deriv_ln_Gamma:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   556
  assumes z: "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0" and d: "d > 0" "d \<le> norm z/2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   557
  shows "uniformly_convergent_on (ball z d)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   558
            (\<lambda>k z. \<Sum>i<k. inverse (of_nat (Suc i)) - inverse (z + of_nat (Suc i)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   559
           (is "uniformly_convergent_on _ (\<lambda>k z. \<Sum>i<k. ?f i z)")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   560
proof (rule weierstrass_m_test'_ev)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   561
  {
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   562
    fix t assume t: "t \<in> ball z d"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   563
    have "norm z = norm (t + (z - t))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   564
    have "norm (t + (z - t)) \<le> norm t + norm (z - t)" by (rule norm_triangle_ineq)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   565
    also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   566
    finally have A: "norm t > norm z / 2" using z by (simp add: field_simps)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   567
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   568
    have "norm t = norm (z + (t - z))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   569
    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   570
    also from t d have "norm (t - z) \<le> norm z / 2" by (simp add: dist_norm norm_minus_commute)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   571
    also from z have "\<dots> < norm z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   572
    finally have B: "norm t < 2 * norm z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   573
    note A B
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   574
  } note ball = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   575
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   576
  show "eventually (\<lambda>n. \<forall>t\<in>ball z d. norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   577
    using eventually_gt_at_top apply eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   578
  proof safe
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   579
    fix t :: 'a assume t: "t \<in> ball z d"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   580
    from z ball[OF t] have t_nz: "t \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   581
    fix n :: nat assume n: "n > nat \<lceil>4 * norm z\<rceil>"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   582
    from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   583
    also from n have "\<dots>  < of_nat n" by linarith
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   584
    finally have n: "of_nat n > 2 * norm t" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   585
    hence "of_nat n > norm t" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   586
    hence t': "t \<noteq> -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   587
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   588
    with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   589
      by (simp add: divide_simps eq_neg_iff_add_eq_0 del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   590
    also have "norm \<dots> = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   591
      by (simp add: norm_divide norm_mult divide_simps add_ac del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   592
    also {
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   593
      from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \<le> of_nat (Suc n) / (2 * norm t)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   594
        by (intro divide_left_mono mult_pos_pos) simp_all
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   595
      also have "\<dots> < norm (of_nat (Suc n) / t) - norm (1 :: 'a)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   596
        using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   597
      also have "\<dots> \<le> norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   598
      finally have "inverse (norm (of_nat (Suc n)/t + 1)) \<le> 4 * norm z / of_nat (Suc n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   599
        using z by (simp add: divide_simps norm_divide mult_ac del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   600
    }
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   601
    also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) =
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   602
                 4 * norm z * inverse (of_nat (Suc n)^2)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   603
                 by (simp add: divide_simps power2_eq_square del: of_nat_Suc)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   604
    finally show "norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   605
      by (simp del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   606
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   607
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   608
  show "summable (\<lambda>n. 4 * norm z * inverse ((of_nat (Suc n))^2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   609
    by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   610
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   611
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   612
lemma summable_deriv_ln_Gamma:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   613
  "z \<noteq> (0 :: 'a :: {real_normed_field,banach}) \<Longrightarrow>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   614
     summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   615
  unfolding summable_iff_convergent
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   616
  by (rule uniformly_convergent_imp_convergent,
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   617
      rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   618
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   619
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   620
definition Polygamma :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   621
  "Polygamma n z = (if n = 0 then
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   622
      (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   623
      (-1)^Suc n * fact n * (\<Sum>k. inverse ((z + of_nat k)^Suc n)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   624
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   625
abbreviation Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   626
  "Digamma \<equiv> Polygamma 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   627
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   628
lemma Digamma_def:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   629
  "Digamma z = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   630
  by (simp add: Polygamma_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   631
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   632
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   633
lemma summable_Digamma:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   634
  assumes "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   635
  shows   "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   636
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   637
  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   638
                       (0 - inverse (z + of_nat 0))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   639
    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   640
              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   641
  from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   642
    show "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   643
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   644
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   645
lemma summable_offset:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   646
  assumes "summable (\<lambda>n. f (n + k) :: 'a :: real_normed_vector)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   647
  shows   "summable f"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   648
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   649
  from assms have "convergent (\<lambda>m. \<Sum>n<m. f (n + k))" by (simp add: summable_iff_convergent)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   650
  hence "convergent (\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   651
    by (intro convergent_add convergent_const)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   652
  also have "(\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k))) = (\<lambda>m. \<Sum>n<m+k. f n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   653
  proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   654
    fix m :: nat
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   655
    have "{..<m+k} = {..<k} \<union> {k..<m+k}" by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   656
    also have "(\<Sum>n\<in>\<dots>. f n) = (\<Sum>n<k. f n) + (\<Sum>n=k..<m+k. f n)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   657
      by (rule sum.union_disjoint) auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   658
    also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))"
66936
cf8d8fc23891 tuned some proofs and added some lemmas
haftmann
parents: 66526
diff changeset
   659
      using sum_shift_bounds_nat_ivl [of f 0 k m] by simp
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   660
    finally show "(\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)) = (\<Sum>n<m+k. f n)" by (simp add: atLeast0LessThan)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   661
  qed
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   662
  finally have "(\<lambda>a. sum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. sum f {..<m + k})"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   663
    by (auto simp: convergent_LIMSEQ_iff dest: LIMSEQ_offset)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   664
  thus ?thesis by (auto simp: summable_iff_convergent convergent_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   665
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   666
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   667
lemma Polygamma_converges:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   668
  fixes z :: "'a :: {real_normed_field,banach}"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   669
  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   670
  shows "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)^n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   671
proof (rule weierstrass_m_test'_ev)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   672
  define e where "e = (1 + d / norm z)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   673
  define m where "m = nat \<lceil>norm z * e\<rceil>"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   674
  {
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   675
    fix t assume t: "t \<in> ball z d"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   676
    have "norm t = norm (z + (t - z))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   677
    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   678
    also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   679
    finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   680
  } note ball = this
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   681
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   682
  show "eventually (\<lambda>k. \<forall>t\<in>ball z d. norm (inverse ((t + of_nat k)^n)) \<le>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   683
            inverse (of_nat (k - m)^n)) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   684
    using eventually_gt_at_top[of m] apply eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   685
  proof (intro ballI)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   686
    fix k :: nat and t :: 'a assume k: "k > m" and t: "t \<in> ball z d"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   687
    from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   688
    also have "\<dots> \<le> norm (of_nat k :: 'a) - norm z * e"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   689
      unfolding m_def by (subst norm_of_nat) linarith
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   690
    also from ball[OF t] have "\<dots> \<le> norm (of_nat k :: 'a) - norm t" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   691
    also have "\<dots> \<le> norm (of_nat k + t)" by (rule norm_diff_ineq)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   692
    finally have "inverse ((norm (t + of_nat k))^n) \<le> inverse (real_of_nat (k - m)^n)" using k n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   693
      by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   694
    thus "norm (inverse ((t + of_nat k)^n)) \<le> inverse (of_nat (k - m)^n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   695
      by (simp add: norm_inverse norm_power power_inverse)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   696
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   697
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   698
  have "summable (\<lambda>k. inverse ((real_of_nat k)^n))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   699
    using inverse_power_summable[of n] n by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   700
  hence "summable (\<lambda>k. inverse ((real_of_nat (k + m - m))^n))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   701
  thus "summable (\<lambda>k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   702
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   703
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   704
lemma Polygamma_converges':
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   705
  fixes z :: "'a :: {real_normed_field,banach}"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   706
  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   707
  shows "summable (\<lambda>k. inverse ((z + of_nat k)^n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   708
  using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   709
  by (simp add: summable_iff_convergent)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   710
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   711
lemma Digamma_LIMSEQ:
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   712
  fixes z :: "'a :: {banach,real_normed_field}"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   713
  assumes z: "z \<noteq> 0"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   714
  shows   "(\<lambda>m. of_real (ln (real m)) - (\<Sum>n<m. inverse (z + of_nat n))) \<longlonglongrightarrow> Digamma z"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   715
proof -
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   716
  have "(\<lambda>n. of_real (ln (real n / (real (Suc n))))) \<longlonglongrightarrow> (of_real (ln 1) :: 'a)"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   717
    by (intro tendsto_intros LIMSEQ_n_over_Suc_n) simp_all
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   718
  hence "(\<lambda>n. of_real (ln (real n / (real n + 1)))) \<longlonglongrightarrow> (0 :: 'a)" by (simp add: add_ac)
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   719
  hence lim: "(\<lambda>n. of_real (ln (real n)) - of_real (ln (real n + 1))) \<longlonglongrightarrow> (0::'a)"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   720
  proof (rule Lim_transform_eventually [rotated])
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   721
    show "eventually (\<lambda>n. of_real (ln (real n / (real n + 1))) =
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   722
            of_real (ln (real n)) - (of_real (ln (real n + 1)) :: 'a)) at_top"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   723
      using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_div)
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   724
  qed
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   725
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   726
  from summable_Digamma[OF z]
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   727
    have "(\<lambda>n. inverse (of_nat (n+1)) - inverse (z + of_nat n))
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   728
              sums (Digamma z + euler_mascheroni)"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   729
    by (simp add: Digamma_def summable_sums)
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   730
  from sums_diff[OF this euler_mascheroni_sum]
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   731
    have "(\<lambda>n. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)) - inverse (z + of_nat n))
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   732
            sums Digamma z" by (simp add: add_ac)
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   733
  hence "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1))) -
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   734
              (\<Sum>n<m. inverse (z + of_nat n))) \<longlonglongrightarrow> Digamma z"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   735
    by (simp add: sums_def sum_subtractf)
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   736
  also have "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)))) =
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   737
                 (\<lambda>m. of_real (ln (m + 1)) :: 'a)"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   738
    by (subst sum_lessThan_telescope) simp_all
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   739
  finally show ?thesis by (rule Lim_transform) (insert lim, simp)
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   740
qed
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
   741
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   742
lemma has_field_derivative_ln_Gamma_complex [derivative_intros]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   743
  fixes z :: complex
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   744
  assumes z: "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   745
  shows   "(ln_Gamma has_field_derivative Digamma z) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   746
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   747
  have not_nonpos_Int [simp]: "t \<notin> \<int>\<^sub>\<le>\<^sub>0" if "Re t > 0" for t
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   748
    using that by (auto elim!: nonpos_Ints_cases')
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   749
  from z have z': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" and z'': "z \<noteq> 0" using nonpos_Ints_subset_nonpos_Reals nonpos_Reals_zero_I
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   750
     by blast+
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   751
  let ?f' = "\<lambda>z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   752
  let ?f = "\<lambda>z k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))" and ?F' = "\<lambda>z. \<Sum>n. ?f' z n"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   753
  define d where "d = min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   754
  from z have d: "d > 0" "norm z/2 \<ge> d" by (auto simp add: complex_nonpos_Reals_iff d_def)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   755
  have ball: "Im t = 0 \<longrightarrow> Re t > 0" if "dist z t < d" for t
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   756
    using no_nonpos_Real_in_ball[OF z, of t] that unfolding d_def by (force simp add: complex_nonpos_Reals_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   757
  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   758
                       (0 - inverse (z + of_nat 0))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   759
    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   760
              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   761
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   762
  have "((\<lambda>z. \<Sum>n. ?f z n) has_field_derivative ?F' z) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   763
    using d z ln_Gamma_series'_aux[OF z']
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   764
    apply (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   765
    apply (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   766
             simp: field_simps sums_iff nonpos_Reals_divide_of_nat_iff
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   767
             simp del: of_nat_Suc)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   768
    apply (auto simp add: complex_nonpos_Reals_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   769
    done
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   770
  with z have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   771
                   ?F' z - euler_mascheroni - inverse z) (at z)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   772
    by (force intro!: derivative_eq_intros simp: Digamma_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   773
  also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   774
  also from sums have "-inverse z = (\<Sum>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   775
    by (simp add: sums_iff)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   776
  also from sums summable_deriv_ln_Gamma[OF z'']
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   777
    have "?F' z + \<dots> =  (\<Sum>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   778
    by (subst suminf_add) (simp_all add: add_ac sums_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   779
  also have "\<dots> - euler_mascheroni = Digamma z" by (simp add: Digamma_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   780
  finally have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   781
                    has_field_derivative Digamma z) (at z)" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   782
  moreover from eventually_nhds_ball[OF d(1), of z]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   783
    have "eventually (\<lambda>z. ln_Gamma z = (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   784
  proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   785
    fix t assume "t \<in> ball z d"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   786
    hence "t \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   787
    from ln_Gamma_series'_aux[OF this]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   788
      show "ln_Gamma t = (\<Sum>k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   789
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   790
  ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   791
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   792
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   793
declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   794
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   795
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   796
lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   797
  by (simp add: Digamma_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   798
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   799
lemma Digamma_plus1:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   800
  assumes "z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   801
  shows   "Digamma (z+1) = Digamma z + 1/z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   802
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   803
  have sums: "(\<lambda>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   804
                  sums (inverse (z + of_nat 0) - 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   805
    by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   806
              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   807
  have "Digamma (z+1) = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) -
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   808
          euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   809
  also have "suminf ?f = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) +
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   810
                         (\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   811
    using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   812
  also have "(\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   813
    using sums by (simp add: sums_iff inverse_eq_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   814
  finally show ?thesis by (simp add: Digamma_def[of z])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   815
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   816
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   817
lemma Polygamma_plus1:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   818
  assumes "z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   819
  shows   "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   820
proof (cases "n = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   821
  assume n: "n \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   822
  let ?f = "\<lambda>k. inverse ((z + of_nat k) ^ Suc n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   823
  have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\<Sum>k. ?f (k+1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   824
    using n by (simp add: Polygamma_def add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   825
  also have "(\<Sum>k. ?f (k+1)) + (\<Sum>k<1. ?f k) = (\<Sum>k. ?f k)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   826
    using Polygamma_converges'[OF assms, of "Suc n"] n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   827
    by (subst suminf_split_initial_segment [symmetric]) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   828
  hence "(\<Sum>k. ?f (k+1)) = (\<Sum>k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   829
  also have "(-1) ^ Suc n * fact n * ((\<Sum>k. ?f k) - inverse (z ^ Suc n)) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   830
               Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   831
    by (simp add: inverse_eq_divide algebra_simps Polygamma_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   832
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   833
qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   834
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   835
lemma Digamma_of_nat:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   836
  "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   837
proof (induction n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   838
  case (Suc n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   839
  have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   840
  also have "\<dots> = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   841
    by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   842
  also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   843
  also have "\<dots> + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   844
    by (simp add: harm_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   845
  finally show ?case .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   846
qed (simp add: harm_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   847
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   848
lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   849
  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   850
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   851
lemma Polygamma_of_real: "x \<noteq> 0 \<Longrightarrow> Polygamma n (of_real x) = of_real (Polygamma n x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   852
  unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   853
  by (simp_all add: suminf_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   854
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   855
lemma Polygamma_Real: "z \<in> \<real> \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Polygamma n z \<in> \<real>"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   856
  by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   857
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   858
lemma Digamma_half_integer:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   859
  "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   860
      (\<Sum>k<n. 2 / (of_nat (2*k+1))) - euler_mascheroni - of_real (2 * ln 2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   861
proof (induction n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   862
  case 0
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   863
  have "Digamma (1/2 :: 'a) = of_real (Digamma (1/2))" by (simp add: Polygamma_of_real [symmetric])
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   864
  also have "Digamma (1/2::real) =
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   865
               (\<Sum>k. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   866
    by (simp add: Digamma_def add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   867
  also have "(\<Sum>k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   868
             (\<Sum>k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   869
    by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   870
  also have "\<dots> = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums']
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   871
    by (subst suminf_mult) (simp_all add: algebra_simps sums_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   872
  finally show ?case by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   873
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   874
  case (Suc n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   875
  have nz: "2 * of_nat n + (1:: 'a) \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   876
     using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   877
  hence nz': "of_nat n + (1/2::'a) \<noteq> 0" by (simp add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   878
  have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   879
  also from nz' have "\<dots> = Digamma (of_nat n + 1 / 2) + 1 / (of_nat n + 1 / 2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   880
    by (rule Digamma_plus1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   881
  also from nz nz' have "1 / (of_nat n + 1 / 2 :: 'a) = 2 / (2 * of_nat n + 1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   882
    by (subst divide_eq_eq) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   883
  also note Suc
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   884
  finally show ?case by (simp add: add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   885
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   886
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   887
lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   888
  using Digamma_half_integer[of 0] by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   889
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   890
lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   891
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   892
  have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   893
  also have "\<dots> = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
   894
  also note euler_mascheroni_less_13_over_22
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
   895
  also note ln2_le_25_over_36
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   896
  finally show ?thesis by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   897
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   898
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   899
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   900
lemma has_field_derivative_Polygamma [derivative_intros]:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   901
  fixes z :: "'a :: {real_normed_field,euclidean_space}"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   902
  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   903
  shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   904
proof (rule has_field_derivative_at_within, cases "n = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   905
  assume n: "n = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   906
  let ?f = "\<lambda>k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   907
  let ?F = "\<lambda>z. \<Sum>k. ?f k z" and ?f' = "\<lambda>k z. inverse ((z + of_nat k)\<^sup>2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   908
  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   909
  from z have summable: "summable (\<lambda>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   910
    by (intro summable_Digamma) force
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   911
  from z have conv: "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)\<^sup>2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   912
    by (intro Polygamma_converges) auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   913
  with d have "summable (\<lambda>k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   914
    by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent )
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   915
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   916
  have "(?F has_field_derivative (\<Sum>k. ?f' k z)) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   917
  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   918
    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   919
    from t d(2)[of t] show "((\<lambda>z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   920
      by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   921
               dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   922
  qed (insert d(1) summable conv, (assumption|simp)+)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   923
  with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   924
    unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   925
    by (force simp: power2_eq_square intro!: derivative_eq_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   926
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   927
  assume n: "n \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   928
  from z have z': "z \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   929
  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
   930
  define n' where "n' = Suc n"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   931
  from n have n': "n' \<ge> 2" by (simp add: n'_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   932
  have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   933
                (\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   934
  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   935
    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   936
    with d have t': "t \<notin> \<int>\<^sub>\<le>\<^sub>0" "t \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   937
    show "((\<lambda>a. inverse ((a + of_nat k) ^ n')) has_field_derivative
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   938
                - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t'
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   939
      by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   940
  next
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   941
    have "uniformly_convergent_on (ball z d)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   942
              (\<lambda>k z. (- of_nat n' :: 'a) * (\<Sum>i<k. inverse ((z + of_nat i) ^ (n'+1))))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   943
      using z' n by (intro uniformly_convergent_mult Polygamma_converges) (simp_all add: n'_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   944
    thus "uniformly_convergent_on (ball z d)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   945
              (\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
   946
      by (subst (asm) sum_distrib_left) simp
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   947
  qed (insert Polygamma_converges'[OF z' n'] d, simp_all)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   948
  also have "(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   949
               (- of_nat n') * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   950
    using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   951
  finally have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   952
                    - of_nat n' * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   953
  from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   954
    show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   955
    unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   956
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   957
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   958
declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   959
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   960
lemma isCont_Polygamma [continuous_intros]:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   961
  fixes f :: "_ \<Rightarrow> 'a :: {real_normed_field,euclidean_space}"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   962
  shows "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Polygamma n (f x)) z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   963
  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Polygamma]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   964
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   965
lemma continuous_on_Polygamma:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   966
  "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A (Polygamma n :: _ \<Rightarrow> 'a :: {real_normed_field,euclidean_space})"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   967
  by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   968
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   969
lemma isCont_ln_Gamma_complex [continuous_intros]:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   970
  fixes f :: "'a::t2_space \<Rightarrow> complex"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   971
  shows "isCont f z \<Longrightarrow> f z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>z. ln_Gamma (f z)) z"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   972
  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   973
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   974
lemma continuous_on_ln_Gamma_complex [continuous_intros]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   975
  fixes A :: "complex set"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   976
  shows "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A ln_Gamma"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
   977
  by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident])
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
   978
     fastforce
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   979
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   980
lemma deriv_Polygamma:
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   981
  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   982
  shows   "deriv (Polygamma m) z =
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   983
             Polygamma (Suc m) (z :: 'a :: {real_normed_field,euclidean_space})"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   984
  by (intro DERIV_imp_deriv has_field_derivative_Polygamma assms)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   985
    thm has_field_derivative_Polygamma
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   986
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   987
lemma higher_deriv_Polygamma:
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   988
  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   989
  shows   "(deriv ^^ n) (Polygamma m) z =
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   990
             Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   991
proof -
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   992
  have "eventually (\<lambda>u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   993
  proof (induction n)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   994
    case (Suc n)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   995
    from Suc.IH have "eventually (\<lambda>z. eventually (\<lambda>u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   996
      by (simp add: eventually_eventually)
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
   997
    hence "eventually (\<lambda>z. deriv ((deriv ^^ n) (Polygamma m)) z =
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   998
             deriv (Polygamma (m + n)) z) (nhds z)"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
   999
      by eventually_elim (intro deriv_cong_ev refl)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1000
    moreover have "eventually (\<lambda>z. z \<in> UNIV - \<int>\<^sub>\<le>\<^sub>0) (nhds z)" using assms
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1001
      by (intro eventually_nhds_in_open open_Diff open_UNIV) auto
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1002
    ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1003
  qed simp_all
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1004
  thus ?thesis by (rule eventually_nhds_x_imp_x)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1005
qed
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1006
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1007
lemma deriv_ln_Gamma_complex:
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1008
  assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1009
  shows   "deriv ln_Gamma z = Digamma (z :: complex)"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1010
  by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_complex assms)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1011
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1012
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1013
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1014
text \<open>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1015
  We define a type class that captures all the fundamental properties of the inverse of the Gamma function
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1016
  and defines the Gamma function upon that. This allows us to instantiate the type class both for
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1017
  the reals and for the complex numbers with a minimal amount of proof duplication.
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1018
\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1019
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1020
class Gamma = real_normed_field + complete_space +
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1021
  fixes rGamma :: "'a \<Rightarrow> 'a"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1022
  assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1023
  assumes differentiable_rGamma_aux1:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1024
    "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1025
     let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k))
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1026
               \<longlonglongrightarrow> d) - scaleR euler_mascheroni 1
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1027
     in  filterlim (\<lambda>y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1028
                        norm (y - z)) (nhds 0) (at z)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1029
  assumes differentiable_rGamma_aux2:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1030
    "let z = - of_nat n
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1031
     in  filterlim (\<lambda>y. (rGamma y - rGamma z - (-1)^n * (prod of_nat {1..n}) * (y - z)) /\<^sub>R
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1032
                        norm (y - z)) (nhds 0) (at z)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1033
  assumes rGamma_series_aux: "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1034
             let fact' = (\<lambda>n. prod of_nat {1..n});
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1035
                 exp = (\<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x^k /\<^sub>R fact k) \<longlonglongrightarrow> e);
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1036
                 pochhammer' = (\<lambda>a n. (\<Prod>n = 0..n. a + of_nat n))
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1037
             in  filterlim (\<lambda>n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1))))
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1038
                     (nhds (rGamma z)) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1039
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1040
subclass banach ..
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1041
end
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1042
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1043
definition "Gamma z = inverse (rGamma z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1044
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1045
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1046
subsection \<open>Basic properties\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1047
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1048
lemma Gamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1049
  and rGamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1050
  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1051
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1052
lemma Gamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1053
  and rGamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1054
  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1055
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1056
lemma Gamma_eq_zero_iff: "Gamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1057
  and rGamma_eq_zero_iff: "rGamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1058
  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1059
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1060
lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1061
  unfolding Gamma_def by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1062
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1063
lemma rGamma_series_LIMSEQ [tendsto_intros]:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1064
  "rGamma_series z \<longlonglongrightarrow> rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1065
proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1066
  case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1067
  hence "z \<noteq> - of_nat n" for n by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1068
  from rGamma_series_aux[OF this] show ?thesis
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1069
    by (simp add: rGamma_series_def[abs_def] fact_prod pochhammer_Suc_prod
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63317
diff changeset
  1070
                  exp_def of_real_def[symmetric] suminf_def sums_def[abs_def] atLeast0AtMost)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1071
qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1072
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1073
lemma Gamma_series_LIMSEQ [tendsto_intros]:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1074
  "Gamma_series z \<longlonglongrightarrow> Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1075
proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1076
  case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1077
  hence "(\<lambda>n. inverse (rGamma_series z n)) \<longlonglongrightarrow> inverse (rGamma z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1078
    by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1079
  also have "(\<lambda>n. inverse (rGamma_series z n)) = Gamma_series z"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1080
    by (simp add: rGamma_series_def Gamma_series_def[abs_def])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1081
  finally show ?thesis by (simp add: Gamma_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1082
qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1083
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1084
lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1085
  using Gamma_series_LIMSEQ[of z] by (simp add: limI)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1086
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1087
lemma rGamma_1 [simp]: "rGamma 1 = 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1088
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1089
  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1090
    using eventually_gt_at_top[of "0::nat"]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1091
    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1092
                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1093
  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1094
  moreover have "rGamma_series 1 \<longlonglongrightarrow> rGamma 1" by (rule tendsto_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1095
  ultimately show ?thesis by (intro LIMSEQ_unique)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1096
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1097
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1098
lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1099
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1100
  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1101
  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1102
    using eventually_gt_at_top[of "0::nat"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1103
  proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1104
    fix n :: nat assume n: "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1105
    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1106
             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1107
      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1108
    also from n have "\<dots> = ?f n * rGamma_series z n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1109
      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1110
    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1111
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1112
  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1113
    by (intro tendsto_intros lim_inverse_n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1114
  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1115
  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1116
    by (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1117
  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1118
    by (intro tendsto_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1119
  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1120
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1121
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1122
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1123
lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1124
proof (induction n arbitrary: z)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1125
  case (Suc n z)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1126
  have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1127
  also note rGamma_plus1 [symmetric]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1128
  finally show ?case by (simp add: add_ac pochhammer_rec')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1129
qed simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1130
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1131
lemma Gamma_plus1: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma (z + 1) = z * Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1132
  using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1133
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1134
lemma pochhammer_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> pochhammer z n = Gamma (z + of_nat n) / Gamma z"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1135
  using pochhammer_rGamma[of z]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1136
  by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1137
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1138
lemma Gamma_0 [simp]: "Gamma 0 = 0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1139
  and rGamma_0 [simp]: "rGamma 0 = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1140
  and Gamma_neg_1 [simp]: "Gamma (- 1) = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1141
  and rGamma_neg_1 [simp]: "rGamma (- 1) = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1142
  and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1143
  and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1144
  and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1145
  and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1146
  by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1147
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1148
lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1149
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1150
lemma Gamma_fact: "Gamma (1 + of_nat n) = fact n"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  1151
  by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1152
        of_nat_Suc [symmetric] del: of_nat_Suc)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1153
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1154
lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  1155
  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc,
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1156
      subst of_nat_Suc, subst Gamma_fact) (rule refl)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1157
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1158
lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1159
proof (cases "n > 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1160
  case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1161
  hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1162
  with True show ?thesis by (subst (asm) of_nat_Suc, subst (asm) Gamma_fact) simp
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1163
qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1164
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1165
lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1166
  by (simp add: Gamma_of_int rGamma_inverse_Gamma)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1167
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1168
lemma Gamma_seriesI:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1169
  assumes "(\<lambda>n. g n / Gamma_series z n) \<longlonglongrightarrow> 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1170
  shows   "g \<longlonglongrightarrow> Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1171
proof (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1172
  have "1/2 > (0::real)" by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1173
  from tendstoD[OF assms, OF this]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1174
    show "eventually (\<lambda>n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1175
    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1176
  from assms have "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> 1 * Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1177
    by (intro tendsto_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1178
  thus "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> Gamma z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1179
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1180
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1181
lemma Gamma_seriesI':
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1182
  assumes "f \<longlonglongrightarrow> rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1183
  assumes "(\<lambda>n. g n * f n) \<longlonglongrightarrow> 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1184
  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1185
  shows   "g \<longlonglongrightarrow> Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1186
proof (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1187
  have "1/2 > (0::real)" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1188
  from tendstoD[OF assms(2), OF this] show "eventually (\<lambda>n. g n * f n / f n = g n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1189
    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1190
  from assms have "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> 1 / rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1191
    by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1192
  thus "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> Gamma z" by (simp add: Gamma_def divide_inverse)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1193
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1194
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1195
lemma Gamma_series'_LIMSEQ: "Gamma_series' z \<longlonglongrightarrow> Gamma z"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1196
  by (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series']
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1197
                                      Gamma_series'_nonpos_Ints_LIMSEQ[of z])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1198
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1199
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1200
subsection \<open>Differentiability\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1201
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1202
lemma has_field_derivative_rGamma_no_nonpos_int:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1203
  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1204
  shows   "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1205
proof (rule has_field_derivative_at_within)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1206
  from assms have "z \<noteq> - of_nat n" for n by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1207
  from differentiable_rGamma_aux1[OF this]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1208
    show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1209
         unfolding Digamma_def suminf_def sums_def[abs_def]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1210
                   has_field_derivative_def has_derivative_def netlimit_at
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1211
    by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1212
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1213
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1214
lemma has_field_derivative_rGamma_nonpos_int:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1215
  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1216
  apply (rule has_field_derivative_at_within)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1217
  using differentiable_rGamma_aux2[of n]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1218
  unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1219
  by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_prod) simp
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1220
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1221
lemma has_field_derivative_rGamma [derivative_intros]:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1222
  "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>norm z\<rfloor>) * fact (nat \<lfloor>norm z\<rfloor>)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1223
      else -rGamma z * Digamma z)) (at z within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1224
using has_field_derivative_rGamma_no_nonpos_int[of z A]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1225
      has_field_derivative_rGamma_nonpos_int[of "nat \<lfloor>norm z\<rfloor>" A]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1226
  by (auto elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1227
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1228
declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1229
declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1230
declare has_field_derivative_rGamma_nonpos_int [derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1231
declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1232
declare has_field_derivative_rGamma [derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1233
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1234
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1235
lemma has_field_derivative_Gamma [derivative_intros]:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1236
  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1237
  unfolding Gamma_def [abs_def]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1238
  by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1239
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1240
declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1241
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1242
(* TODO: Hide ugly facts properly *)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1243
hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1244
          differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1245
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1246
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1247
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1248
(* TODO: differentiable etc. *)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1249
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1250
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1251
subsection \<open>Continuity\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1252
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1253
lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1254
  by (rule DERIV_continuous_on has_field_derivative_rGamma)+
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1255
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1256
lemma continuous_on_Gamma [continuous_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A Gamma"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1257
  by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1258
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1259
lemma isCont_rGamma [continuous_intros]:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1260
  "isCont f z \<Longrightarrow> isCont (\<lambda>x. rGamma (f x)) z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1261
  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_rGamma]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1262
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1263
lemma isCont_Gamma [continuous_intros]:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1264
  "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Gamma (f x)) z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1265
  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Gamma]])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1266
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1267
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1268
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1269
text \<open>The complex Gamma function\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1270
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1271
instantiation complex :: Gamma
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1272
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1273
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1274
definition rGamma_complex :: "complex \<Rightarrow> complex" where
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1275
  "rGamma_complex z = lim (rGamma_series z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1276
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1277
lemma rGamma_series_complex_converges:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1278
        "convergent (rGamma_series (z :: complex))" (is "?thesis1")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1279
  and rGamma_complex_altdef:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1280
        "rGamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1281
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1282
  have "?thesis1 \<and> ?thesis2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1283
  proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1284
    case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1285
    have "rGamma_series z \<longlonglongrightarrow> exp (- ln_Gamma z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1286
    proof (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1287
      from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1288
      from this(1) uniformly_convergent_imp_convergent[OF this(2), of z]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1289
        have "ln_Gamma_series z \<longlonglongrightarrow> lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1290
      thus "(\<lambda>n. exp (-ln_Gamma_series z n)) \<longlonglongrightarrow> exp (- ln_Gamma z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1291
        unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1292
      from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1293
        show "eventually (\<lambda>n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1294
        by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1295
    qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1296
    with False show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1297
      by (auto simp: convergent_def rGamma_complex_def intro!: limI)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1298
  next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1299
    case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1300
    then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1301
    also have "rGamma_series \<dots> \<longlonglongrightarrow> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1302
      by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1303
    finally show ?thesis using True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1304
      by (auto simp: rGamma_complex_def convergent_def intro!: limI)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1305
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1306
  thus "?thesis1" "?thesis2" by blast+
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1307
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1308
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1309
context
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1310
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1311
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1312
(* TODO: duplication *)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1313
private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1314
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1315
  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1316
  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1317
    using eventually_gt_at_top[of "0::nat"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1318
  proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1319
    fix n :: nat assume n: "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1320
    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1321
             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1322
      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1323
    also from n have "\<dots> = ?f n * rGamma_series z n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1324
      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1325
    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1326
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1327
  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1328
    using rGamma_series_complex_converges
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1329
    by (intro tendsto_intros lim_inverse_n)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1330
       (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1331
  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1332
  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1333
    by (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1334
  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1335
    using rGamma_series_complex_converges
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1336
    by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1337
  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1338
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1339
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1340
private lemma has_field_derivative_rGamma_complex_no_nonpos_Int:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1341
  assumes "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1342
  shows   "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1343
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1344
  have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1345
  proof (subst DERIV_cong_ev[OF refl _ refl])
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1346
    from that have "eventually (\<lambda>t. t \<in> ball z (Re z/2)) (nhds z)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1347
      by (intro eventually_nhds_in_nhd) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1348
    thus "eventually (\<lambda>t. rGamma t = exp (- ln_Gamma t)) (nhds z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1349
      using no_nonpos_Int_in_ball_complex[OF that]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1350
      by (auto elim!: eventually_mono simp: rGamma_complex_altdef)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1351
  next
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1352
    have "z \<notin> \<real>\<^sub>\<le>\<^sub>0" using that by (simp add: complex_nonpos_Reals_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1353
    with that show "((\<lambda>t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1354
     by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros simp: rGamma_complex_altdef)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1355
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1356
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1357
  from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1358
  proof (induction "nat \<lfloor>1 - Re z\<rfloor>" arbitrary: z)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1359
    case (Suc n z)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1360
    from Suc.prems have z: "z \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1361
    from Suc.hyps have "n = nat \<lfloor>- Re z\<rfloor>" by linarith
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1362
    hence A: "n = nat \<lfloor>1 - Re (z + 1)\<rfloor>" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1363
    from Suc.prems have B: "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1364
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1365
    have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1)) z) has_field_derivative
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1366
      -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1367
      by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1368
    also have "(\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) = rGamma"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1369
      by (simp add: rGamma_complex_plus1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1370
    also from z have "Digamma (z + 1) * z - 1 = z * Digamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1371
      by (subst Digamma_plus1) (simp_all add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1372
    also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1373
      by (simp add: rGamma_complex_plus1[of z, symmetric])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1374
    finally show ?case .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1375
  qed (intro diff, simp)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1376
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1377
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1378
private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1379
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1380
  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1381
    using eventually_gt_at_top[of "0::nat"]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1382
    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1383
                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1384
  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1385
  thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1386
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1387
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1388
private lemma has_field_derivative_rGamma_complex_nonpos_Int:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1389
  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1390
proof (induction n)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1391
  case 0
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1392
  have A: "(0::complex) + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1393
  have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1394
    by (rule derivative_eq_intros DERIV_chain refl
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1395
             has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1396
    thus ?case by (simp add: rGamma_complex_plus1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1397
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1398
  case (Suc n)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1399
  hence A: "(rGamma has_field_derivative (-1)^n * fact n)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1400
                (at (- of_nat (Suc n) + 1 :: complex))" by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1401
   have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1402
             (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1403
     by (rule derivative_eq_intros refl A DERIV_chain)+
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1404
        (simp add: algebra_simps rGamma_complex_altdef)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1405
  thus ?case by (simp add: rGamma_complex_plus1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1406
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1407
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1408
instance proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1409
  fix z :: complex show "(rGamma z = 0) \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1410
    by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1411
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1412
  fix z :: complex assume "\<And>n. z \<noteq> - of_nat n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1413
  hence "z \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1414
  from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1415
    show "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k))
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1416
                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma z +
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1417
              rGamma z * d * (y - z)) /\<^sub>R  cmod (y - z)) \<midarrow>z\<rightarrow> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1418
      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1419
                    netlimit_at of_real_def[symmetric] suminf_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1420
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1421
  fix n :: nat
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1422
  from has_field_derivative_rGamma_complex_nonpos_Int[of n]
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1423
  show "let z = - of_nat n in (\<lambda>y. (rGamma y - rGamma z - (- 1) ^ n * prod of_nat {1..n} *
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1424
                  (y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1425
    by (simp add: has_field_derivative_def has_derivative_def fact_prod netlimit_at Let_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1426
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1427
  fix z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1428
  from rGamma_series_complex_converges[of z] have "rGamma_series z \<longlonglongrightarrow> rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1429
    by (simp add: convergent_LIMSEQ_iff rGamma_complex_def)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1430
  thus "let fact' = \<lambda>n. prod of_nat {1..n};
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1431
            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1432
            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1433
        in  (\<lambda>n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma z"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1434
    by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63317
diff changeset
  1435
                  of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1436
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1437
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1438
end
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1439
end
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1440
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1441
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1442
lemma Gamma_complex_altdef:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1443
  "Gamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1444
  unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1445
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1446
lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1447
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1448
  have "rGamma_series (cnj z) = (\<lambda>n. cnj (rGamma_series z n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1449
    by (intro ext) (simp_all add: rGamma_series_def exp_cnj)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1450
  also have "... \<longlonglongrightarrow> cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1451
  finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1452
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1453
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1454
lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1455
  unfolding Gamma_def by (simp add: cnj_rGamma)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1456
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1457
lemma Gamma_complex_real:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1458
  "z \<in> \<real> \<Longrightarrow> Gamma z \<in> (\<real> :: complex set)" and rGamma_complex_real: "z \<in> \<real> \<Longrightarrow> rGamma z \<in> \<real>"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1459
  by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1460
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1461
lemma field_differentiable_rGamma: "rGamma field_differentiable (at z within A)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1462
  using has_field_derivative_rGamma[of z] unfolding field_differentiable_def by blast
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1463
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1464
lemma holomorphic_rGamma [holomorphic_intros]: "rGamma holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1465
  unfolding holomorphic_on_def by (auto intro!: field_differentiable_rGamma)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1466
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1467
lemma analytic_rGamma: "rGamma analytic_on A"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1468
  unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_rGamma)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1469
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1470
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1471
lemma field_differentiable_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma field_differentiable (at z within A)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1472
  using has_field_derivative_Gamma[of z] unfolding field_differentiable_def by auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1473
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1474
lemma holomorphic_Gamma [holomorphic_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1475
  unfolding holomorphic_on_def by (auto intro!: field_differentiable_Gamma)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1476
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1477
lemma analytic_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma analytic_on A"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1478
  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1479
     (auto intro!: holomorphic_Gamma)
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1480
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1481
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1482
lemma field_differentiable_ln_Gamma_complex:
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1483
  "z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma field_differentiable (at (z::complex) within A)"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1484
  by (rule field_differentiable_within_subset[of _ _ UNIV])
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1485
     (force simp: field_differentiable_def intro!: derivative_intros)+
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1486
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1487
lemma holomorphic_ln_Gamma [holomorphic_intros]: "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> ln_Gamma holomorphic_on A"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1488
  unfolding holomorphic_on_def by (auto intro!: field_differentiable_ln_Gamma_complex)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1489
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1490
lemma analytic_ln_Gamma: "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> ln_Gamma analytic_on A"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1491
  by (rule analytic_on_subset[of _ "UNIV - \<real>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1492
     (auto intro!: holomorphic_ln_Gamma)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1493
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1494
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1495
lemma has_field_derivative_rGamma_complex' [derivative_intros]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1496
  "(rGamma has_field_derivative (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-Re z\<rfloor>) * fact (nat \<lfloor>-Re z\<rfloor>) else
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1497
        -rGamma z * Digamma z)) (at z within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1498
  using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1499
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1500
declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1501
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1502
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1503
lemma field_differentiable_Polygamma:
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1504
  fixes z :: complex
62533
bc25f3916a99 new material to Blochj's theorem, as well as supporting lemmas
paulson <lp15@cam.ac.uk>
parents: 62398
diff changeset
  1505
  shows
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1506
  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Polygamma n field_differentiable (at z within A)"
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1507
  using has_field_derivative_Polygamma[of z n] unfolding field_differentiable_def by auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1508
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1509
lemma holomorphic_on_Polygamma [holomorphic_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n holomorphic_on A"
62534
6855b348e828 complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
paulson <lp15@cam.ac.uk>
parents: 62533
diff changeset
  1510
  unfolding holomorphic_on_def by (auto intro!: field_differentiable_Polygamma)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1511
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1512
lemma analytic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n analytic_on A"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1513
  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1514
     (auto intro!: holomorphic_on_Polygamma)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1515
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1516
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1517
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1518
text \<open>The real Gamma function\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1519
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1520
lemma rGamma_series_real:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1521
  "eventually (\<lambda>n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1522
  using eventually_gt_at_top[of "0 :: nat"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1523
proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1524
  fix n :: nat assume n: "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1525
  have "Re (rGamma_series (of_real x) n) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1526
          Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1527
    using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1528
  also from n have "\<dots> = Re (of_real ((pochhammer x (Suc n)) /
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1529
                              (fact n * (exp (x * ln (real_of_nat n))))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1530
    by (subst exp_of_real) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1531
  also from n have "\<dots> = rGamma_series x n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1532
    by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1533
  finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" ..
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1534
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1535
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1536
instantiation real :: Gamma
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1537
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1538
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1539
definition "rGamma_real x = Re (rGamma (of_real x :: complex))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1540
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1541
instance proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1542
  fix x :: real
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1543
  have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1544
  also have "of_real \<dots> = rGamma (of_real x :: complex)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1545
    by (intro of_real_Re rGamma_complex_real) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1546
  also have "\<dots> = 0 \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0" by (simp add: rGamma_eq_zero_iff of_real_in_nonpos_Ints_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1547
  also have "\<dots> \<longleftrightarrow> (\<exists>n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1548
  finally show "(rGamma x) = 0 \<longleftrightarrow> (\<exists>n. x = - real_of_nat n)" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1549
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1550
  fix x :: real assume "\<And>n. x \<noteq> - of_nat n"
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63417
diff changeset
  1551
  hence x: "complex_of_real x \<notin> \<int>\<^sub>\<le>\<^sub>0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1552
    by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases')
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63417
diff changeset
  1553
  then have "x \<noteq> 0" by auto
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63417
diff changeset
  1554
  with x have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1555
    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1556
                  simp: Polygamma_of_real rGamma_real_def [abs_def])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1557
  thus "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (x + of_nat k))
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1558
                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma x +
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1559
              rGamma x * d * (y - x)) /\<^sub>R  norm (y - x)) \<midarrow>x\<rightarrow> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1560
      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1561
                    netlimit_at of_real_def[symmetric] suminf_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1562
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1563
  fix n :: nat
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1564
  have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1565
    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1566
                  simp: Polygamma_of_real rGamma_real_def [abs_def])
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1567
  thus "let x = - of_nat n in (\<lambda>y. (rGamma y - rGamma x - (- 1) ^ n * prod of_nat {1..n} *
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1568
                  (y - x)) /\<^sub>R norm (y - x)) \<midarrow>x::real\<rightarrow> 0"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1569
    by (simp add: has_field_derivative_def has_derivative_def fact_prod netlimit_at Let_def)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1570
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1571
  fix x :: real
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1572
  have "rGamma_series x \<longlonglongrightarrow> rGamma x"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1573
  proof (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1574
    show "(\<lambda>n. Re (rGamma_series (of_real x) n)) \<longlonglongrightarrow> rGamma x" unfolding rGamma_real_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1575
      by (intro tendsto_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1576
  qed (insert rGamma_series_real, simp add: eq_commute)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1577
  thus "let fact' = \<lambda>n. prod of_nat {1..n};
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1578
            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1579
            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1580
        in  (\<lambda>n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma x"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  1581
    by (simp add: fact_prod pochhammer_Suc_prod rGamma_series_def [abs_def] exp_def
63367
6c731c8b7f03 simplified definitions of combinatorial functions
haftmann
parents: 63317
diff changeset
  1582
                  of_real_def [symmetric] suminf_def sums_def [abs_def] atLeast0AtMost)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1583
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1584
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1585
end
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1586
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1587
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1588
lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1589
  unfolding rGamma_real_def using rGamma_complex_real by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1590
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1591
lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1592
  unfolding Gamma_def by (simp add: rGamma_complex_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1593
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1594
lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1595
  by (rule sym, rule limI, rule tendsto_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1596
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1597
lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1598
  by (rule sym, rule limI, rule tendsto_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1599
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1600
lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1601
  using rGamma_complex_real[OF Reals_of_real[of x]]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1602
  by (elim Reals_cases)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1603
     (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1604
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1605
lemma ln_Gamma_series_complex_of_real:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1606
  "x > 0 \<Longrightarrow> n > 0 \<Longrightarrow> ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1607
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1608
  assume xn: "x > 0" "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1609
  have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \<ge> 1" for k
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1610
    using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1611
  with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1612
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1613
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1614
lemma ln_Gamma_real_converges:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1615
  assumes "(x::real) > 0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1616
  shows   "convergent (ln_Gamma_series x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1617
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1618
  have "(\<lambda>n. ln_Gamma_series (complex_of_real x) n) \<longlonglongrightarrow> ln_Gamma (of_real x)" using assms
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1619
    by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1620
  moreover from eventually_gt_at_top[of "0::nat"]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1621
    have "eventually (\<lambda>n. complex_of_real (ln_Gamma_series x n) =
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1622
            ln_Gamma_series (complex_of_real x) n) sequentially"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1623
    by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1624
  ultimately have "(\<lambda>n. complex_of_real (ln_Gamma_series x n)) \<longlonglongrightarrow> ln_Gamma (of_real x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1625
    by (subst tendsto_cong) assumption+
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1626
  from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1627
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1628
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1629
lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \<Longrightarrow> ln_Gamma_series x \<longlonglongrightarrow> ln_Gamma x"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1630
  using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1631
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1632
lemma ln_Gamma_complex_of_real: "x > 0 \<Longrightarrow> ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1633
proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1634
  assume x: "x > 0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1635
  show "eventually (\<lambda>n. of_real (ln_Gamma_series x n) =
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1636
            ln_Gamma_series (complex_of_real x) n) sequentially"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1637
    using eventually_gt_at_top[of "0::nat"]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1638
    by eventually_elim (simp add: ln_Gamma_series_complex_of_real x)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1639
qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1640
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1641
lemma Gamma_real_pos_exp: "x > (0 :: real) \<Longrightarrow> Gamma x = exp (ln_Gamma x)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1642
  by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1643
                 ln_Gamma_complex_of_real exp_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1644
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1645
lemma ln_Gamma_real_pos: "x > 0 \<Longrightarrow> ln_Gamma x = ln (Gamma x :: real)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1646
  unfolding Gamma_real_pos_exp by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1647
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1648
lemma ln_Gamma_complex_conv_fact: "n > 0 \<Longrightarrow> ln_Gamma (of_nat n :: complex) = ln (fact (n - 1))"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1649
  using ln_Gamma_complex_of_real[of "real n"] Gamma_fact[of "n - 1", where 'a = real]
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1650
  by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric])
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1651
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1652
lemma ln_Gamma_real_conv_fact: "n > 0 \<Longrightarrow> ln_Gamma (real n) = ln (fact (n - 1))"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1653
  using Gamma_fact[of "n - 1", where 'a = real]
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1654
  by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric])
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1655
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1656
lemma Gamma_real_pos: "x > (0::real) \<Longrightarrow> Gamma x > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1657
  by (simp add: Gamma_real_pos_exp)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1658
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1659
lemma has_field_derivative_ln_Gamma_real [derivative_intros]:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1660
  assumes x: "x > (0::real)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1661
  shows "(ln_Gamma has_field_derivative Digamma x) (at x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1662
proof (subst DERIV_cong_ev[OF refl _ refl])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1663
  from assms show "((Re \<circ> ln_Gamma \<circ> complex_of_real) has_field_derivative Digamma x) (at x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1664
    by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1665
             simp: Polygamma_of_real o_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1666
  from eventually_nhds_in_nhd[of x "{0<..}"] assms
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1667
    show "eventually (\<lambda>y. ln_Gamma y = (Re \<circ> ln_Gamma \<circ> of_real) y) (nhds x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1668
    by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1669
qed
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1670
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1671
lemma field_differentiable_ln_Gamma_real:
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1672
  "x > 0 \<Longrightarrow> ln_Gamma field_differentiable (at (x::real) within A)"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1673
  by (rule field_differentiable_within_subset[of _ _ UNIV])
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1674
     (auto simp: field_differentiable_def intro!: derivative_intros)+
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1675
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1676
declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros]
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1677
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1678
lemma deriv_ln_Gamma_real:
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1679
  assumes "z > 0"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1680
  shows   "deriv ln_Gamma z = Digamma (z :: real)"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  1681
  by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_real assms)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1682
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1683
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1684
lemma has_field_derivative_rGamma_real' [derivative_intros]:
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1685
  "(rGamma has_field_derivative (if x \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-x\<rfloor>) * fact (nat \<lfloor>-x\<rfloor>) else
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1686
        -rGamma x * Digamma x)) (at x within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1687
  using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1688
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1689
declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1690
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1691
lemma Polygamma_real_odd_pos:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1692
  assumes "(x::real) \<notin> \<int>\<^sub>\<le>\<^sub>0" "odd n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1693
  shows   "Polygamma n x > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1694
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1695
  from assms have "x \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1696
  with assms show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1697
    unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1698
    by (auto simp: zero_less_power_eq simp del: power_Suc
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1699
             dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1700
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1701
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1702
lemma Polygamma_real_even_neg:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1703
  assumes "(x::real) > 0" "n > 0" "even n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1704
  shows   "Polygamma n x < 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1705
  using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1706
  by (auto intro!: mult_pos_pos suminf_pos)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1707
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1708
lemma Polygamma_real_strict_mono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1709
  assumes "x > 0" "x < (y::real)" "even n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1710
  shows   "Polygamma n x < Polygamma n y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1711
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1712
  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1713
    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1714
  then guess \<xi> by (elim exE conjE) note \<xi> = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1715
  note \<xi>(3)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1716
  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> > 0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1717
    by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1718
  finally show ?thesis by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1719
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1720
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1721
lemma Polygamma_real_strict_antimono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1722
  assumes "x > 0" "x < (y::real)" "odd n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1723
  shows   "Polygamma n x > Polygamma n y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1724
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1725
  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1726
    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1727
  then guess \<xi> by (elim exE conjE) note \<xi> = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1728
  note \<xi>(3)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1729
  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> < 0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1730
    by (intro mult_pos_neg Polygamma_real_even_neg) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1731
  finally show ?thesis by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1732
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1733
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1734
lemma Polygamma_real_mono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1735
  assumes "x > 0" "x \<le> (y::real)" "even n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1736
  shows   "Polygamma n x \<le> Polygamma n y"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1737
  using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1738
  by (cases "x = y") simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1739
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1740
lemma Digamma_real_strict_mono: "(0::real) < x \<Longrightarrow> x < y \<Longrightarrow> Digamma x < Digamma y"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1741
  by (rule Polygamma_real_strict_mono) simp_all
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1742
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1743
lemma Digamma_real_mono: "(0::real) < x \<Longrightarrow> x \<le> y \<Longrightarrow> Digamma x \<le> Digamma y"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1744
  by (rule Polygamma_real_mono) simp_all
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  1745
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1746
lemma Digamma_real_ge_three_halves_pos:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1747
  assumes "x \<ge> 3/2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1748
  shows   "Digamma (x :: real) > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1749
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1750
  have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1751
  also from assms have "\<dots> \<le> Digamma x" by (intro Polygamma_real_mono) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1752
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1753
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1754
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1755
lemma ln_Gamma_real_strict_mono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1756
  assumes "x \<ge> 3/2" "x < y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1757
  shows   "ln_Gamma (x :: real) < ln_Gamma y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1758
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1759
  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> ln_Gamma y - ln_Gamma x = (y - x) * Digamma \<xi>"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1760
    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1761
  then guess \<xi> by (elim exE conjE) note \<xi> = this
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1762
  note \<xi>(3)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1763
  also from \<xi>(1,2) assms have "(y - x) * Digamma \<xi> > 0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1764
    by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1765
  finally show ?thesis by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1766
qed
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1767
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1768
lemma Gamma_real_strict_mono:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1769
  assumes "x \<ge> 3/2" "x < y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1770
  shows   "Gamma (x :: real) < Gamma y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1771
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1772
  from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1773
  also have "\<dots> < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1774
  also from Gamma_real_pos_exp[of y] assms have "\<dots> = Gamma y" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1775
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1776
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1777
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1778
lemma log_convex_Gamma_real: "convex_on {0<..} (ln \<circ> Gamma :: real \<Rightarrow> real)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1779
  by (rule convex_on_realI[of _ _ Digamma])
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1780
     (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1781
           simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1782
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1783
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1784
subsection \<open>The uniqueness of the real Gamma function\<close>
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1785
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1786
text \<open>
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1787
  The following is a proof of the Bohr--Mollerup theorem, which states that
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1788
  any log-convex function $G$ on the positive reals that fulfils $G(1) = 1$ and
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1789
  satisfies the functional equation $G(x+1) = x G(x)$ must be equal to the
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1790
  Gamma function.
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1791
  In principle, if $G$ is a holomorphic complex function, one could then extend
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1792
  this from the positive reals to the entire complex plane (minus the non-positive
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1793
  integers, where the Gamma function is not defined).
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1794
\<close>
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1795
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1796
context
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1797
  fixes G :: "real \<Rightarrow> real"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1798
  assumes G_1: "G 1 = 1"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1799
  assumes G_plus1: "x > 0 \<Longrightarrow> G (x + 1) = x * G x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1800
  assumes G_pos: "x > 0 \<Longrightarrow> G x > 0"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1801
  assumes log_convex_G: "convex_on {0<..} (ln \<circ> G)"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1802
begin
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1803
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1804
private lemma G_fact: "G (of_nat n + 1) = fact n"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1805
  using G_plus1[of "real n + 1" for n]
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1806
  by (induction n) (simp_all add: G_1 G_plus1)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1807
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1808
private definition S :: "real \<Rightarrow> real \<Rightarrow> real" where
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1809
  "S x y = (ln (G y) - ln (G x)) / (y - x)"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1810
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1811
private lemma S_eq:
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1812
  "n \<ge> 2 \<Longrightarrow> S (of_nat n) (of_nat n + x) = (ln (G (real n + x)) - ln (fact (n - 1))) / x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1813
  by (subst G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1814
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1815
private lemma G_lower:
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1816
  assumes x: "x > 0" and n: "n \<ge> 1"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1817
  shows  "Gamma_series x n \<le> G x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1818
proof -
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1819
  have "(ln \<circ> G) (real (Suc n)) \<le> ((ln \<circ> G) (real (Suc n) + x) -
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1820
          (ln \<circ> G) (real (Suc n) - 1)) / (real (Suc n) + x - (real (Suc n) - 1)) *
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1821
           (real (Suc n) - (real (Suc n) - 1)) + (ln \<circ> G) (real (Suc n) - 1)"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1822
    using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1823
  hence "S (of_nat n) (of_nat (Suc n)) \<le> S (of_nat (Suc n)) (of_nat (Suc n) + x)"
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1824
    unfolding S_def using x by (simp add: field_simps)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1825
  also have "S (of_nat n) (of_nat (Suc n)) = ln (fact n) - ln (fact (n-1))"
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1826
    unfolding S_def using n
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1827
    by (subst (1 2) G_fact [symmetric]) (simp_all add: add_ac of_nat_diff)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1828
  also have "\<dots> = ln (fact n / fact (n-1))" by (subst ln_div) simp_all
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1829
  also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1830
  finally have "x * ln (real n) + ln (fact n) \<le> ln (G (real (Suc n) + x))"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1831
    using x n by (subst (asm) S_eq) (simp_all add: field_simps)
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1832
  also have "x * ln (real n) + ln (fact n) = ln (exp (x * ln (real n)) * fact n)"
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1833
    using x by (simp add: ln_mult)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1834
  finally have "exp (x * ln (real n)) * fact n \<le> G (real (Suc n) + x)" using x
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1835
    by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1836
  also have "G (real (Suc n) + x) = pochhammer x (Suc n) * G x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1837
    using G_plus1[of "real (Suc n) + x" for n] G_plus1[of x] x
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1838
    by (induction n) (simp_all add: pochhammer_Suc add_ac)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1839
  finally show "Gamma_series x n \<le> G x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1840
    using x by (simp add: field_simps pochhammer_pos Gamma_series_def)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1841
qed
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1842
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1843
private lemma G_upper:
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1844
  assumes x: "x > 0" "x \<le> 1" and n: "n \<ge> 2"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1845
  shows  "G x \<le> Gamma_series x n * (1 + x / real n)"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1846
proof -
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1847
  have "(ln \<circ> G) (real n + x) \<le> ((ln \<circ> G) (real n + 1) -
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1848
          (ln \<circ> G) (real n)) / (real n + 1 - (real n)) *
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1849
           ((real n + x) - real n) + (ln \<circ> G) (real n)"
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1850
    using x n by (intro convex_onD_Icc' convex_on_subset[OF log_convex_G]) auto
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1851
  hence "S (of_nat n) (of_nat n + x) \<le> S (of_nat n) (of_nat n + 1)"
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1852
    unfolding S_def using x by (simp add: field_simps)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1853
  also from n have "S (of_nat n) (of_nat n + 1) = ln (fact n) - ln (fact (n-1))"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1854
    by (subst (1 2) G_fact [symmetric]) (simp add: S_def add_ac of_nat_diff)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1855
  also have "\<dots> = ln (fact n / (fact (n-1)))" using n by (subst ln_div) simp_all
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1856
  also from n have "fact n / fact (n - 1) = n" by (cases n) simp_all
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1857
  finally have "ln (G (real n + x)) \<le> x * ln (real n) + ln (fact (n - 1))"
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1858
    using x n by (subst (asm) S_eq) (simp_all add: field_simps)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1859
  also have "\<dots> = ln (exp (x * ln (real n)) * fact (n - 1))" using x
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1860
    by (simp add: ln_mult)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1861
  finally have "G (real n + x) \<le> exp (x * ln (real n)) * fact (n - 1)" using x
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1862
    by (subst (asm) ln_le_cancel_iff) (simp_all add: G_pos)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1863
  also have "G (real n + x) = pochhammer x n * G x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1864
    using G_plus1[of "real n + x" for n] x
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1865
    by (induction n) (simp_all add: pochhammer_Suc add_ac)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1866
  finally have "G x \<le> exp (x * ln (real n)) * fact (n- 1) / pochhammer x n"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1867
    using x by (simp add: field_simps pochhammer_pos)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1868
  also from n have "fact (n - 1) = fact n / n" by (cases n) simp_all
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1869
  also have "exp (x * ln (real n)) * \<dots> / pochhammer x n =
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1870
               Gamma_series x n * (1 + x / real n)" using n x
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1871
    by (simp add: Gamma_series_def divide_simps pochhammer_Suc)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1872
  finally show ?thesis .
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1873
qed
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1874
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1875
private lemma G_eq_Gamma_aux:
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1876
  assumes x: "x > 0" "x \<le> 1"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1877
  shows   "G x = Gamma x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1878
proof (rule antisym)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1879
  show "G x \<ge> Gamma x"
63952
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1880
  proof (rule tendsto_upperbound)
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1881
    from G_lower[of x] show "eventually (\<lambda>n. Gamma_series x n \<le> G x) sequentially"
65578
e4997c181cce New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents: 64969
diff changeset
  1882
      using  x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "1::nat"]])
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1883
  qed (simp_all add: Gamma_series_LIMSEQ)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1884
next
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1885
  show "G x \<le> Gamma x"
63952
354808e9f44b new material connected with HOL Light measure theory, plus more rationalisation
paulson <lp15@cam.ac.uk>
parents: 63918
diff changeset
  1886
  proof (rule tendsto_lowerbound)
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1887
    have "(\<lambda>n. Gamma_series x n * (1 + x / real n)) \<longlonglongrightarrow> Gamma x * (1 + 0)"
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1888
      by (rule tendsto_intros real_tendsto_divide_at_top
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1889
               Gamma_series_LIMSEQ filterlim_real_sequentially)+
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1890
    thus "(\<lambda>n. Gamma_series x n * (1 + x / real n)) \<longlonglongrightarrow> Gamma x" by simp
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1891
  next
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1892
    from G_upper[of x] show "eventually (\<lambda>n. Gamma_series x n * (1 + x / real n) \<ge> G x) sequentially"
65578
e4997c181cce New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents: 64969
diff changeset
  1893
      using x by (auto intro: eventually_mono[OF eventually_ge_at_top[of "2::nat"]])
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1894
  qed simp_all
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1895
qed
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1896
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1897
theorem Gamma_pos_real_unique:
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1898
  assumes x: "x > 0"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1899
  shows   "G x = Gamma x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1900
proof -
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1901
  have G_eq: "G (real n + x) = Gamma (real n + x)" if "x \<in> {0<..1}" for n x using that
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1902
  proof (induction n)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1903
    case (Suc n)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1904
    from Suc have "x + real n > 0" by simp
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1905
    hence "x + real n \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1906
    with Suc show ?case using G_plus1[of "real n + x"] Gamma_plus1[of "real n + x"]
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1907
      by (auto simp: add_ac)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1908
  qed (simp_all add: G_eq_Gamma_aux)
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  1909
63725
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1910
  show ?thesis
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1911
  proof (cases "frac x = 0")
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1912
    case True
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1913
    hence "x = of_int (floor x)" by (simp add: frac_def)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1914
    with x have x_eq: "x = of_nat (nat (floor x) - 1) + 1" by simp
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1915
    show ?thesis by (subst (1 2) x_eq, rule G_eq) simp_all
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1916
  next
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1917
    case False
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1918
    from assms have x_eq: "x = of_nat (nat (floor x)) + frac x"
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1919
      by (simp add: frac_def)
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1920
    have frac_le_1: "frac x \<le> 1" unfolding frac_def by linarith
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1921
    show ?thesis
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1922
      by (subst (1 2) x_eq, rule G_eq, insert False frac_le_1) simp_all
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1923
  qed
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1924
qed
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1925
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1926
end
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1927
4c00ba1ad11a Bohr-Mollerup theorem for the Gamma function
Manuel Eberl <eberlm@in.tum.de>
parents: 63721
diff changeset
  1928
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1929
subsection \<open>Beta function\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1930
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1931
definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1932
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1933
lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1934
  by (simp add: inverse_eq_divide Beta_def Gamma_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1935
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1936
lemma Beta_commute: "Beta a b = Beta b a"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1937
  unfolding Beta_def by (simp add: ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1938
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1939
lemma has_field_derivative_Beta1 [derivative_intros]:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1940
  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1941
  shows   "((\<lambda>x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y))))
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1942
               (at x within A)" unfolding Beta_altdef
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1943
  by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1944
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1945
lemma Beta_pole1: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1946
  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1947
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1948
lemma Beta_pole2: "y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1949
  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  1950
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1951
lemma Beta_zero: "x + y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Beta x y = 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1952
  by (auto simp add: Beta_def elim!: nonpos_Ints_cases')
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  1953
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1954
lemma has_field_derivative_Beta2 [derivative_intros]:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1955
  assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1956
  shows   "((\<lambda>y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y))))
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1957
               (at y within A)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1958
  using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1959
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1960
lemma Beta_plus1_plus1:
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1961
  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1962
  shows   "Beta (x + 1) y + Beta x (y + 1) = Beta x y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1963
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1964
  have "Beta (x + 1) y + Beta x (y + 1) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1965
            (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1966
    by (simp add: Beta_altdef add_divide_distrib algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1967
  also have "\<dots> = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1968
    by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1969
  also from assms have "\<dots> = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1970
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1971
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1972
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1973
lemma Beta_plus1_left:
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1974
  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1975
  shows   "(x + y) * Beta (x + 1) y = x * Beta x y"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1976
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1977
  have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1978
    unfolding Beta_altdef by (simp only: ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1979
  also have "\<dots> = x * Beta x y" unfolding Beta_altdef
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1980
     by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1981
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1982
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1983
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  1984
lemma Beta_plus1_right:
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1985
  assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1986
  shows   "(x + y) * Beta x (y + 1) = y * Beta x y"
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1987
  using Beta_plus1_left[of y x] assms by (simp_all add: Beta_commute add.commute)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1988
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1989
lemma Gamma_Gamma_Beta:
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  1990
  assumes "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1991
  shows   "Gamma x * Gamma y = Beta x y * Gamma (x + y)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1992
  unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1993
  by (simp add: rGamma_inverse_Gamma)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1994
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1995
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1996
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1997
subsection \<open>Legendre duplication theorem\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1998
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  1999
context
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2000
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2001
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2002
private lemma Gamma_legendre_duplication_aux:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2003
  fixes z :: "'a :: Gamma"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2004
  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2005
  shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2006
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2007
  let ?powr = "\<lambda>b a. exp (a * of_real (ln (of_nat b)))"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2008
  let ?h = "\<lambda>n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) *
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2009
                exp (1/2 * of_real (ln (real_of_nat n)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2010
  {
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2011
    fix z :: 'a assume z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2012
    let ?g = "\<lambda>n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n /
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2013
                      Gamma_series' (2*z) (2*n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2014
    have "eventually (\<lambda>n. ?g n = ?h n) sequentially" using eventually_gt_at_top
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2015
    proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2016
      fix n :: nat assume n: "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2017
      let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2018
      have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2019
      have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) /
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2020
                (pochhammer z n * pochhammer (z + 1/2) n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2021
        by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2022
      have B: "Gamma_series' (2*z) (2*n) =
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2023
                       ?f' * ?powr 2 (2*z) * ?powr n (2*z) /
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2024
                       (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2025
        by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2026
      from z have "pochhammer z n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2027
      moreover from z have "pochhammer (z + 1/2) n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2028
      ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2029
         ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2030
        using n unfolding A B by (simp add: divide_simps exp_minus)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2031
      also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2032
        by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2033
      finally show "?g n = ?h n" by (simp only: mult_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2034
    qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2035
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2036
    moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2037
    hence "?g \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
62397
5ae24f33d343 Substantial new material for multivariate analysis. Also removal of some duplicates.
paulson <lp15@cam.ac.uk>
parents: 62131
diff changeset
  2038
      using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "op*2" "2*z"]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2039
      by (intro tendsto_intros Gamma_series'_LIMSEQ)
66447
a1f5c5c26fa6 Replaced subseq with strict_mono
eberlm <eberlm@in.tum.de>
parents: 66286
diff changeset
  2040
         (simp_all add: o_def strict_mono_def Gamma_eq_zero_iff)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2041
    ultimately have "?h \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2042
      by (rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2043
  } note lim = this
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2044
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2045
  from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2046
  from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2047
    by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2048
  with lim[of "1/2 :: 'a"] have "?h \<longlonglongrightarrow> 2 * Gamma (1 / 2 :: 'a)" by (simp add: exp_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2049
  from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2050
    by (simp add: divide_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2051
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2052
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2053
lemma Gamma_reflection_complex:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2054
  fixes z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2055
  shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2056
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2057
  let ?g = "\<lambda>z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  2058
  define g where [abs_def]: "g z = (if z \<in> \<int> then of_real pi else ?g z)" for z :: complex
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2059
  let ?h = "\<lambda>z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  2060
  define h where [abs_def]: "h z = (if z \<in> \<int> then 0 else ?h z)" for z :: complex
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2061
62072
bf3d9f113474 isabelle update_cartouches -c -t;
wenzelm
parents: 62055
diff changeset
  2062
  \<comment> \<open>@{term g} is periodic with period 1.\<close>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2063
  interpret g: periodic_fun_simple' g
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2064
  proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2065
    fix z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2066
    show "g (z + 1) = g z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2067
    proof (cases "z \<in> \<int>")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2068
      case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2069
      hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2070
      also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2071
        using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2072
        by (subst Beta_plus1_left [symmetric]) auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2073
      also have "\<dots> * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2074
        using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2075
        by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2076
      also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2077
        using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2078
      finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2079
    qed (simp add: g_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2080
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2081
62072
bf3d9f113474 isabelle update_cartouches -c -t;
wenzelm
parents: 62055
diff changeset
  2082
  \<comment> \<open>@{term g} is entire.\<close>
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2083
  have g_g' [derivative_intros]: "(g has_field_derivative (h z * g z)) (at z)" for z :: complex
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2084
  proof (cases "z \<in> \<int>")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2085
    let ?h' = "\<lambda>z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) +
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2086
                     of_real pi * cos (z * of_real pi))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2087
    case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2088
    from False have "eventually (\<lambda>t. t \<in> UNIV - \<int>) (nhds z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2089
      by (intro eventually_nhds_in_open) (auto simp: open_Diff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2090
    hence "eventually (\<lambda>t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2091
    moreover {
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2092
      from False Ints_diff[of 1 "1-z"] have "1 - z \<notin> \<int>" by auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2093
      hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2094
        by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2095
      also from False have "sin (of_real pi * z) \<noteq> 0" by (subst sin_eq_0) auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2096
      hence "?h' z = h z * g z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2097
        using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2098
      finally have "(?g has_field_derivative (h z * g z)) (at z)" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2099
    }
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2100
    ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2101
  next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2102
    case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2103
    then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2104
    let ?t = "(\<lambda>z::complex. if z = 0 then 1 else sin z / z) \<circ> (\<lambda>z. of_real pi * z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2105
    have deriv_0: "(g has_field_derivative 0) (at 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2106
    proof (subst DERIV_cong_ev[OF refl _ refl])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2107
      show "eventually (\<lambda>z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2108
        using eventually_nhds_ball[OF zero_less_one, of "0::complex"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2109
      proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2110
        fix z :: complex assume z: "z \<in> ball 0 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2111
        show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2112
        proof (cases "z = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2113
          assume z': "z \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2114
          with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases simp: dist_0_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2115
          from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2116
          with z'' z' show ?thesis by (simp add: g_def ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2117
        qed (simp add: g_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2118
      qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2119
      have "(?t has_field_derivative (0 * of_real pi)) (at 0)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2120
        using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2121
        by (intro DERIV_chain) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2122
      thus "((\<lambda>z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2123
        by (auto intro!: derivative_eq_intros simp: o_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2124
    qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2125
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2126
    have "((g \<circ> (\<lambda>x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2127
      using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2128
    also have "g \<circ> (\<lambda>x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2129
    finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2130
  qed
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  2131
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2132
  have g_holo [holomorphic_intros]: "g holomorphic_on A" for A
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  2133
    by (rule holomorphic_on_subset[of _ UNIV])
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2134
       (force simp: holomorphic_on_open intro!: derivative_intros)+
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2135
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2136
  have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" if "Re z > -1" "Re z < 2" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2137
  proof (cases "z \<in> \<int>")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2138
    case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2139
    with that have "z = 0 \<or> z = 1" by (force elim!: Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2140
    moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2141
      using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2142
    moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2143
        using fraction_not_in_ints[where 'a = complex, of 2 1]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2144
        by (simp add: g_def power2_eq_square Beta_def algebra_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2145
    ultimately show ?thesis by force
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2146
  next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2147
    case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2148
    hence z: "z/2 \<notin> \<int>" "(z+1)/2 \<notin> \<int>" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2149
    hence z': "z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "(z+1)/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2150
    from z have "1-z/2 \<notin> \<int>" "1-((z+1)/2) \<notin> \<int>"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2151
      using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2152
    hence z'': "1-z/2 \<notin> \<int>\<^sub>\<le>\<^sub>0" "1-((z+1)/2) \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2153
    from z have "g (z/2) * g ((z+1)/2) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2154
      (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) *
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2155
      (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2156
      by (simp add: g_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2157
    also from z' Gamma_legendre_duplication_aux[of "z/2"]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2158
      have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2159
      by (simp add: add_divide_distrib)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2160
    also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2161
      have "Gamma (1-z/2) * Gamma (1-(z+1)/2) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2162
              Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2163
      by (simp add: add_divide_distrib ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2164
    finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) *
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2165
                    (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2166
      by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2167
    also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2168
      using cos_sin_eq[of "- of_real pi * z/2", symmetric]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2169
      by (simp add: ring_distribs add_divide_distrib ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2170
    also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2171
      by (subst sin_times_cos) (simp add: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2172
    also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2173
      using \<open>z \<notin> \<int>\<close> by (simp add: g_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2174
    finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2175
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2176
  have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2177
  proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  2178
    define r where "r = \<lfloor>Re z / 2\<rfloor>"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2179
    have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2180
    also have "of_int (2*r) = 2 * of_int r" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2181
    also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2182
    hence "Gamma (1/2)^2 * g (z - 2 * of_int r) =
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2183
                   g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2184
      unfolding r_def by (intro g_eq[symmetric]) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2185
    also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2186
    also have "g \<dots> = g (z/2)" by (rule g.minus_of_int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2187
    also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2188
    also have "g \<dots> = g ((z+1)/2)" by (rule g.minus_of_int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2189
    finally show ?thesis ..
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2190
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2191
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2192
  have g_nz [simp]: "g z \<noteq> 0" for z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2193
  unfolding g_def using Ints_diff[of 1 "1 - z"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2194
    by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int)
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  2195
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2196
  have h_altdef: "h z = deriv g z / g z" for z :: complex
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2197
    using DERIV_imp_deriv[OF g_g', of z] by simp
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  2198
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2199
  have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2200
  proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2201
    have "((\<lambda>t. g (t/2) * g ((t+1)/2)) has_field_derivative
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2202
                       (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2203
      by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2204
    hence "((\<lambda>t. Gamma (1/2)^2 * g t) has_field_derivative
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2205
              Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2206
      by (subst (1 2) g_eq[symmetric]) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2207
    from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2208
      have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2209
      using fraction_not_in_ints[where 'a = complex, of 2 1]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2210
      by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2211
    moreover have "(g has_field_derivative (g z * h z)) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2212
      using g_g'[of z] by (simp add: ac_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2213
    ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2214
      by (intro DERIV_unique)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2215
    thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2216
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2217
64969
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2218
  have h_holo [holomorphic_intros]: "h holomorphic_on A" for A
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2219
    unfolding h_altdef [abs_def]
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2220
    by (rule holomorphic_on_subset[of _ UNIV]) (auto intro!: holomorphic_intros)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2221
  define h' where "h' = deriv h"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2222
  have h_h': "(h has_field_derivative h' z) (at z)" for z unfolding h'_def
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2223
    by (auto intro!: holomorphic_derivI[of _ UNIV] holomorphic_intros)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2224
  have h'_holo [holomorphic_intros]: "h' holomorphic_on A" for A unfolding h'_def
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2225
    by (rule holomorphic_on_subset[of _ UNIV]) (auto intro!: holomorphic_intros)
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2226
  have h'_cont: "continuous_on UNIV h'"
a6953714799d Simplified Gamma_Function
eberlm <eberlm@in.tum.de>
parents: 64272
diff changeset
  2227
    by (intro holomorphic_on_imp_continuous_on holomorphic_intros)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2228
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2229
  have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2230
  proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2231
    have "((\<lambda>t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2232
                       ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2233
      by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2234
    hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2235
      by (subst (asm) h_eq[symmetric])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2236
    from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2237
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2238
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2239
  have h'_zero: "h' z = 0" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2240
  proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  2241
    define m where "m = max 1 \<bar>Re z\<bar>"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  2242
    define B where "B = {t. abs (Re t) \<le> m \<and> abs (Im t) \<le> abs (Im z)}"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2243
    have "closed ({t. Re t \<ge> -m} \<inter> {t. Re t \<le> m} \<inter>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2244
                  {t. Im t \<ge> -\<bar>Im z\<bar>} \<inter> {t. Im t \<le> \<bar>Im z\<bar>})"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2245
      (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2246
                                 closed_halfspace_Im_ge closed_halfspace_Im_le)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2247
    also have "?B = B" unfolding B_def by fastforce
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2248
    finally have "closed B" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2249
    moreover have "bounded B" unfolding bounded_iff
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2250
    proof (intro ballI exI)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2251
      fix t assume t: "t \<in> B"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2252
      have "norm t \<le> \<bar>Re t\<bar> + \<bar>Im t\<bar>" by (rule cmod_le)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2253
      also from t have "\<bar>Re t\<bar> \<le> m" unfolding B_def by blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2254
      also from t have "\<bar>Im t\<bar> \<le> \<bar>Im z\<bar>" unfolding B_def by blast
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2255
      finally show "norm t \<le> m + \<bar>Im z\<bar>" by - simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2256
    qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2257
    ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2258
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  2259
    define M where "M = (SUP z:B. norm (h' z))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2260
    have "compact (h' ` B)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2261
      by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2262
    hence bdd: "bdd_above ((\<lambda>z. norm (h' z)) ` B)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2263
      using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2264
    have "norm (h' z) \<le> M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2265
    also have "M \<le> M/2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2266
    proof (subst M_def, subst cSUP_le_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2267
      have "z \<in> B" unfolding B_def m_def by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2268
      thus "B \<noteq> {}" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2269
    next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2270
      show "\<forall>z\<in>B. norm (h' z) \<le> M/2"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2271
      proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2272
        fix t :: complex assume t: "t \<in> B"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2273
        from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2274
        also have "norm \<dots> = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2275
        also have "norm (h' (t/2) + h' ((t+1)/2)) \<le> norm (h' (t/2)) + norm (h' ((t+1)/2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2276
          by (rule norm_triangle_ineq)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2277
        also from t have "abs (Re ((t + 1)/2)) \<le> m" unfolding m_def B_def by auto
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2278
        with t have "t/2 \<in> B" "(t+1)/2 \<in> B" unfolding B_def by auto
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2279
        hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \<le> M + M" unfolding M_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2280
          by (intro add_mono cSUP_upper bdd) (auto simp: B_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2281
        also have "(M + M) / 4 = M / 2" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2282
        finally show "norm (h' t) \<le> M/2" by - simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2283
      qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2284
    qed (insert bdd, auto simp: cball_eq_empty)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2285
    hence "M \<le> 0" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2286
    finally show "h' z = 0" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2287
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2288
  have h_h'_2: "(h has_field_derivative 0) (at z)" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2289
    using h_h'[of z] h'_zero[of z] by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2290
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2291
  have g_real: "g z \<in> \<real>" if "z \<in> \<real>" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2292
    unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2293
  have h_real: "h z \<in> \<real>" if "z \<in> \<real>" for z
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2294
    unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2295
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2296
  from h'_zero h_h'_2 have "\<exists>c. \<forall>z\<in>UNIV. h z = c"
66936
cf8d8fc23891 tuned some proofs and added some lemmas
haftmann
parents: 66526
diff changeset
  2297
    by (intro has_field_derivative_zero_constant) simp_all
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2298
  then obtain c where c: "\<And>z. h z = c" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2299
  have "\<exists>u. u \<in> closed_segment 0 1 \<and> Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2300
    by (intro complex_mvt_line g_g')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2301
  then guess u by (elim exE conjE) note u = this
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2302
  from u(1) have u': "u \<in> \<real>" unfolding closed_segment_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2303
    by (auto simp: scaleR_conv_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2304
  from u' g_real[of u] g_nz[of u] have "Re (g u) \<noteq> 0" by (auto elim!: Reals_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2305
  with u(2) c[of u] g_real[of u] g_nz[of u] u'
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2306
    have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2307
  with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2308
  with c have A: "h z * g z = 0" for z by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2309
  hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2310
  hence "\<exists>c'. \<forall>z\<in>UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2311
  then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm)
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63417
diff changeset
  2312
  from this[of 0] have "c' = pi" unfolding g_def by simp
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63417
diff changeset
  2313
  with c have "g z = pi" by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2314
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2315
  show ?thesis
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2316
  proof (cases "z \<in> \<int>")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2317
    case False
62072
bf3d9f113474 isabelle update_cartouches -c -t;
wenzelm
parents: 62055
diff changeset
  2318
    with \<open>g z = pi\<close> show ?thesis by (auto simp: g_def divide_simps)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2319
  next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2320
    case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2321
    then obtain n where n: "z = of_int n" by (elim Ints_cases)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2322
    with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2323
    moreover have "of_int (1 - n) \<in> \<int>\<^sub>\<le>\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2324
    ultimately show ?thesis using n
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2325
      by (cases "n \<le> 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2326
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2327
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2328
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2329
lemma rGamma_reflection_complex:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2330
  "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2331
  using Gamma_reflection_complex[of z]
62390
842917225d56 more canonical names
nipkow
parents: 62131
diff changeset
  2332
    by (simp add: Gamma_def divide_simps split: if_split_asm)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2333
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2334
lemma rGamma_reflection_complex':
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2335
  "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2336
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2337
  have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2338
    using rGamma_plus1[of "-z", symmetric] by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2339
  also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2340
    by (rule rGamma_reflection_complex)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2341
  finally show ?thesis by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2342
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2343
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2344
lemma Gamma_reflection_complex':
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2345
  "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2346
  using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def divide_simps mult_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2347
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2348
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2349
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2350
lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2351
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2352
  from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2353
    have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2354
  hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2355
  also have "\<dots> = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2356
  finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2357
  moreover have "Gamma (1/2 :: real) \<ge> 0" using Gamma_real_pos[of "1/2"] by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2358
  ultimately show ?thesis by (rule real_sqrt_unique [symmetric])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2359
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2360
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2361
lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2362
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2363
  have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2364
  also have "\<dots> = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2365
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2366
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2367
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2368
lemma Gamma_legendre_duplication:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2369
  fixes z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2370
  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2371
  shows "Gamma z * Gamma (z + 1/2) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2372
             exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2373
  using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2374
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2375
end
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2376
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2377
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2378
subsection \<open>Limits and residues\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2379
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2380
text \<open>
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2381
  The inverse of the Gamma function has simple zeros:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2382
\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2383
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2384
lemma rGamma_zeros:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2385
  "(\<lambda>z. rGamma z / (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n * fact n :: 'a :: Gamma)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2386
proof (subst tendsto_cong)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2387
  let ?f = "\<lambda>z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2388
  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2389
    show "eventually (\<lambda>z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2390
    by (subst pochhammer_rGamma[of _ "Suc n"])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2391
       (auto elim!: eventually_mono simp: divide_simps pochhammer_rec' eq_neg_iff_add_eq_0)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2392
  have "isCont ?f (- of_nat n)" by (intro continuous_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2393
  thus "?f \<midarrow> (- of_nat n) \<rightarrow> (- 1) ^ n * fact n" unfolding isCont_def
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2394
    by (simp add: pochhammer_same)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2395
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2396
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2397
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2398
text \<open>
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2399
  The simple zeros of the inverse of the Gamma function correspond to simple poles of the Gamma function,
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2400
  and their residues can easily be computed from the limit we have just proven:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2401
\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2402
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2403
lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2404
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2405
  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2406
    have "eventually (\<lambda>z. rGamma z \<noteq> (0 :: 'a)) (at (- of_nat n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2407
    by (auto elim!: eventually_mono nonpos_Ints_cases'
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2408
             simp: rGamma_eq_zero_iff dist_of_nat dist_minus)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2409
  with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident]
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2410
    have "filterlim (\<lambda>z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2411
    unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2412
                            (simp_all add: filterlim_at)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2413
  moreover have "(\<lambda>z. inverse (rGamma z) :: 'a) = Gamma"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2414
    by (intro ext) (simp add: rGamma_inverse_Gamma)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2415
  ultimately show ?thesis by (simp only: )
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2416
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2417
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2418
lemma Gamma_residues:
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2419
  "(\<lambda>z. Gamma z * (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n / fact n :: 'a :: Gamma)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2420
proof (subst tendsto_cong)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2421
  let ?c = "(- 1) ^ n / fact n :: 'a"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2422
  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2423
    show "eventually (\<lambda>z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n)))
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2424
            (at (- of_nat n))"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2425
    by (auto elim!: eventually_mono simp: divide_simps rGamma_inverse_Gamma)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2426
  have "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow>
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2427
          inverse ((- 1) ^ n * fact n :: 'a)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2428
    by (intro tendsto_intros rGamma_zeros) simp_all
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2429
  also have "inverse ((- 1) ^ n * fact n) = ?c"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2430
    by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2431
  finally show "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> ?c" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2432
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2433
66286
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2434
lemma is_pole_Gamma: "is_pole Gamma (- of_nat n)"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2435
  unfolding is_pole_def using Gamma_poles .
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2436
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2437
lemma Gamme_residue:
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2438
  "residue Gamma (- of_nat n) = (-1) ^ n / fact n"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2439
proof (rule residue_simple')
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2440
  show "open (- (\<int>\<^sub>\<le>\<^sub>0 - {-of_nat n}) :: complex set)"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2441
    by (intro open_Compl closed_subset_Ints) auto
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2442
  show "Gamma holomorphic_on (- (\<int>\<^sub>\<le>\<^sub>0 - {-of_nat n}) - {- of_nat n})"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2443
    by (rule holomorphic_Gamma) auto
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2444
  show "(\<lambda>w. Gamma w * (w - - of_nat n)) \<midarrow>- of_nat n \<rightarrow> (- 1) ^ n / fact n"
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2445
    using Gamma_residues[of n] by simp
1c977b13414f poles and residues of the Gamma function
eberlm <eberlm@in.tum.de>
parents: 65587
diff changeset
  2446
qed auto
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2447
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2448
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2449
subsection \<open>Alternative definitions\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2450
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2451
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2452
subsubsection \<open>Variant of the Euler form\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2453
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2454
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2455
definition Gamma_series_euler' where
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2456
  "Gamma_series_euler' z n =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2457
     inverse z * (\<Prod>k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2458
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2459
context
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2460
begin
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2461
private lemma Gamma_euler'_aux1:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2462
  fixes z :: "'a :: {real_normed_field,banach}"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2463
  assumes n: "n > 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2464
  shows "exp (z * of_real (ln (of_nat n + 1))) = (\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2465
proof -
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2466
  have "(\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2467
          exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
  2468
    by (subst exp_sum [symmetric]) (simp_all add: sum_distrib_left)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2469
  also have "(\<Sum>k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\<Prod>k=1..n. 1 + 1 / real_of_nat k)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2470
    by (subst ln_prod [symmetric]) (auto intro!: add_pos_nonneg)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2471
  also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2472
    by (intro prod.cong) (simp_all add: divide_simps)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2473
  also have "(\<Prod>k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2474
    by (induction n) (simp_all add: prod_nat_ivl_Suc' divide_simps)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2475
  finally show ?thesis ..
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2476
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2477
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2478
lemma Gamma_series_euler':
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2479
  assumes z: "(z :: 'a :: Gamma) \<notin> \<int>\<^sub>\<le>\<^sub>0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2480
  shows "(\<lambda>n. Gamma_series_euler' z n) \<longlonglongrightarrow> Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2481
proof (rule Gamma_seriesI, rule Lim_transform_eventually)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2482
  let ?f = "\<lambda>n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2483
  let ?r = "\<lambda>n. ?f n / Gamma_series z n"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2484
  let ?r' = "\<lambda>n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2485
  from z have z': "z \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2486
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  2487
  have "eventually (\<lambda>n. ?r' n = ?r n) sequentially"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2488
    using z by (auto simp: divide_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac
65578
e4997c181cce New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
paulson <lp15@cam.ac.uk>
parents: 64969
diff changeset
  2489
                     intro: eventually_mono eventually_gt_at_top[of "0::nat"] dest: pochhammer_eq_0_imp_nonpos_Int)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2490
  moreover have "?r' \<longlonglongrightarrow> exp (z * of_real (ln 1))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2491
    by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2492
  ultimately show "?r \<longlonglongrightarrow> 1" by (force dest!: Lim_transform_eventually)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2493
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2494
  from eventually_gt_at_top[of "0::nat"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2495
    show "eventually (\<lambda>n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2496
  proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2497
    fix n :: nat assume n: "n > 0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2498
    from n z' have "Gamma_series_euler' z n =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2499
      exp (z * of_real (ln (of_nat n + 1))) / (z * (\<Prod>k=1..n. (1 + z / of_nat k)))"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2500
      by (subst Gamma_euler'_aux1)
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2501
         (simp_all add: Gamma_series_euler'_def prod.distrib
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2502
                        prod_inversef[symmetric] divide_inverse)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2503
    also have "(\<Prod>k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2504
      by (cases n) (simp_all add: pochhammer_prod fact_prod atLeastLessThanSuc_atLeastAtMost
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2505
        prod_dividef [symmetric] field_simps prod.atLeast_Suc_atMost_Suc_shift)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2506
    also have "z * \<dots> = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2507
    finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2508
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2509
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2510
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2511
end
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2512
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2513
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2514
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2515
subsubsection \<open>Weierstrass form\<close>
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2516
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2517
definition Gamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2518
  "Gamma_series_weierstrass z n =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2519
     exp (-euler_mascheroni * z) / z * (\<Prod>k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2520
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2521
definition rGamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2522
  "rGamma_series_weierstrass z n =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2523
     exp (euler_mascheroni * z) * z * (\<Prod>k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2524
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2525
lemma Gamma_series_weierstrass_nonpos_Ints:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2526
  "eventually (\<lambda>k. Gamma_series_weierstrass (- of_nat n) k = 0) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2527
  using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_weierstrass_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2528
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2529
lemma rGamma_series_weierstrass_nonpos_Ints:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2530
  "eventually (\<lambda>k. rGamma_series_weierstrass (- of_nat n) k = 0) sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2531
  using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_weierstrass_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2532
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2533
lemma Gamma_weierstrass_complex: "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma (z :: complex)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2534
proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2535
  case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2536
  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2537
  also from True have "Gamma_series_weierstrass \<dots> \<longlonglongrightarrow> Gamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2538
    by (simp add: tendsto_cong[OF Gamma_series_weierstrass_nonpos_Ints] Gamma_nonpos_Int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2539
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2540
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2541
  case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2542
  hence z: "z \<noteq> 0" by auto
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2543
  let ?f = "(\<lambda>x. \<Prod>x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2544
  have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \<ge> 1" for n :: nat
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2545
    using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2546
  have "(\<lambda>n. \<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k)) \<longlonglongrightarrow> ln_Gamma z + euler_mascheroni * z + ln z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2547
    using ln_Gamma_series'_aux[OF False]
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2548
    by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
  2549
                   sum_shift_bounds_Suc_ivl sums_def atLeast0LessThan)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2550
  from tendsto_exp[OF this] False z have "?f \<longlonglongrightarrow> z * exp (euler_mascheroni * z) * Gamma z"
64267
b9a1486e79be setsum -> sum
nipkow
parents: 63992
diff changeset
  2551
    by (simp add: exp_add exp_sum exp_diff mult_ac Gamma_complex_altdef A)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2552
  from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2553
    show "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma z"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2554
    by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2555
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2556
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2557
lemma tendsto_complex_of_real_iff: "((\<lambda>x. complex_of_real (f x)) \<longlongrightarrow> of_real c) F = (f \<longlongrightarrow> c) F"
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2558
  by (rule tendsto_of_real_iff)
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2559
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2560
lemma Gamma_weierstrass_real: "Gamma_series_weierstrass x \<longlonglongrightarrow> Gamma (x :: real)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2561
  using Gamma_weierstrass_complex[of "of_real x"] unfolding Gamma_series_weierstrass_def[abs_def]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2562
  by (subst tendsto_complex_of_real_iff [symmetric])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2563
     (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2564
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2565
lemma rGamma_weierstrass_complex: "rGamma_series_weierstrass z \<longlonglongrightarrow> rGamma (z :: complex)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2566
proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2567
  case True
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2568
  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2569
  also from True have "rGamma_series_weierstrass \<dots> \<longlonglongrightarrow> rGamma z"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2570
    by (simp add: tendsto_cong[OF rGamma_series_weierstrass_nonpos_Ints] rGamma_nonpos_Int)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2571
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2572
next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2573
  case False
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2574
  have "rGamma_series_weierstrass z = (\<lambda>n. inverse (Gamma_series_weierstrass z n))"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2575
    by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2576
                  exp_minus divide_inverse prod_inversef[symmetric] mult_ac)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2577
  also from False have "\<dots> \<longlonglongrightarrow> inverse (Gamma z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2578
    by (intro tendsto_intros Gamma_weierstrass_complex) (simp add: Gamma_eq_zero_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2579
  finally show ?thesis by (simp add: Gamma_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2580
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2581
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2582
subsubsection \<open>Binomial coefficient form\<close>
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2583
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2584
lemma Gamma_gbinomial:
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2585
  "(\<lambda>n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \<longlonglongrightarrow> rGamma (z+1)"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2586
proof (cases "z = 0")
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2587
  case False
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2588
  show ?thesis
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2589
  proof (rule Lim_transform_eventually)
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2590
    let ?powr = "\<lambda>a b. exp (b * of_real (ln (of_nat a)))"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2591
    show "eventually (\<lambda>n. rGamma_series z n / z =
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2592
            ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2593
    proof (intro always_eventually allI)
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2594
      fix n :: nat
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2595
      from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2596
        by (simp add: gbinomial_pochhammer' pochhammer_rec)
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2597
      also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2598
        by (simp add: rGamma_series_def divide_simps exp_minus)
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2599
      finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" ..
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2600
    qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2601
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2602
    from False have "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma z / z" by (intro tendsto_intros)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2603
    also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z]
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2604
      by (simp add: field_simps)
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2605
    finally show "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma (z+1)" .
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2606
  qed
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2607
qed (simp_all add: binomial_gbinomial [symmetric])
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2608
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2609
lemma gbinomial_minus': "(a + of_nat b) gchoose b = (- 1) ^ b * (- (a + 1) gchoose b)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2610
  by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric])
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2611
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2612
lemma gbinomial_asymptotic:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2613
  fixes z :: "'a :: Gamma"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  2614
  shows "(\<lambda>n. (z gchoose n) / ((-1)^n / exp ((z+1) * of_real (ln (real n))))) \<longlonglongrightarrow>
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2615
           inverse (Gamma (- z))"
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  2616
  unfolding rGamma_inverse_Gamma [symmetric] using Gamma_gbinomial[of "-z-1"]
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2617
  by (subst (asm) gbinomial_minus')
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2618
     (simp add: add_ac mult_ac divide_inverse power_inverse [symmetric])
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2619
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  2620
lemma fact_binomial_limit:
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2621
  "(\<lambda>n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \<longlonglongrightarrow> 1 / fact k"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2622
proof (rule Lim_transform_eventually)
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2623
  have "(\<lambda>n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n))))
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2624
            \<longlonglongrightarrow> 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \<longlonglongrightarrow> _")
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2625
    using Gamma_gbinomial[of "of_nat k :: 'a"]
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2626
    by (simp add: binomial_gbinomial add_ac Gamma_def divide_simps exp_of_real [symmetric] exp_minus)
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2627
  also have "Gamma (of_nat (Suc k)) = fact k" by (simp add: Gamma_fact)
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2628
  finally show "?f \<longlonglongrightarrow> 1 / fact k" .
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2629
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2630
  show "eventually (\<lambda>n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2631
    using eventually_gt_at_top[of "0::nat"]
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2632
  proof eventually_elim
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2633
    fix n :: nat assume n: "n > 0"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2634
    from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2635
      by (simp add: exp_of_nat_mult)
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2636
    thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2637
  qed
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2638
qed
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2639
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2640
lemma binomial_asymptotic':
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2641
  "(\<lambda>n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \<longlonglongrightarrow> 1"
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2642
  using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2643
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2644
lemma gbinomial_Beta:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2645
  assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2646
  shows   "((z::'a::Gamma) gchoose n) = inverse ((z + 1) * Beta (z - of_nat n + 1) (of_nat n + 1))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2647
using assms
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2648
proof (induction n arbitrary: z)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2649
  case 0
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2650
  hence "z + 2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2651
    using plus_one_in_nonpos_Ints_imp[of "z+1"] by (auto simp: add.commute)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2652
  with 0 show ?case
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2653
    by (auto simp: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric] add.commute)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2654
next
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2655
  case (Suc n z)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2656
  show ?case
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2657
  proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2658
    case True
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2659
    with Suc.prems have "z = 0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2660
      by (auto elim!: nonpos_Ints_cases simp: algebra_simps one_plus_of_int_in_nonpos_Ints_iff)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2661
    show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2662
    proof (cases "n = 0")
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2663
      case True
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2664
      with \<open>z = 0\<close> show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2665
        by (simp add: Beta_def Gamma_eq_zero_iff Gamma_plus1 [symmetric])
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2666
    next
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2667
      case False
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2668
      with \<open>z = 0\<close> show ?thesis
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2669
        by (simp_all add: Beta_pole1 one_minus_of_nat_in_nonpos_Ints_iff gbinomial_1)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2670
    qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2671
  next
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2672
    case False
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2673
    have "(z gchoose (Suc n)) = ((z - 1 + 1) gchoose (Suc n))" by simp
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2674
    also have "\<dots> = (z - 1 gchoose n) * ((z - 1) + 1) / of_nat (Suc n)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2675
      by (subst gbinomial_factors) (simp add: field_simps)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  2676
    also from False have "\<dots> = inverse (of_nat (Suc n) * Beta (z - of_nat n) (of_nat (Suc n)))"
63295
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2677
      (is "_ = inverse ?x") by (subst Suc.IH) (simp_all add: field_simps Beta_pole1)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2678
    also have "of_nat (Suc n) \<notin> (\<int>\<^sub>\<le>\<^sub>0 :: 'a set)" by (subst of_nat_in_nonpos_Ints_iff) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2679
    hence "?x = (z + 1) * Beta (z - of_nat (Suc n) + 1) (of_nat (Suc n) + 1)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2680
      by (subst Beta_plus1_right [symmetric]) simp_all
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2681
    finally show ?thesis .
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2682
  qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2683
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2684
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2685
lemma gbinomial_Gamma:
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2686
  assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2687
  shows   "(z gchoose n) = Gamma (z + 1) / (fact n * Gamma (z - of_nat n + 1))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2688
proof -
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2689
  have "(z gchoose n) = Gamma (z + 2) / (z + 1) / (fact n * Gamma (z - of_nat n + 1))"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2690
    by (subst gbinomial_Beta[OF assms]) (simp_all add: Beta_def Gamma_fact [symmetric] add_ac)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2691
  also from assms have "Gamma (z + 2) / (z + 1) = Gamma (z + 1)"
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2692
    using Gamma_plus1[of "z+1"] by (auto simp add: divide_simps mult_ac add_ac)
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2693
  finally show ?thesis .
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2694
qed
52792bb9126e Facts about HK integration, complex powers, Gamma function
eberlm
parents: 63040
diff changeset
  2695
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2696
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2697
subsubsection \<open>Integral form\<close>
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2698
66526
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2699
lemma integrable_on_powr_from_0':
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2700
  assumes a: "a > (-1::real)" and c: "c \<ge> 0"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2701
  shows   "(\<lambda>x. x powr a) integrable_on {0<..c}"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2702
proof -
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2703
  from c have *: "{0<..c} - {0..c} = {}" "{0..c} - {0<..c} = {0}" by auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2704
  show ?thesis
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2705
  by (rule integrable_spike_set [OF integrable_on_powr_from_0[OF a c]]) (simp_all add: *)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2706
qed
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2707
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2708
lemma absolutely_integrable_Gamma_integral:
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2709
  assumes "Re z > 0" "a > 0"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2710
  shows   "(\<lambda>t. complex_of_real t powr (z - 1) / of_real (exp (a * t))) 
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2711
             absolutely_integrable_on {0<..}" (is "?f absolutely_integrable_on _")
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2712
proof -
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2713
  have "((\<lambda>x. (Re z - 1) * (ln x / x)) \<longlongrightarrow> (Re z - 1) * 0) at_top"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2714
    by (intro tendsto_intros ln_x_over_x_tendsto_0)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2715
  hence "((\<lambda>x. ((Re z - 1) * ln x) / x) \<longlongrightarrow> 0) at_top" by simp
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2716
  from order_tendstoD(2)[OF this, of "a/2"] and \<open>a > 0\<close>
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2717
    have "eventually (\<lambda>x. (Re z - 1) * ln x / x < a/2) at_top" by simp
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2718
  from eventually_conj[OF this eventually_gt_at_top[of 0]]
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2719
    obtain x0 where "\<forall>x\<ge>x0. (Re z - 1) * ln x / x < a/2 \<and> x > 0"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2720
      by (auto simp: eventually_at_top_linorder)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2721
  hence "x0 > 0" by simp
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2722
  have "x powr (Re z - 1) / exp (a * x) < exp (-(a/2) * x)" if "x \<ge> x0" for x
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2723
  proof -
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2724
    from that and \<open>\<forall>x\<ge>x0. _\<close> have x: "(Re z - 1) * ln x / x < a / 2" "x > 0" by auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2725
    have "x powr (Re z - 1) = exp ((Re z - 1) * ln x)"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2726
      using \<open>x > 0\<close> by (simp add: powr_def)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2727
    also from x have "(Re z - 1) * ln x < (a * x) / 2" by (simp add: field_simps)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2728
    finally show ?thesis by (simp add: field_simps exp_add [symmetric])
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2729
  qed
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2730
  note x0 = \<open>x0 > 0\<close> this
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2731
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2732
  have "?f absolutely_integrable_on ({0<..x0} \<union> {x0..})"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2733
  proof (rule set_integrable_Un)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2734
    show "?f absolutely_integrable_on {0<..x0}"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2735
    proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2])
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2736
      show "set_integrable lebesgue {0<..x0} (\<lambda>x. x powr (Re z - 1))" using x0(1) assms
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2737
        by (intro nonnegative_absolutely_integrable_1 integrable_on_powr_from_0') auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2738
      show "set_borel_measurable lebesgue {0<..x0}
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2739
              (\<lambda>x. complex_of_real x powr (z - 1) / complex_of_real (exp (a * x)))"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2740
        by (intro measurable_completion)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2741
           (auto intro!: borel_measurable_continuous_on_indicator continuous_intros)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2742
      fix x :: real 
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2743
      have "x powr (Re z - 1) / exp (a * x) \<le> x powr (Re z - 1) / 1" if "x \<ge> 0"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2744
        using that assms by (intro divide_left_mono) auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2745
      thus "norm (indicator {0<..x0} x *\<^sub>R ?f x) \<le> 
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2746
               norm (indicator {0<..x0} x *\<^sub>R x powr (Re z - 1))"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2747
        by (simp_all add: norm_divide norm_powr_real_powr indicator_def)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2748
    qed
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2749
  next
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2750
    show "?f absolutely_integrable_on {x0..}"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2751
    proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2])
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2752
      show "set_integrable lebesgue {x0..} (\<lambda>x. exp (-(a/2) * x))" using assms
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2753
        by (intro nonnegative_absolutely_integrable_1 integrable_on_exp_minus_to_infinity) auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2754
      show "set_borel_measurable lebesgue {x0..}
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2755
              (\<lambda>x. complex_of_real x powr (z - 1) / complex_of_real (exp (a * x)))" using x0(1)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2756
        by (intro measurable_completion)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2757
           (auto intro!: borel_measurable_continuous_on_indicator continuous_intros)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2758
      fix x :: real 
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2759
      show "norm (indicator {x0..} x *\<^sub>R ?f x) \<le> 
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2760
               norm (indicator {x0..} x *\<^sub>R exp (-(a/2) * x))" using x0
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2761
        by (auto simp: norm_divide norm_powr_real_powr indicator_def less_imp_le)
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2762
    qed
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2763
  qed auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2764
  also have "{0<..x0} \<union> {x0..} = {0<..}" using x0(1) by auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2765
  finally show ?thesis .
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2766
qed
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2767
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2768
lemma integrable_Gamma_integral_bound:
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2769
  fixes a c :: real
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2770
  assumes a: "a > -1" and c: "c \<ge> 0"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2771
  defines "f \<equiv> \<lambda>x. if x \<in> {0..c} then x powr a else exp (-x/2)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2772
  shows   "f integrable_on {0..}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2773
proof -
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2774
  have "f integrable_on {0..c}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2775
    by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]])
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2776
       (insert a c, simp_all add: f_def)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2777
  moreover have A: "(\<lambda>x. exp (-x/2)) integrable_on {c..}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2778
    using integrable_on_exp_minus_to_infinity[of "1/2"] by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2779
  have "f integrable_on {c..}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2780
    by (rule integrable_spike_finite[of "{c}", OF _ _ A]) (simp_all add: f_def)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  2781
  ultimately show "f integrable_on {0..}"
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  2782
    by (rule integrable_Un') (insert c, auto simp: max_def)
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  2783
qed
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2784
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2785
lemma Gamma_integral_complex:
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2786
  assumes z: "Re z > 0"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2787
  shows   "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2788
proof -
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2789
  have A: "((\<lambda>t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n))
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2790
          has_integral (fact n / pochhammer z (n+1))) {0..1}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2791
    if "Re z > 0" for n z using that
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2792
  proof (induction n arbitrary: z)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2793
    case 0
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2794
    have "((\<lambda>t. complex_of_real t powr (z - 1)) has_integral
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2795
            (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2796
      by (intro fundamental_theorem_of_calculus_interior)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2797
         (auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_complex)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2798
    thus ?case by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2799
  next
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2800
    case (Suc n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2801
    let ?f = "\<lambda>t. complex_of_real t powr z / z"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2802
    let ?f' = "\<lambda>t. complex_of_real t powr (z - 1)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2803
    let ?g = "\<lambda>t. (1 - complex_of_real t) ^ Suc n"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2804
    let ?g' = "\<lambda>t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2805
    have "((\<lambda>t. ?f' t * ?g t) has_integral
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2806
            (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2807
      (is "(_ has_integral ?I) _")
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2808
    proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g])
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2809
      from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2810
        by (auto intro!: continuous_intros)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2811
    next
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2812
      fix t :: real assume t: "t \<in> {0<..<1}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2813
      show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2814
        by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2815
      show "(?g has_vector_derivative ?g' t) (at t)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2816
        by (rule has_vector_derivative_real_complex derivative_eq_intros refl)+ simp_all
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2817
    next
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2818
      from Suc.prems have [simp]: "z \<noteq> 0" by auto
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2819
      from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2820
      have [simp]: "z + of_nat n \<noteq> 0" "z + 1 + of_nat n \<noteq> 0" for n
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2821
        using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2822
      have "((\<lambda>x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2823
              fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2824
        (is "(?A has_integral ?B) _")
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2825
        using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2826
      also have "?A = (\<lambda>t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2827
      also have "?B = - (of_nat (Suc n) * fact n / pochhammer z (n+2))"
63417
c184ec919c70 more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
haftmann
parents: 63367
diff changeset
  2828
        by (simp add: divide_simps pochhammer_rec
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  2829
              prod_shift_bounds_cl_Suc_ivl del: of_nat_Suc)
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2830
      finally show "((\<lambda>t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2831
        by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2832
    qed (simp_all add: bounded_bilinear_mult)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2833
    thus ?case by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2834
  qed
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2835
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2836
  have B: "((\<lambda>t. if t \<in> {0..of_nat n} then
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2837
             of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2838
           has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2839
  proof (cases "n > 0")
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2840
    case [simp]: True
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2841
    hence [simp]: "n \<noteq> 0" by auto
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2842
    with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0]
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2843
      have "((\<lambda>x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2844
              has_integral fact n * of_nat n / pochhammer z (n+1)) ((\<lambda>x. real n * x)`{0..1})"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2845
      (is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2846
    also from True have "((\<lambda>x. real n*x)`{0..1}) = {0..real n}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2847
      by (subst image_mult_atLeastAtMost) simp_all
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2848
    also have "?f = (\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2849
      using True by (intro ext) (simp add: field_simps)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2850
    finally have "((\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2851
                    has_integral ?I) {0..real n}" (is ?P) .
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2852
    also have "?P \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2853
                        has_integral ?I) {0..real n}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2854
      by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric])
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2855
    also have "\<dots> \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln x - ln (of_nat n))) * (1 - of_real x / of_nat n) ^ n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2856
                        has_integral ?I) {0..real n}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2857
      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2858
    finally have \<dots> .
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2859
    note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"]
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2860
    have "((\<lambda>x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2861
            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2862
      by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric])
66936
cf8d8fc23891 tuned some proofs and added some lemmas
haftmann
parents: 66526
diff changeset
  2863
         (simp add: algebra_simps)
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2864
    also have "?P \<longleftrightarrow> ((\<lambda>x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2865
            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2866
      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2867
    also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) =
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2868
                 (of_nat n powr z * fact n / pochhammer z (n+1))"
65587
16a8991ab398 New material (and some tidying) purely in the Analysis directory
paulson <lp15@cam.ac.uk>
parents: 65578
diff changeset
  2869
      by (auto simp add: powr_def algebra_simps exp_diff exp_of_real)
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2870
    finally show ?thesis by (subst has_integral_restrict) simp_all
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2871
  next
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2872
    case False
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2873
    thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2874
  qed
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2875
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2876
  have "eventually (\<lambda>n. Gamma_series z n =
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2877
          of_nat n powr z * fact n / pochhammer z (n+1)) sequentially"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2878
    using eventually_gt_at_top[of "0::nat"]
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2879
    by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2880
  from this and Gamma_series_LIMSEQ[of z]
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2881
    have C: "(\<lambda>k. of_nat k powr z * fact k / pochhammer z (k+1)) \<longlonglongrightarrow> Gamma z"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2882
    by (rule Lim_transform_eventually)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2883
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2884
  {
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2885
    fix x :: real assume x: "x \<ge> 0"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2886
    have lim_exp: "(\<lambda>k. (1 - x / real k) ^ k) \<longlonglongrightarrow> exp (-x)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2887
      using tendsto_exp_limit_sequentially[of "-x"] by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2888
    have "(\<lambda>k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k))
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2889
            \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))" (is ?P)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2890
      by (intro tendsto_intros lim_exp)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2891
    also from eventually_gt_at_top[of "nat \<lceil>x\<rceil>"]
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2892
      have "eventually (\<lambda>k. of_nat k > x) sequentially" by eventually_elim linarith
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2893
    hence "?P \<longleftrightarrow> (\<lambda>k. if x \<le> of_nat k then
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2894
                 of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2895
                   \<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2896
      by (intro tendsto_cong) (auto elim!: eventually_mono)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2897
    finally have \<dots> .
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2898
  }
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2899
  hence D: "\<forall>x\<in>{0..}. (\<lambda>k. if x \<in> {0..real k} then
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2900
              of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2901
             \<longlonglongrightarrow> of_real x powr (z - 1) / of_real (exp x)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2902
    by (simp add: exp_minus field_simps cong: if_cong)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2903
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2904
  have "((\<lambda>x. (Re z - 1) * (ln x / x)) \<longlongrightarrow> (Re z - 1) * 0) at_top"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2905
    by (intro tendsto_intros ln_x_over_x_tendsto_0)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2906
  hence "((\<lambda>x. ((Re z - 1) * ln x) / x) \<longlongrightarrow> 0) at_top" by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2907
  from order_tendstoD(2)[OF this, of "1/2"]
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2908
    have "eventually (\<lambda>x. (Re z - 1) * ln x / x < 1/2) at_top" by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2909
  from eventually_conj[OF this eventually_gt_at_top[of 0]]
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2910
    obtain x0 where "\<forall>x\<ge>x0. (Re z - 1) * ln x / x < 1/2 \<and> x > 0"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2911
    by (auto simp: eventually_at_top_linorder)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2912
  hence x0: "x0 > 0" "\<And>x. x \<ge> x0 \<Longrightarrow> (Re z - 1) * ln x < x / 2" by auto
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2913
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2914
  define h where "h = (\<lambda>x. if x \<in> {0..x0} then x powr (Re z - 1) else exp (-x/2))"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2915
  have le_h: "x powr (Re z - 1) * exp (-x) \<le> h x" if x: "x \<ge> 0" for x
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2916
  proof (cases "x > x0")
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2917
    case True
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2918
    from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2919
      by (simp add: powr_def exp_diff exp_minus field_simps exp_add)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2920
    also from x0(2)[of x] True have "\<dots> < exp (-x/2)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2921
      by (simp add: field_simps)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2922
    finally show ?thesis using True by (auto simp add: h_def)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2923
  next
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2924
    case False
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2925
    from x have "x powr (Re z - 1) * exp (- x) \<le> x powr (Re z - 1) * 1"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2926
      by (intro mult_left_mono) simp_all
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2927
    with False show ?thesis by (auto simp add: h_def)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2928
  qed
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2929
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2930
  have E: "\<forall>x\<in>{0..}. cmod (if x \<in> {0..real k} then of_real x powr (z - 1) *
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2931
                   (1 - complex_of_real x / of_nat k) ^ k else 0) \<le> h x"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2932
    (is "\<forall>x\<in>_. ?f x \<le> _") for k
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2933
  proof safe
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2934
    fix x :: real assume x: "x \<ge> 0"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2935
    {
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2936
      fix x :: real and n :: nat assume x: "x \<le> of_nat n"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2937
      have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2938
      also have "norm \<dots> = \<bar>(1 - x / real n)\<bar>" by (subst norm_of_real) (rule refl)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2939
      also from x have "\<dots> = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: divide_simps)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2940
      finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" .
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2941
    } note D = this
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2942
    from D[of x k] x
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2943
      have "?f x \<le> (if of_nat k \<ge> x \<and> k > 0 then x powr (Re z - 1) * (1 - x / real k) ^ k else 0)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2944
      by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2945
    also have "\<dots> \<le> x powr (Re z - 1) * exp  (-x)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2946
      by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2947
    also from x have "\<dots> \<le> h x" by (rule le_h)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2948
    finally show "?f x \<le> h x" .
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2949
  qed
63594
bd218a9320b5 HOL-Multivariate_Analysis: rename theories for more descriptive names
hoelzl
parents: 63539
diff changeset
  2950
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2951
  have F: "h integrable_on {0..}" unfolding h_def
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2952
    by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2953
  show ?thesis
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2954
    by (rule has_integral_dominated_convergence[OF B F E D C])
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2955
qed
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2956
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2957
lemma Gamma_integral_real:
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2958
  assumes x: "x > (0 :: real)"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2959
  shows   "((\<lambda>t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2960
proof -
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2961
  have A: "((\<lambda>t. complex_of_real t powr (complex_of_real x - 1) /
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2962
          complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2963
    using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2964
  have "((\<lambda>t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}"
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2965
    by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric])
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2966
  from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def)
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2967
qed
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2968
66526
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2969
lemma absolutely_integrable_Gamma_integral':
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2970
  assumes "Re z > 0"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2971
  shows   "(\<lambda>t. complex_of_real t powr (z - 1) / of_real (exp t)) absolutely_integrable_on {0<..}"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2972
  using absolutely_integrable_Gamma_integral [OF assms zero_less_one] by simp
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2973
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2974
lemma Gamma_integral_complex':
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2975
  assumes z: "Re z > 0"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2976
  shows   "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0<..}"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2977
proof -
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2978
  have "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2979
    by (rule Gamma_integral_complex) fact+
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2980
  hence "((\<lambda>t. if t \<in> {0<..} then of_real t powr (z - 1) / of_real (exp t) else 0) 
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2981
           has_integral Gamma z) {0..}"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2982
    by (rule has_integral_spike [of "{0}", rotated 2]) auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2983
  also have "?this = ?thesis"
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2984
    by (subst has_integral_restrict) auto
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2985
  finally show ?thesis .
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2986
qed
322120e880c5 More material on infinite sums
eberlm <eberlm@in.tum.de>
parents: 66512
diff changeset
  2987
63296
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2988
3951a15a05d1 Integral form of Gamma function
eberlm
parents: 63295
diff changeset
  2989
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2990
subsection \<open>The Weierstraß product formula for the sine\<close>
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  2991
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2992
lemma sin_product_formula_complex:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2993
  fixes z :: complex
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2994
  shows "(\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k^2)) \<longlonglongrightarrow> sin (of_real pi * z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2995
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2996
  let ?f = "rGamma_series_weierstrass"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2997
  have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n))
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2998
            \<longlonglongrightarrow> (- of_real pi * inverse z) * (rGamma z * rGamma (- z))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  2999
    by (intro tendsto_intros rGamma_weierstrass_complex)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3000
  also have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3001
                    (\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3002
  proof
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3003
    fix n :: nat
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3004
    have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) =
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3005
              of_real pi * z * (\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3006
      by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3007
                    divide_simps prod.distrib[symmetric] power2_eq_square)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3008
    also have "(\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) =
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3009
                 (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3010
      by (intro prod.cong) (simp_all add: power2_eq_square field_simps)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3011
    finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \<dots>"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3012
      by (simp add: divide_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3013
  qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3014
  also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3015
    by (subst rGamma_reflection_complex') (simp add: divide_simps)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3016
  finally show ?thesis .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3017
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3018
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3019
lemma sin_product_formula_real:
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3020
  "(\<lambda>n. pi * (x::real) * (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3021
proof -
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3022
  from sin_product_formula_complex[of "of_real x"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3023
    have "(\<lambda>n. of_real pi * of_real x * (\<Prod>k=1..n. 1 - (of_real x)^2 / (of_nat k)^2))
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3024
              \<longlonglongrightarrow> sin (of_real pi * of_real x :: complex)" (is "?f \<longlonglongrightarrow> ?y") .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3025
  also have "?f = (\<lambda>n. of_real (pi * x * (\<Prod>k=1..n. 1 - x^2 / (of_nat k^2))))" by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3026
  also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3027
  finally show ?thesis by (subst (asm) tendsto_of_real_iff)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3028
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3029
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3030
lemma sin_product_formula_real':
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3031
  assumes "x \<noteq> (0::real)"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3032
  shows   "(\<lambda>n. (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x) / (pi * x)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3033
  using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3034
  by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3035
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3036
theorem wallis: "(\<lambda>n. \<Prod>k=1..n. (4*real k^2) / (4*real k^2 - 1)) \<longlonglongrightarrow> pi / 2"
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3037
proof -
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  3038
  from tendsto_inverse[OF tendsto_mult[OF
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3039
         sin_product_formula_real[of "1/2"] tendsto_const[of "2/pi"]]]
66512
89b6455b63b6 starting to unscramble bounded_variation_absolutely_integrable_interval
paulson <lp15@cam.ac.uk>
parents: 66453
diff changeset
  3040
    have "(\<lambda>n. (\<Prod>k=1..n. inverse (1 - (1 / 2)\<^sup>2 / (real k)\<^sup>2))) \<longlonglongrightarrow> pi/2"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3041
    by (simp add: prod_inversef [symmetric])
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3042
  also have "(\<lambda>n. (\<Prod>k=1..n. inverse (1 - (1 / 2)\<^sup>2 / (real k)\<^sup>2))) =
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3043
               (\<lambda>n. (\<Prod>k=1..n. (4*real k^2)/(4*real k^2 - 1)))"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3044
    by (intro ext prod.cong refl) (simp add: divide_simps)
63721
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3045
  finally show ?thesis .
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3046
qed
492bb53c3420 More analysis lemmas
Manuel Eberl <eberlm@in.tum.de>
parents: 63627
diff changeset
  3047
62085
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  3048
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  3049
subsection \<open>The Solution to the Basel problem\<close>
5b7758af429e Tuned approximations in Multivariate_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents: 62072
diff changeset
  3050
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3051
theorem inverse_squares_sums: "(\<lambda>n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3052
proof -
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  3053
  define P where "P x n = (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  3054
  define K where "K = (\<Sum>n. inverse (real_of_nat (Suc n))^2)"
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  3055
  define f where [abs_def]: "f x = (\<Sum>n. P x n / of_nat (Suc n)^2)" for x
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  3056
  define g where [abs_def]: "g x = (1 - sin (pi * x) / (pi * x))" for x
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3057
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3058
  have sums: "(\<lambda>n. P x n / of_nat (Suc n)^2) sums (if x = 0 then K else g x / x^2)" for x
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3059
  proof (cases "x = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3060
    assume x: "x = 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3061
    have "summable (\<lambda>n. inverse ((real_of_nat (Suc n))\<^sup>2))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3062
      using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3063
    thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3064
  next
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3065
    assume x: "x \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3066
    have "(\<lambda>n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3067
      unfolding P_def using x by (intro telescope_sums' sin_product_formula_real')
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3068
    also have "(\<lambda>n. P x n - P x (Suc n)) = (\<lambda>n. (x^2 / of_nat (Suc n)^2) * P x n)"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3069
      unfolding P_def by (simp add: prod_nat_ivl_Suc' algebra_simps)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3070
    also have "P x 0 = 1" by (simp add: P_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3071
    finally have "(\<lambda>n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3072
    from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3073
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3074
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3075
  have "continuous_on (ball 0 1) f"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3076
  proof (rule uniform_limit_theorem; (intro always_eventually allI)?)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3077
    show "uniform_limit (ball 0 1) (\<lambda>n x. \<Sum>k<n. P x k / of_nat (Suc k)^2) f sequentially"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3078
    proof (unfold f_def, rule weierstrass_m_test)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3079
      fix n :: nat and x :: real assume x: "x \<in> ball 0 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3080
      {
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3081
        fix k :: nat assume k: "k \<ge> 1"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3082
        from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3083
        also from k have "\<dots> \<le> of_nat k^2" by simp
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3084
        finally have "(1 - x^2 / of_nat k^2) \<in> {0..1}" using k
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3085
          by (simp_all add: field_simps del: of_nat_Suc)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3086
      }
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3087
      hence "(\<Prod>k=1..n. abs (1 - x^2 / of_nat k^2)) \<le> (\<Prod>k=1..n. 1)" by (intro prod_mono) simp
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3088
      thus "norm (P x n / (of_nat (Suc n)^2)) \<le> 1 / of_nat (Suc n)^2"
64272
f76b6dda2e56 setprod -> prod
nipkow
parents: 64267
diff changeset
  3089
        unfolding P_def by (simp add: field_simps abs_prod del: of_nat_Suc)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3090
    qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3091
  qed (auto simp: P_def intro!: continuous_intros)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3092
  hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3093
  hence "(f \<midarrow> 0 \<rightarrow> f 0)" by (simp add: isCont_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3094
  also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3095
  finally have "f \<midarrow> 0 \<rightarrow> K" .
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3096
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3097
  moreover have "f \<midarrow> 0 \<rightarrow> pi^2 / 6"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3098
  proof (rule Lim_transform_eventually)
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62534
diff changeset
  3099
    define f' where [abs_def]: "f' x = (\<Sum>n. - sin_coeff (n+3) * pi ^ (n+2) * x^n)" for x
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3100
    have "eventually (\<lambda>x. x \<noteq> (0::real)) (at 0)"
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3101
      by (auto simp add: eventually_at intro!: exI[of _ 1])
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3102
    thus "eventually (\<lambda>x. f' x = f x) (at 0)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3103
    proof eventually_elim
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3104
      fix x :: real assume x: "x \<noteq> 0"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3105
      have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3106
      with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3107
      have "(\<lambda>n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3108
        by (simp add: eval_nat_numeral)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3109
      from sums_divide[OF this, of "x^3 * pi"] x
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3110
        have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3111
        by (simp add: divide_simps eval_nat_numeral power_mult_distrib mult_ac)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3112
      with x have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3113
        by (simp add: g_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3114
      hence "f' x = g x / x^2" by (simp add: sums_iff f'_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3115
      also have "\<dots> = f x" using sums[of x] x by (simp add: sums_iff g_def f_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3116
      finally show "f' x = f x" .
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3117
    qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3118
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3119
    have "isCont f' 0" unfolding f'_def
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3120
    proof (intro isCont_powser_converges_everywhere)
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3121
      fix x :: real show "summable (\<lambda>n. -sin_coeff (n+3) * pi^(n+2) * x^n)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3122
      proof (cases "x = 0")
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3123
        assume x: "x \<noteq> 0"
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3124
        from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3125
               sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3126
          show ?thesis by (simp add: mult_ac power_mult_distrib divide_simps eval_nat_numeral)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3127
      qed (simp only: summable_0_powser)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3128
    qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3129
    hence "f' \<midarrow> 0 \<rightarrow> f' 0" by (simp add: isCont_def)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3130
    also have "f' 0 = pi * pi / fact 3" unfolding f'_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3131
      by (subst powser_zero) (simp add: sin_coeff_def)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3132
    finally show "f' \<midarrow> 0 \<rightarrow> pi^2 / 6" by (simp add: eval_nat_numeral)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3133
  qed
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3134
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3135
  ultimately have "K = pi^2 / 6" by (rule LIM_unique)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3136
  moreover from inverse_power_summable[of 2]
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3137
    have "summable (\<lambda>n. (inverse (real_of_nat (Suc n)))\<^sup>2)"
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3138
    by (subst summable_Suc_iff) (simp add: power_inverse)
62131
1baed43f453e nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
paulson
parents: 62085
diff changeset
  3139
  ultimately show ?thesis unfolding K_def
62049
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3140
    by (auto simp add: sums_iff power_divide inverse_eq_divide)
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3141
qed
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3142
b0f941e207cf Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
eberlm
parents:
diff changeset
  3143
end