| author | haftmann | 
| Thu, 08 Nov 2018 09:11:52 +0100 | |
| changeset 69260 | 0a9688695a1b | 
| parent 69064 | 5840724b1d71 | 
| child 69529 | 4ab9657b3257 | 
| permissions | -rw-r--r-- | 
| 63992 | 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 | 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 | 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 | 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  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
233  | 
subsection \<open>The Euler form and the logarithmic Gamma function\<close>  | 
| 
62049
 
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>  | 
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
236  | 
We define the Gamma function by first defining its multiplicative inverse \<open>rGamma\<close>.  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
237  | 
This is more convenient because \<open>rGamma\<close> 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.  | 
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
239  | 
(e.g. \<open>rGamma\<close> fulfils \<open>rGamma z = z * rGamma (z + 1)\<close> everywhere, whereas the \<open>\<Gamma>\<close> function  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
240  | 
does not fulfil the analogous equation on the non-positive integers)  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
241  | 
|
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
242  | 
We define the \<open>\<Gamma>\<close> function (resp.\ its reciprocale) in the Euler form. This form has the advantage  | 
| 
62131
 
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  | 
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
244  | 
(due to division by 0). The functional equation \<open>Gamma (z + 1) = z * Gamma z\<close> 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  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
248  | 
definition%important Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
 | 
| 
62049
 
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  | 
text \<open>  | 
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
308  | 
We now show that the series that defines the \<open>\<Gamma>\<close> function in the Euler form converges  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
309  | 
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
 | 
310  | 
real part.  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
311  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
312  | 
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
 | 
313  | 
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
 | 
314  | 
continuous.  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
315  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
316  | 
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
 | 
317  | 
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
 | 
318  | 
\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
319  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
320  | 
definition%important ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
 | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
321  | 
"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
 | 
322  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
323  | 
definition%unimportant ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
 | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
324  | 
"ln_Gamma_series' z n =  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
325  | 
- 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
 | 
326  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
327  | 
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
 | 
328  | 
"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
 | 
329  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
330  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
331  | 
text \<open>  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
332  | 
We now show that the log-Gamma series converges locally uniformly for all complex numbers except  | 
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
333  | 
the non-positive integers. We do this by proving that the series is locally Cauchy.  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
334  | 
\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
335  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
336  | 
context  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
337  | 
begin  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
338  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
339  | 
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
 | 
340  | 
fixes z :: complex and k :: nat  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
341  | 
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
 | 
342  | 
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
 | 
343  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
344  | 
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
 | 
345  | 
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
 | 
346  | 
by (simp add: algebra_simps)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
347  | 
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
 | 
348  | 
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
 | 
349  | 
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
 | 
350  | 
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
 | 
351  | 
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
 | 
352  | 
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
 | 
353  | 
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
 | 
354  | 
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
 | 
355  | 
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
 | 
356  | 
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
 | 
357  | 
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
 | 
358  | 
by (intro Ln_approx_linear) (simp add: norm_divide)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
359  | 
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
 | 
360  | 
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
 | 
361  | 
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
 | 
362  | 
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
 | 
363  | 
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
 | 
364  | 
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
 | 
365  | 
also note add_divide_distrib [symmetric]  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
366  | 
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
 | 
367  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
368  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
369  | 
lemma ln_Gamma_series_complex_converges:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
370  | 
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
 | 
371  | 
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
 | 
372  | 
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
 | 
373  | 
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
 | 
374  | 
fix e :: real assume e: "e > 0"  | 
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69064 
diff
changeset
 | 
375  | 
define e'' where "e'' = (SUP t\<in>ball z d. norm t + norm t^2)"  | 
| 63040 | 376  | 
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
 | 
377  | 
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
 | 
378  | 
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
 | 
379  | 
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
 | 
380  | 
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
 | 
381  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
382  | 
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
 | 
383  | 
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
 | 
384  | 
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
 | 
385  | 
by (rule cSUP_upper[OF _ bdd])  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
386  | 
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
 | 
387  | 
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
 | 
388  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
389  | 
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
 | 
390  | 
by (rule inverse_power_summable) simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
391  | 
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
 | 
392  | 
|
| 63040 | 393  | 
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
 | 
394  | 
  {
 | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
395  | 
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
 | 
396  | 
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
 | 
397  | 
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
 | 
398  | 
by (simp_all add: le_of_int_ceiling)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
399  | 
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
 | 
400  | 
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
 | 
401  | 
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
 | 
402  | 
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
 | 
403  | 
moreover have "(\<Sum>k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \<ge> N" for m n  | 
| 62072 | 404  | 
using M[OF order.trans[OF \<open>N \<ge> M\<close> that]] unfolding real_norm_def  | 
| 64267 | 405  | 
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
 | 
406  | 
moreover note calculation  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
407  | 
} note N = this  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
408  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
409  | 
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
 | 
410  | 
unfolding dist_complex_def  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
411  | 
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
 | 
412  | 
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
 | 
413  | 
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
 | 
414  | 
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
 | 
415  | 
by (simp add: dist_commute)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
416  | 
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
 | 
417  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
418  | 
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
 | 
419  | 
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
 | 
420  | 
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
 | 
421  | 
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
 | 
422  | 
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
 | 
423  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
424  | 
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
 | 
425  | 
(-(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
 | 
426  | 
((\<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
 | 
427  | 
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
 | 
428  | 
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
 | 
429  | 
                 (\<Sum>k\<in>{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn
 | 
| 64267 | 430  | 
by (simp add: sum_diff)  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
431  | 
    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
 | 
432  | 
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
 | 
433  | 
(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn  | 
| 64267 | 434  | 
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
 | 
435  | 
also have "... = (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N  | 
| 64267 | 436  | 
by (intro sum_cong_Suc)  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
437  | 
(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
 | 
438  | 
    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
 | 
439  | 
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
 | 
440  | 
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
 | 
441  | 
(\<Sum>k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N  | 
| 64267 | 442  | 
by (intro sum.cong) simp_all  | 
443  | 
also note sum.distrib [symmetric]  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
444  | 
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
 | 
445  | 
(\<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 | 446  | 
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
 | 
447  | 
also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)"  | 
| 64267 | 448  | 
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
 | 
449  | 
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
 | 
450  | 
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
 | 
451  | 
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
 | 
452  | 
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
 | 
453  | 
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
 | 
454  | 
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
 | 
455  | 
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
 | 
456  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
457  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
458  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
459  | 
end  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
460  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
461  | 
lemma ln_Gamma_series_complex_converges':  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
462  | 
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
 | 
463  | 
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
 | 
464  | 
proof -  | 
| 63040 | 465  | 
define d' where "d' = Re z"  | 
466  | 
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
 | 
467  | 
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
 | 
468  | 
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
 | 
469  | 
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
 | 
470  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
471  | 
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
 | 
472  | 
proof (cases "Re z > 0")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
473  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
474  | 
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
 | 
475  | 
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
 | 
476  | 
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
 | 
477  | 
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
 | 
478  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
479  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
480  | 
case False  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
481  | 
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
 | 
482  | 
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
 | 
483  | 
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
 | 
484  | 
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
 | 
485  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
486  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
487  | 
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
 | 
488  | 
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
 | 
489  | 
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
 | 
490  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
491  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
492  | 
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
 | 
493  | 
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
 | 
494  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
495  | 
theorem ln_Gamma_complex_LIMSEQ: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma_series z \<longlonglongrightarrow> ln_Gamma z"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
496  | 
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
 | 
497  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
498  | 
lemma exp_ln_Gamma_series_complex:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
499  | 
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
 | 
500  | 
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
 | 
501  | 
proof -  | 
| 
63417
 
c184ec919c70
more lemmas to emphasize {0::nat..(<)n} as canonical representation of intervals on nat
 
haftmann 
parents: 
63367 
diff
changeset
 | 
502  | 
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
 | 
503  | 
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
 | 
504  | 
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
 | 
505  | 
(of_nat n) powr z / (z * (\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))))"  | 
| 64267 | 506  | 
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
 | 
507  | 
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 | 508  | 
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
 | 
509  | 
also have "... = (\<Prod>k=1..n. z + k) / fact n"  | 
| 64272 | 510  | 
by (simp add: fact_prod)  | 
511  | 
(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
 | 
512  | 
also from m have "z * ... = (\<Prod>k=0..n. z + k) / fact n"  | 
| 64272 | 513  | 
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
 | 
514  | 
also have "(\<Prod>k=0..n. z + k) = pochhammer z (Suc n)"  | 
| 64272 | 515  | 
unfolding pochhammer_prod  | 
516  | 
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
 | 
517  | 
also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n"  | 
| 66936 | 518  | 
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
 | 
519  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
520  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
521  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
522  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
523  | 
lemma ln_Gamma_series'_aux:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
524  | 
assumes "(z::complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
525  | 
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
 | 
526  | 
(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
 | 
527  | 
unfolding sums_def  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
528  | 
proof (rule Lim_transform)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
529  | 
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
 | 
530  | 
(is "?g \<longlonglongrightarrow> _")  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
531  | 
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
 | 
532  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
533  | 
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
 | 
534  | 
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
 | 
535  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
536  | 
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
 | 
537  | 
have "(\<Sum>k<n. ?f k) = (\<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k))"  | 
| 64267 | 538  | 
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
 | 
539  | 
subst atLeastLessThanSuc_atLeastAtMost) simp_all  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
540  | 
also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"  | 
| 64267 | 541  | 
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
 | 
542  | 
also from n have "\<dots> - ?g n = 0"  | 
| 64267 | 543  | 
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
 | 
544  | 
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
 | 
545  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
546  | 
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
 | 
547  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
548  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
549  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
550  | 
lemma uniformly_summable_deriv_ln_Gamma:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
551  | 
  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
 | 
552  | 
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
 | 
553  | 
(\<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
 | 
554  | 
(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
 | 
555  | 
proof (rule weierstrass_m_test'_ev)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
556  | 
  {
 | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
557  | 
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
 | 
558  | 
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
 | 
559  | 
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
 | 
560  | 
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
 | 
561  | 
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
 | 
562  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
563  | 
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
 | 
564  | 
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
 | 
565  | 
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
 | 
566  | 
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
 | 
567  | 
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
 | 
568  | 
note A B  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
569  | 
} note ball = this  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
570  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
571  | 
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
 | 
572  | 
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
 | 
573  | 
proof safe  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
574  | 
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
 | 
575  | 
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
 | 
576  | 
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
 | 
577  | 
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
 | 
578  | 
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
 | 
579  | 
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
 | 
580  | 
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
 | 
581  | 
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
 | 
582  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
583  | 
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
 | 
584  | 
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
 | 
585  | 
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
 | 
586  | 
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
 | 
587  | 
    also {
 | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
588  | 
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
 | 
589  | 
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
 | 
590  | 
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
 | 
591  | 
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
 | 
592  | 
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
 | 
593  | 
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
 | 
594  | 
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
 | 
595  | 
}  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
596  | 
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
 | 
597  | 
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
 | 
598  | 
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
 | 
599  | 
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
 | 
600  | 
by (simp del: of_nat_Suc)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
601  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
602  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
603  | 
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
 | 
604  | 
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
 | 
605  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
606  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
607  | 
|
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
608  | 
subsection \<open>The Polygamma functions\<close>  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
609  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
610  | 
lemma summable_deriv_ln_Gamma:  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
611  | 
  "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
 | 
612  | 
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
 | 
613  | 
unfolding summable_iff_convergent  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
614  | 
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
 | 
615  | 
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
 | 
616  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
617  | 
definition%important 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
 | 
618  | 
"Polygamma n z = (if n = 0 then  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
619  | 
(\<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
 | 
620  | 
(-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
 | 
621  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
622  | 
abbreviation%important Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
 | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
623  | 
"Digamma \<equiv> Polygamma 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
624  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
625  | 
lemma Digamma_def:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
626  | 
"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
 | 
627  | 
by (simp add: Polygamma_def)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
628  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
629  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
630  | 
lemma summable_Digamma:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
631  | 
  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
 | 
632  | 
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
 | 
633  | 
proof -  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
634  | 
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
 | 
635  | 
(0 - inverse (z + of_nat 0))"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
636  | 
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
 | 
637  | 
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
 | 
638  | 
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
 | 
639  | 
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
 | 
640  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
641  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
642  | 
lemma summable_offset:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
643  | 
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
 | 
644  | 
shows "summable f"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
645  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
646  | 
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
 | 
647  | 
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
 | 
648  | 
by (intro convergent_add convergent_const)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
649  | 
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
 | 
650  | 
proof  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
651  | 
fix m :: nat  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
652  | 
    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
 | 
653  | 
also have "(\<Sum>n\<in>\<dots>. f n) = (\<Sum>n<k. f n) + (\<Sum>n=k..<m+k. f n)"  | 
| 64267 | 654  | 
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
 | 
655  | 
also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))"  | 
| 66936 | 656  | 
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
 | 
657  | 
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
 | 
658  | 
qed  | 
| 64267 | 659  | 
  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
 | 
660  | 
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
 | 
661  | 
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
 | 
662  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
663  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
664  | 
lemma Polygamma_converges:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
665  | 
  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
 | 
666  | 
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
 | 
667  | 
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
 | 
668  | 
proof (rule weierstrass_m_test'_ev)  | 
| 63040 | 669  | 
define e where "e = (1 + d / norm z)"  | 
670  | 
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
 | 
671  | 
  {
 | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
672  | 
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
 | 
673  | 
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
 | 
674  | 
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
 | 
675  | 
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
 | 
676  | 
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
 | 
677  | 
} note ball = this  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
678  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
679  | 
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
 | 
680  | 
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
 | 
681  | 
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
 | 
682  | 
proof (intro ballI)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
683  | 
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
 | 
684  | 
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
 | 
685  | 
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
 | 
686  | 
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
 | 
687  | 
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
 | 
688  | 
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
 | 
689  | 
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
 | 
690  | 
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
 | 
691  | 
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
 | 
692  | 
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
 | 
693  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
694  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
695  | 
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
 | 
696  | 
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
 | 
697  | 
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
 | 
698  | 
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
 | 
699  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
700  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
701  | 
lemma Polygamma_converges':  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
702  | 
  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
 | 
703  | 
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
 | 
704  | 
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
 | 
705  | 
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
 | 
706  | 
by (simp add: summable_iff_convergent)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
707  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
708  | 
theorem Digamma_LIMSEQ:  | 
| 63721 | 709  | 
  fixes z :: "'a :: {banach,real_normed_field}"
 | 
710  | 
assumes z: "z \<noteq> 0"  | 
|
711  | 
shows "(\<lambda>m. of_real (ln (real m)) - (\<Sum>n<m. inverse (z + of_nat n))) \<longlonglongrightarrow> Digamma z"  | 
|
712  | 
proof -  | 
|
713  | 
have "(\<lambda>n. of_real (ln (real n / (real (Suc n))))) \<longlonglongrightarrow> (of_real (ln 1) :: 'a)"  | 
|
714  | 
by (intro tendsto_intros LIMSEQ_n_over_Suc_n) simp_all  | 
|
715  | 
hence "(\<lambda>n. of_real (ln (real n / (real n + 1)))) \<longlonglongrightarrow> (0 :: 'a)" by (simp add: add_ac)  | 
|
716  | 
hence lim: "(\<lambda>n. of_real (ln (real n)) - of_real (ln (real n + 1))) \<longlonglongrightarrow> (0::'a)"  | 
|
717  | 
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
 | 
718  | 
show "eventually (\<lambda>n. of_real (ln (real n / (real n + 1))) =  | 
| 63721 | 719  | 
of_real (ln (real n)) - (of_real (ln (real n + 1)) :: 'a)) at_top"  | 
720  | 
using eventually_gt_at_top[of "0::nat"] by eventually_elim (simp add: ln_div)  | 
|
721  | 
qed  | 
|
722  | 
||
723  | 
from summable_Digamma[OF z]  | 
|
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
724  | 
have "(\<lambda>n. inverse (of_nat (n+1)) - inverse (z + of_nat n))  | 
| 63721 | 725  | 
sums (Digamma z + euler_mascheroni)"  | 
726  | 
by (simp add: Digamma_def summable_sums)  | 
|
727  | 
from sums_diff[OF this euler_mascheroni_sum]  | 
|
728  | 
have "(\<lambda>n. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)) - inverse (z + of_nat n))  | 
|
729  | 
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
 | 
730  | 
hence "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1))) -  | 
| 63721 | 731  | 
(\<Sum>n<m. inverse (z + of_nat n))) \<longlonglongrightarrow> Digamma z"  | 
| 64267 | 732  | 
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
 | 
733  | 
also have "(\<lambda>m. (\<Sum>n<m. of_real (ln (real (Suc n) + 1)) - of_real (ln (real n + 1)))) =  | 
| 63721 | 734  | 
(\<lambda>m. of_real (ln (m + 1)) :: 'a)"  | 
| 64267 | 735  | 
by (subst sum_lessThan_telescope) simp_all  | 
| 63721 | 736  | 
finally show ?thesis by (rule Lim_transform) (insert lim, simp)  | 
737  | 
qed  | 
|
738  | 
||
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
739  | 
theorem Polygamma_LIMSEQ:  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
740  | 
  fixes z :: "'a :: {banach,real_normed_field}"
 | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
741  | 
assumes "z \<noteq> 0" and "n > 0"  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
742  | 
shows "(\<lambda>k. inverse ((z + of_nat k)^Suc n)) sums ((-1) ^ Suc n * Polygamma n z / fact n)"  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
743  | 
using Polygamma_converges'[OF assms(1), of "Suc n"] assms(2)  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
744  | 
by (simp add: sums_iff Polygamma_def)  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
745  | 
|
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
746  | 
theorem has_field_derivative_ln_Gamma_complex [derivative_intros]:  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
747  | 
fixes z :: complex  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
748  | 
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
 | 
749  | 
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
 | 
750  | 
proof -  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
751  | 
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
 | 
752  | 
using that by (auto elim!: nonpos_Ints_cases')  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
753  | 
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
 | 
754  | 
by blast+  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
755  | 
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
 | 
756  | 
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 | 757  | 
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
 | 
758  | 
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
 | 
759  | 
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
 | 
760  | 
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
 | 
761  | 
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
 | 
762  | 
(0 - inverse (z + of_nat 0))"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
763  | 
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
 | 
764  | 
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
 | 
765  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
766  | 
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
 | 
767  | 
using d z ln_Gamma_series'_aux[OF z']  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
768  | 
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
 | 
769  | 
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
 | 
770  | 
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
 | 
771  | 
simp del: of_nat_Suc)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
772  | 
apply (auto simp add: complex_nonpos_Reals_iff)  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
773  | 
done  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
774  | 
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
 | 
775  | 
?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
 | 
776  | 
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
 | 
777  | 
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
 | 
778  | 
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
 | 
779  | 
by (simp add: sums_iff)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
780  | 
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
 | 
781  | 
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
 | 
782  | 
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
 | 
783  | 
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
 | 
784  | 
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
 | 
785  | 
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
 | 
786  | 
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
 | 
787  | 
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
 | 
788  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
789  | 
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
 | 
790  | 
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
 | 
791  | 
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
 | 
792  | 
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
 | 
793  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
794  | 
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
 | 
795  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
796  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
797  | 
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
 | 
798  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
799  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
800  | 
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
 | 
801  | 
by (simp add: Digamma_def)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
802  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
803  | 
lemma Digamma_plus1:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
804  | 
assumes "z \<noteq> 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
805  | 
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
 | 
806  | 
proof -  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
807  | 
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
 | 
808  | 
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
 | 
809  | 
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
 | 
810  | 
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
 | 
811  | 
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
 | 
812  | 
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
 | 
813  | 
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
 | 
814  | 
(\<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
 | 
815  | 
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
 | 
816  | 
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
 | 
817  | 
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
 | 
818  | 
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
 | 
819  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
820  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
821  | 
theorem Polygamma_plus1:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
822  | 
assumes "z \<noteq> 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
823  | 
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
 | 
824  | 
proof (cases "n = 0")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
825  | 
assume n: "n \<noteq> 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
826  | 
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
 | 
827  | 
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
 | 
828  | 
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
 | 
829  | 
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
 | 
830  | 
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
 | 
831  | 
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
 | 
832  | 
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
 | 
833  | 
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
 | 
834  | 
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
 | 
835  | 
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
 | 
836  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
837  | 
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
 | 
838  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
839  | 
theorem Digamma_of_nat:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
840  | 
  "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
 | 
841  | 
proof (induction n)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
842  | 
case (Suc n)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
843  | 
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
 | 
844  | 
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
 | 
845  | 
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
 | 
846  | 
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
 | 
847  | 
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
 | 
848  | 
by (simp add: harm_Suc)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
849  | 
finally show ?case .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
850  | 
qed (simp add: harm_def)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
851  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
852  | 
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
 | 
853  | 
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
 | 
854  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
855  | 
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
 | 
856  | 
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
 | 
857  | 
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
 | 
858  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
859  | 
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
 | 
860  | 
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
 | 
861  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
862  | 
lemma Digamma_half_integer:  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
863  | 
  "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
 | 
864  | 
(\<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
 | 
865  | 
proof (induction n)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
866  | 
case 0  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
867  | 
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
 | 
868  | 
also have "Digamma (1/2::real) =  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
869  | 
(\<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
 | 
870  | 
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
 | 
871  | 
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
 | 
872  | 
(\<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
 | 
873  | 
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
 | 
874  | 
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
 | 
875  | 
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
 | 
876  | 
finally show ?case by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
877  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
878  | 
case (Suc n)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
879  | 
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
 | 
880  | 
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
 | 
881  | 
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
 | 
882  | 
have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
883  | 
also from nz' have "\<dots> = Digamma (of_nat n + 1/2) + 1 / (of_nat n + 1/2)"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
884  | 
by (rule Digamma_plus1)  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
885  | 
also from nz nz' have "1 / (of_nat n + 1/2 :: 'a) = 2 / (2 * of_nat n + 1)"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
886  | 
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
 | 
887  | 
also note Suc  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
888  | 
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
 | 
889  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
890  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
891  | 
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
 | 
892  | 
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
 | 
893  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
894  | 
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
 | 
895  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
896  | 
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
 | 
897  | 
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
 | 
898  | 
also note euler_mascheroni_less_13_over_22  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
899  | 
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
 | 
900  | 
finally show ?thesis by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
901  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
902  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
903  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
904  | 
theorem has_field_derivative_Polygamma [derivative_intros]:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
905  | 
  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
 | 
906  | 
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
 | 
907  | 
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
 | 
908  | 
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
 | 
909  | 
assume n: "n = 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
910  | 
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
 | 
911  | 
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
 | 
912  | 
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
 | 
913  | 
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
 | 
914  | 
by (intro summable_Digamma) force  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
915  | 
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
 | 
916  | 
by (intro Polygamma_converges) auto  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
917  | 
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
 | 
918  | 
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
 | 
919  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
920  | 
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
 | 
921  | 
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
 | 
922  | 
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
 | 
923  | 
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
 | 
924  | 
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
 | 
925  | 
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
 | 
926  | 
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
 | 
927  | 
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
 | 
928  | 
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
 | 
929  | 
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
 | 
930  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
931  | 
assume n: "n \<noteq> 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
932  | 
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
 | 
933  | 
from no_nonpos_Int_in_ball'[OF z] guess d . note d = this  | 
| 63040 | 934  | 
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
 | 
935  | 
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
 | 
936  | 
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
 | 
937  | 
(\<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
 | 
938  | 
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
 | 
939  | 
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
 | 
940  | 
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
 | 
941  | 
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
 | 
942  | 
- 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
 | 
943  | 
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
 | 
944  | 
next  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
945  | 
have "uniformly_convergent_on (ball z d)  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
946  | 
(\<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
 | 
947  | 
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
 | 
948  | 
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
 | 
949  | 
(\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))"  | 
| 64267 | 950  | 
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
 | 
951  | 
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
 | 
952  | 
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
 | 
953  | 
(- 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
 | 
954  | 
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
 | 
955  | 
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
 | 
956  | 
- 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
 | 
957  | 
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
 | 
958  | 
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
 | 
959  | 
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
 | 
960  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
961  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
962  | 
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
 | 
963  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
964  | 
lemma isCont_Polygamma [continuous_intros]:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
965  | 
  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
 | 
966  | 
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
 | 
967  | 
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
 | 
968  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
969  | 
lemma continuous_on_Polygamma:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
970  | 
  "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
 | 
971  | 
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
 | 
972  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
973  | 
lemma isCont_ln_Gamma_complex [continuous_intros]:  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
974  | 
fixes f :: "'a::t2_space \<Rightarrow> complex"  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
975  | 
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
 | 
976  | 
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
 | 
977  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
978  | 
lemma continuous_on_ln_Gamma_complex [continuous_intros]:  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
979  | 
fixes A :: "complex set"  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
980  | 
  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
 | 
981  | 
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
 | 
982  | 
fastforce  | 
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
983  | 
|
| 64969 | 984  | 
lemma deriv_Polygamma:  | 
985  | 
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
 | 
986  | 
shows "deriv (Polygamma m) z =  | 
| 64969 | 987  | 
             Polygamma (Suc m) (z :: 'a :: {real_normed_field,euclidean_space})"
 | 
988  | 
by (intro DERIV_imp_deriv has_field_derivative_Polygamma assms)  | 
|
989  | 
thm has_field_derivative_Polygamma  | 
|
990  | 
||
991  | 
lemma higher_deriv_Polygamma:  | 
|
992  | 
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
 | 
993  | 
shows "(deriv ^^ n) (Polygamma m) z =  | 
| 64969 | 994  | 
             Polygamma (m + n) (z :: 'a :: {real_normed_field,euclidean_space})"
 | 
995  | 
proof -  | 
|
996  | 
have "eventually (\<lambda>u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)"  | 
|
997  | 
proof (induction n)  | 
|
998  | 
case (Suc n)  | 
|
999  | 
from Suc.IH have "eventually (\<lambda>z. eventually (\<lambda>u. (deriv ^^ n) (Polygamma m) u = Polygamma (m + n) u) (nhds z)) (nhds z)"  | 
|
1000  | 
by (simp add: eventually_eventually)  | 
|
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1001  | 
hence "eventually (\<lambda>z. deriv ((deriv ^^ n) (Polygamma m)) z =  | 
| 64969 | 1002  | 
deriv (Polygamma (m + n)) z) (nhds z)"  | 
1003  | 
by eventually_elim (intro deriv_cong_ev refl)  | 
|
1004  | 
moreover have "eventually (\<lambda>z. z \<in> UNIV - \<int>\<^sub>\<le>\<^sub>0) (nhds z)" using assms  | 
|
1005  | 
by (intro eventually_nhds_in_open open_Diff open_UNIV) auto  | 
|
1006  | 
ultimately show ?case by eventually_elim (simp_all add: deriv_Polygamma)  | 
|
1007  | 
qed simp_all  | 
|
1008  | 
thus ?thesis by (rule eventually_nhds_x_imp_x)  | 
|
1009  | 
qed  | 
|
1010  | 
||
1011  | 
lemma deriv_ln_Gamma_complex:  | 
|
1012  | 
assumes "z \<notin> \<real>\<^sub>\<le>\<^sub>0"  | 
|
1013  | 
shows "deriv ln_Gamma z = Digamma (z :: complex)"  | 
|
1014  | 
by (intro DERIV_imp_deriv has_field_derivative_ln_Gamma_complex assms)  | 
|
1015  | 
||
1016  | 
||
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1017  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1018  | 
text \<open>  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1019  | 
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
 | 
1020  | 
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
 | 
1021  | 
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
 | 
1022  | 
\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1023  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1024  | 
class%unimportant Gamma = real_normed_field + complete_space +  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1025  | 
fixes rGamma :: "'a \<Rightarrow> 'a"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1026  | 
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
 | 
1027  | 
assumes differentiable_rGamma_aux1:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1028  | 
"(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1029  | 
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
 | 
1030  | 
\<longlonglongrightarrow> d) - scaleR euler_mascheroni 1  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1031  | 
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
 | 
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 differentiable_rGamma_aux2:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1034  | 
"let z = - of_nat n  | 
| 64272 | 1035  | 
     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
 | 
1036  | 
norm (y - z)) (nhds 0) (at z)"  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1037  | 
assumes rGamma_series_aux: "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>  | 
| 64272 | 1038  | 
             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
 | 
1039  | 
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
 | 
1040  | 
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
 | 
1041  | 
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
 | 
1042  | 
(nhds (rGamma z)) sequentially"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1043  | 
begin  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1044  | 
subclass banach ..  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1045  | 
end  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1046  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1047  | 
definition "Gamma z = inverse (rGamma z)"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1048  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1049  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1050  | 
subsection \<open>Basic properties\<close>  | 
| 
 
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_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
 | 
1053  | 
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
 | 
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  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1056  | 
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
 | 
1057  | 
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
 | 
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  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1060  | 
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
 | 
1061  | 
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
 | 
1062  | 
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
 | 
1063  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1064  | 
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
 | 
1065  | 
unfolding Gamma_def by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1066  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1067  | 
lemma rGamma_series_LIMSEQ [tendsto_intros]:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1068  | 
"rGamma_series z \<longlonglongrightarrow> rGamma z"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1069  | 
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
 | 
1070  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1071  | 
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
 | 
1072  | 
from rGamma_series_aux[OF this] show ?thesis  | 
| 64272 | 1073  | 
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
 | 
1074  | 
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
 | 
1075  | 
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
 | 
1076  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1077  | 
theorem Gamma_series_LIMSEQ [tendsto_intros]:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1078  | 
"Gamma_series z \<longlonglongrightarrow> Gamma z"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1079  | 
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
 | 
1080  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1081  | 
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
 | 
1082  | 
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
 | 
1083  | 
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
 | 
1084  | 
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
 | 
1085  | 
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
 | 
1086  | 
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
 | 
1087  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1088  | 
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
 | 
1089  | 
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
 | 
1090  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1091  | 
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
 | 
1092  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1093  | 
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
 | 
1094  | 
using eventually_gt_at_top[of "0::nat"]  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1095  | 
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
 | 
1096  | 
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
 | 
1097  | 
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
 | 
1098  | 
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
 | 
1099  | 
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
 | 
1100  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1101  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1102  | 
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
 | 
1103  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1104  | 
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
 | 
1105  | 
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
 | 
1106  | 
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
 | 
1107  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1108  | 
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
 | 
1109  | 
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
 | 
1110  | 
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
 | 
1111  | 
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
 | 
1112  | 
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
 | 
1113  | 
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
 | 
1114  | 
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
 | 
1115  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1116  | 
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
 | 
1117  | 
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
 | 
1118  | 
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
 | 
1119  | 
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
 | 
1120  | 
by (rule Lim_transform_eventually)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1121  | 
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
 | 
1122  | 
by (intro tendsto_intros)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1123  | 
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
 | 
1124  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1125  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1126  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1127  | 
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
 | 
1128  | 
proof (induction n arbitrary: z)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1129  | 
case (Suc n z)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1130  | 
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
 | 
1131  | 
also note rGamma_plus1 [symmetric]  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1132  | 
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
 | 
1133  | 
qed simp_all  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1134  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1135  | 
theorem Gamma_plus1: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma (z + 1) = z * Gamma z"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1136  | 
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
 | 
1137  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1138  | 
theorem 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
 | 
1139  | 
using pochhammer_rGamma[of z]  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1140  | 
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
 | 
1141  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1142  | 
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
 | 
1143  | 
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
 | 
1144  | 
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
 | 
1145  | 
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
 | 
1146  | 
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
 | 
1147  | 
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
 | 
1148  | 
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
 | 
1149  | 
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
 | 
1150  | 
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
 | 
1151  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1152  | 
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
 | 
1153  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1154  | 
theorem Gamma_fact: "Gamma (1 + of_nat n) = fact n"  | 
| 68403 | 1155  | 
by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff flip: of_nat_Suc)  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1156  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1157  | 
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
 | 
1158  | 
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
 | 
1159  | 
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
 | 
1160  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1161  | 
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
 | 
1162  | 
proof (cases "n > 0")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1163  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1164  | 
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
 | 
1165  | 
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
 | 
1166  | 
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
 | 
1167  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1168  | 
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
 | 
1169  | 
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
 | 
1170  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1171  | 
lemma Gamma_seriesI:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1172  | 
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
 | 
1173  | 
shows "g \<longlonglongrightarrow> Gamma z"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1174  | 
proof (rule Lim_transform_eventually)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1175  | 
have "1/2 > (0::real)" by simp  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1176  | 
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
 | 
1177  | 
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
 | 
1178  | 
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
 | 
1179  | 
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
 | 
1180  | 
by (intro tendsto_intros)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1181  | 
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
 | 
1182  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1183  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1184  | 
lemma Gamma_seriesI':  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1185  | 
assumes "f \<longlonglongrightarrow> rGamma z"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1186  | 
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
 | 
1187  | 
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
 | 
1188  | 
shows "g \<longlonglongrightarrow> Gamma z"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1189  | 
proof (rule Lim_transform_eventually)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1190  | 
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
 | 
1191  | 
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
 | 
1192  | 
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
 | 
1193  | 
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
 | 
1194  | 
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
 | 
1195  | 
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
 | 
1196  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1197  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1198  | 
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
 | 
1199  | 
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
 | 
1200  | 
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
 | 
1201  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1202  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1203  | 
subsection \<open>Differentiability\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1204  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1205  | 
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
 | 
1206  | 
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
 | 
1207  | 
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
 | 
1208  | 
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
 | 
1209  | 
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
 | 
1210  | 
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
 | 
1211  | 
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
 | 
1212  | 
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
 | 
1213  | 
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
 | 
1214  | 
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
 | 
1215  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1216  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1217  | 
lemma has_field_derivative_rGamma_nonpos_int:  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1218  | 
"(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
 | 
1219  | 
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
 | 
1220  | 
using differentiable_rGamma_aux2[of n]  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1221  | 
unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at  | 
| 64272 | 1222  | 
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
 | 
1223  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1224  | 
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
 | 
1225  | 
"(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
 | 
1226  | 
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
 | 
1227  | 
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
 | 
1228  | 
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
 | 
1229  | 
by (auto elim!: nonpos_Ints_cases')  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1230  | 
|
| 
 
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 [THEN DERIV_chain2, 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 [THEN DERIV_chain2, derivative_intros]  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1233  | 
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
 | 
1234  | 
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
 | 
1235  | 
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
 | 
1236  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1237  | 
theorem 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
 | 
1238  | 
"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
 | 
1239  | 
unfolding Gamma_def [abs_def]  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1240  | 
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
 | 
1241  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1242  | 
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
 | 
1243  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1244  | 
(* TODO: Hide ugly facts properly *)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1245  | 
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
 | 
1246  | 
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
 | 
1247  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1248  | 
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
 | 
1249  | 
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
 | 
1250  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1251  | 
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
 | 
1252  | 
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
 | 
1253  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1254  | 
lemma isCont_rGamma [continuous_intros]:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1255  | 
"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
 | 
1256  | 
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
 | 
1257  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1258  | 
lemma isCont_Gamma [continuous_intros]:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1259  | 
"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
 | 
1260  | 
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
 | 
1261  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1262  | 
subsection%unimportant \<open>The complex Gamma function\<close>  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1263  | 
|
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1264  | 
instantiation%unimportant complex :: Gamma  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1265  | 
begin  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1266  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1267  | 
definition%unimportant rGamma_complex :: "complex \<Rightarrow> complex" where  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1268  | 
"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
 | 
1269  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1270  | 
lemma rGamma_series_complex_converges:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1271  | 
"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
 | 
1272  | 
and rGamma_complex_altdef:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1273  | 
"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
 | 
1274  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1275  | 
have "?thesis1 \<and> ?thesis2"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1276  | 
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
 | 
1277  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1278  | 
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
 | 
1279  | 
proof (rule Lim_transform_eventually)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1280  | 
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
 | 
1281  | 
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
 | 
1282  | 
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
 | 
1283  | 
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
 | 
1284  | 
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
 | 
1285  | 
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
 | 
1286  | 
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
 | 
1287  | 
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
 | 
1288  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1289  | 
with False show ?thesis  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1290  | 
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
 | 
1291  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1292  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1293  | 
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
 | 
1294  | 
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
 | 
1295  | 
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
 | 
1296  | 
finally show ?thesis using True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1297  | 
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
 | 
1298  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1299  | 
thus "?thesis1" "?thesis2" by blast+  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1300  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1301  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1302  | 
context%unimportant  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1303  | 
begin  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1304  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1305  | 
(* TODO: duplication *)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1306  | 
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
 | 
1307  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1308  | 
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
 | 
1309  | 
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
 | 
1310  | 
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
 | 
1311  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1312  | 
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
 | 
1313  | 
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
 | 
1314  | 
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
 | 
1315  | 
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
 | 
1316  | 
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
 | 
1317  | 
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
 | 
1318  | 
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
 | 
1319  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1320  | 
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
 | 
1321  | 
using rGamma_series_complex_converges  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1322  | 
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
 | 
1323  | 
(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
 | 
1324  | 
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
 | 
1325  | 
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
 | 
1326  | 
by (rule Lim_transform_eventually)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1327  | 
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
 | 
1328  | 
using rGamma_series_complex_converges  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1329  | 
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
 | 
1330  | 
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
 | 
1331  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1332  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1333  | 
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
 | 
1334  | 
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
 | 
1335  | 
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
 | 
1336  | 
proof -  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1337  | 
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
 | 
1338  | 
proof (subst DERIV_cong_ev[OF refl _ refl])  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1339  | 
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
 | 
1340  | 
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
 | 
1341  | 
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
 | 
1342  | 
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
 | 
1343  | 
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
 | 
1344  | 
next  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1345  | 
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
 | 
1346  | 
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
 | 
1347  | 
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
 | 
1348  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1349  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1350  | 
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
 | 
1351  | 
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
 | 
1352  | 
case (Suc n z)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1353  | 
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
 | 
1354  | 
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
 | 
1355  | 
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
 | 
1356  | 
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
 | 
1357  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1358  | 
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
 | 
1359  | 
-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
 | 
1360  | 
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
 | 
1361  | 
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
 | 
1362  | 
by (simp add: rGamma_complex_plus1)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1363  | 
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
 | 
1364  | 
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
 | 
1365  | 
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
 | 
1366  | 
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
 | 
1367  | 
finally show ?case .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1368  | 
qed (intro diff, simp)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1369  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1370  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1371  | 
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
 | 
1372  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1373  | 
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
 | 
1374  | 
using eventually_gt_at_top[of "0::nat"]  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1375  | 
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
 | 
1376  | 
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
 | 
1377  | 
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
 | 
1378  | 
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
 | 
1379  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1380  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1381  | 
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
 | 
1382  | 
"(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
 | 
1383  | 
proof (induction n)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1384  | 
case 0  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1385  | 
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
 | 
1386  | 
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
 | 
1387  | 
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
 | 
1388  | 
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
 | 
1389  | 
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
 | 
1390  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1391  | 
case (Suc n)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1392  | 
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
 | 
1393  | 
(at (- of_nat (Suc n) + 1 :: complex))" by simp  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1394  | 
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
 | 
1395  | 
(- 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
 | 
1396  | 
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
 | 
1397  | 
(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
 | 
1398  | 
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
 | 
1399  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1400  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1401  | 
instance proof  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1402  | 
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
 | 
1403  | 
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
 | 
1404  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1405  | 
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
 | 
1406  | 
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
 | 
1407  | 
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
 | 
1408  | 
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
 | 
1409  | 
\<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
 | 
1410  | 
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
 | 
1411  | 
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
 | 
1412  | 
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
 | 
1413  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1414  | 
fix n :: nat  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1415  | 
from has_field_derivative_rGamma_complex_nonpos_Int[of n]  | 
| 64272 | 1416  | 
  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
 | 
1417  | 
(y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0"  | 
| 64272 | 1418  | 
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
 | 
1419  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1420  | 
fix z :: complex  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1421  | 
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
 | 
1422  | 
by (simp add: convergent_LIMSEQ_iff rGamma_complex_def)  | 
| 64272 | 1423  | 
  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
 | 
1424  | 
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
 | 
1425  | 
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
 | 
1426  | 
in (\<lambda>n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma z"  | 
| 64272 | 1427  | 
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
 | 
1428  | 
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
 | 
1429  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1430  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1431  | 
end  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1432  | 
end  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1433  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1434  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1435  | 
lemma Gamma_complex_altdef:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1436  | 
"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
 | 
1437  | 
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
 | 
1438  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1439  | 
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
 | 
1440  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1441  | 
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
 | 
1442  | 
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
 | 
1443  | 
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
 | 
1444  | 
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
 | 
1445  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1446  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1447  | 
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
 | 
1448  | 
unfolding Gamma_def by (simp add: cnj_rGamma)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1449  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1450  | 
lemma Gamma_complex_real:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1451  | 
"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
 | 
1452  | 
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
 | 
1453  | 
|
| 
62534
 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 
paulson <lp15@cam.ac.uk> 
parents: 
62533 
diff
changeset
 | 
1454  | 
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
 | 
1455  | 
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
 | 
1456  | 
|
| 64969 | 1457  | 
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
 | 
1458  | 
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
 | 
1459  | 
|
| 68721 | 1460  | 
lemma holomorphic_rGamma' [holomorphic_intros]:  | 
1461  | 
assumes "f holomorphic_on A"  | 
|
1462  | 
shows "(\<lambda>x. rGamma (f x)) holomorphic_on A"  | 
|
1463  | 
proof -  | 
|
1464  | 
have "rGamma \<circ> f holomorphic_on A" using assms  | 
|
1465  | 
by (intro holomorphic_on_compose assms holomorphic_rGamma)  | 
|
1466  | 
thus ?thesis by (simp only: o_def)  | 
|
1467  | 
qed  | 
|
1468  | 
||
| 64969 | 1469  | 
lemma analytic_rGamma: "rGamma analytic_on A"  | 
1470  | 
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
 | 
1471  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1472  | 
|
| 
62534
 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 
paulson <lp15@cam.ac.uk> 
parents: 
62533 
diff
changeset
 | 
1473  | 
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
 | 
1474  | 
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
 | 
1475  | 
|
| 64969 | 1476  | 
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
 | 
1477  | 
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
 | 
1478  | 
|
| 68721 | 1479  | 
lemma holomorphic_Gamma' [holomorphic_intros]:  | 
1480  | 
assumes "f holomorphic_on A" and "\<And>x. x \<in> A \<Longrightarrow> f x \<notin> \<int>\<^sub>\<le>\<^sub>0"  | 
|
1481  | 
shows "(\<lambda>x. Gamma (f x)) holomorphic_on A"  | 
|
1482  | 
proof -  | 
|
1483  | 
have "Gamma \<circ> f holomorphic_on A" using assms  | 
|
1484  | 
by (intro holomorphic_on_compose assms holomorphic_Gamma) auto  | 
|
1485  | 
thus ?thesis by (simp only: o_def)  | 
|
1486  | 
qed  | 
|
1487  | 
||
| 64969 | 1488  | 
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
 | 
1489  | 
by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)  | 
| 64969 | 1490  | 
(auto intro!: holomorphic_Gamma)  | 
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1491  | 
|
| 
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1492  | 
|
| 
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1493  | 
lemma field_differentiable_ln_Gamma_complex:  | 
| 64969 | 1494  | 
"z \<notin> \<real>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma field_differentiable (at (z::complex) within A)"  | 
1495  | 
by (rule field_differentiable_within_subset[of _ _ UNIV])  | 
|
1496  | 
(force simp: field_differentiable_def intro!: derivative_intros)+  | 
|
1497  | 
||
1498  | 
lemma holomorphic_ln_Gamma [holomorphic_intros]: "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> ln_Gamma holomorphic_on A"
 | 
|
1499  | 
unfolding holomorphic_on_def by (auto intro!: field_differentiable_ln_Gamma_complex)  | 
|
1500  | 
||
1501  | 
lemma analytic_ln_Gamma: "A \<inter> \<real>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> ln_Gamma analytic_on A"
 | 
|
1502  | 
by (rule analytic_on_subset[of _ "UNIV - \<real>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)  | 
|
1503  | 
(auto intro!: holomorphic_ln_Gamma)  | 
|
1504  | 
||
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1505  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1506  | 
lemma has_field_derivative_rGamma_complex' [derivative_intros]:  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1507  | 
"(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
 | 
1508  | 
-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
 | 
1509  | 
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
 | 
1510  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1511  | 
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
 | 
1512  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1513  | 
|
| 
62534
 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 
paulson <lp15@cam.ac.uk> 
parents: 
62533 
diff
changeset
 | 
1514  | 
lemma field_differentiable_Polygamma:  | 
| 64969 | 1515  | 
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
 | 
1516  | 
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
 | 
1517  | 
"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
 | 
1518  | 
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
 | 
1519  | 
|
| 64969 | 1520  | 
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
 | 
1521  | 
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
 | 
1522  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1523  | 
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
 | 
1524  | 
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
 | 
1525  | 
(auto intro!: holomorphic_on_Polygamma)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1526  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1527  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1528  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1529  | 
subsection%unimportant \<open>The real Gamma function\<close>  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1530  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1531  | 
lemma rGamma_series_real:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1532  | 
"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
 | 
1533  | 
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
 | 
1534  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1535  | 
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
 | 
1536  | 
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
 | 
1537  | 
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
 | 
1538  | 
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
 | 
1539  | 
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
 | 
1540  | 
(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
 | 
1541  | 
by (subst exp_of_real) simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1542  | 
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
 | 
1543  | 
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
 | 
1544  | 
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
 | 
1545  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1546  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1547  | 
instantiation%unimportant real :: Gamma  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1548  | 
begin  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1549  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1550  | 
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
 | 
1551  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1552  | 
instance proof  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1553  | 
fix x :: real  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1554  | 
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
 | 
1555  | 
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
 | 
1556  | 
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
 | 
1557  | 
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
 | 
1558  | 
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
 | 
1559  | 
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
 | 
1560  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1561  | 
fix x :: real assume "\<And>n. x \<noteq> - of_nat n"  | 
| 63539 | 1562  | 
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
 | 
1563  | 
by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases')  | 
| 63539 | 1564  | 
then have "x \<noteq> 0" by auto  | 
1565  | 
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
 | 
1566  | 
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
 | 
1567  | 
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
 | 
1568  | 
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
 | 
1569  | 
\<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
 | 
1570  | 
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
 | 
1571  | 
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
 | 
1572  | 
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
 | 
1573  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1574  | 
fix n :: nat  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1575  | 
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
 | 
1576  | 
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
 | 
1577  | 
simp: Polygamma_of_real rGamma_real_def [abs_def])  | 
| 64272 | 1578  | 
  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
 | 
1579  | 
(y - x)) /\<^sub>R norm (y - x)) \<midarrow>x::real\<rightarrow> 0"  | 
| 64272 | 1580  | 
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
 | 
1581  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1582  | 
fix x :: real  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1583  | 
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
 | 
1584  | 
proof (rule Lim_transform_eventually)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1585  | 
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
 | 
1586  | 
by (intro tendsto_intros)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1587  | 
qed (insert rGamma_series_real, simp add: eq_commute)  | 
| 64272 | 1588  | 
  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
 | 
1589  | 
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
 | 
1590  | 
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
 | 
1591  | 
in (\<lambda>n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma x"  | 
| 64272 | 1592  | 
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
 | 
1593  | 
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
 | 
1594  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1595  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1596  | 
end  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1597  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1598  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1599  | 
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
 | 
1600  | 
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
 | 
1601  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1602  | 
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
 | 
1603  | 
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
 | 
1604  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1605  | 
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
 | 
1606  | 
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
 | 
1607  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1608  | 
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
 | 
1609  | 
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
 | 
1610  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1611  | 
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
 | 
1612  | 
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
 | 
1613  | 
by (elim Reals_cases)  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1614  | 
(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
 | 
1615  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1616  | 
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
 | 
1617  | 
"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
 | 
1618  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1619  | 
assume xn: "x > 0" "n > 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1620  | 
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
 | 
1621  | 
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
 | 
1622  | 
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
 | 
1623  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1624  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1625  | 
lemma ln_Gamma_real_converges:  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1626  | 
assumes "(x::real) > 0"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1627  | 
shows "convergent (ln_Gamma_series x)"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1628  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1629  | 
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
 | 
1630  | 
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
 | 
1631  | 
moreover from eventually_gt_at_top[of "0::nat"]  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1632  | 
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
 | 
1633  | 
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
 | 
1634  | 
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
 | 
1635  | 
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
 | 
1636  | 
by (subst tendsto_cong) assumption+  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1637  | 
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
 | 
1638  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1639  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1640  | 
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
 | 
1641  | 
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
 | 
1642  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1643  | 
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
 | 
1644  | 
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
 | 
1645  | 
assume x: "x > 0"  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1646  | 
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
 | 
1647  | 
ln_Gamma_series (complex_of_real x) n) sequentially"  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1648  | 
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
 | 
1649  | 
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
 | 
1650  | 
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
 | 
1651  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1652  | 
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
 | 
1653  | 
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
 | 
1654  | 
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
 | 
1655  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1656  | 
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
 | 
1657  | 
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
 | 
1658  | 
|
| 64969 | 1659  | 
lemma ln_Gamma_complex_conv_fact: "n > 0 \<Longrightarrow> ln_Gamma (of_nat n :: complex) = ln (fact (n - 1))"  | 
1660  | 
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
 | 
1661  | 
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
 | 
1662  | 
|
| 64969 | 1663  | 
lemma ln_Gamma_real_conv_fact: "n > 0 \<Longrightarrow> ln_Gamma (real n) = ln (fact (n - 1))"  | 
1664  | 
using Gamma_fact[of "n - 1", where 'a = real]  | 
|
1665  | 
by (simp add: ln_Gamma_real_pos of_nat_diff Ln_of_real [symmetric])  | 
|
1666  | 
||
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
1667  | 
lemma Gamma_real_pos [simp, intro]: "x > (0::real) \<Longrightarrow> Gamma x > 0"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
1668  | 
by (simp add: Gamma_real_pos_exp)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
1669  | 
|
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
1670  | 
lemma Gamma_real_nonneg [simp, intro]: "x > (0::real) \<Longrightarrow> Gamma x \<ge> 0"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1671  | 
by (simp add: Gamma_real_pos_exp)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1672  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1673  | 
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
 | 
1674  | 
assumes x: "x > (0::real)"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1675  | 
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
 | 
1676  | 
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
 | 
1677  | 
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
 | 
1678  | 
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
 | 
1679  | 
simp: Polygamma_of_real o_def)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1680  | 
  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
 | 
1681  | 
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
 | 
1682  | 
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
 | 
1683  | 
qed  | 
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1684  | 
|
| 
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1685  | 
lemma field_differentiable_ln_Gamma_real:  | 
| 64969 | 1686  | 
"x > 0 \<Longrightarrow> ln_Gamma field_differentiable (at (x::real) within A)"  | 
1687  | 
by (rule field_differentiable_within_subset[of _ _ UNIV])  | 
|
1688  | 
(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
 | 
1689  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1690  | 
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
 | 
1691  | 
|
| 64969 | 1692  | 
lemma deriv_ln_Gamma_real:  | 
1693  | 
assumes "z > 0"  | 
|
1694  | 
shows "deriv ln_Gamma z = Digamma (z :: real)"  | 
|
1695  | 
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
 | 
1696  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1697  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1698  | 
lemma has_field_derivative_rGamma_real' [derivative_intros]:  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1699  | 
"(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
 | 
1700  | 
-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
 | 
1701  | 
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
 | 
1702  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1703  | 
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
 | 
1704  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1705  | 
lemma Polygamma_real_odd_pos:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1706  | 
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
 | 
1707  | 
shows "Polygamma n x > 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1708  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1709  | 
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
 | 
1710  | 
with assms show ?thesis  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1711  | 
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
 | 
1712  | 
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
 | 
1713  | 
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
 | 
1714  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1715  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1716  | 
lemma Polygamma_real_even_neg:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1717  | 
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
 | 
1718  | 
shows "Polygamma n x < 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1719  | 
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
 | 
1720  | 
by (auto intro!: mult_pos_pos suminf_pos)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1721  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1722  | 
lemma Polygamma_real_strict_mono:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1723  | 
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
 | 
1724  | 
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
 | 
1725  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1726  | 
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
 | 
1727  | 
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
 | 
1728  | 
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
 | 
1729  | 
note \<xi>(3)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1730  | 
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
 | 
1731  | 
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
 | 
1732  | 
finally show ?thesis by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1733  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1734  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1735  | 
lemma Polygamma_real_strict_antimono:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1736  | 
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
 | 
1737  | 
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
 | 
1738  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1739  | 
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
 | 
1740  | 
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
 | 
1741  | 
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
 | 
1742  | 
note \<xi>(3)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1743  | 
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
 | 
1744  | 
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
 | 
1745  | 
finally show ?thesis by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1746  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1747  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1748  | 
lemma Polygamma_real_mono:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1749  | 
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
 | 
1750  | 
shows "Polygamma n x \<le> Polygamma n y"  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1751  | 
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
 | 
1752  | 
by (cases "x = y") simp_all  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1753  | 
|
| 63721 | 1754  | 
lemma Digamma_real_strict_mono: "(0::real) < x \<Longrightarrow> x < y \<Longrightarrow> Digamma x < Digamma y"  | 
1755  | 
by (rule Polygamma_real_strict_mono) simp_all  | 
|
1756  | 
||
1757  | 
lemma Digamma_real_mono: "(0::real) < x \<Longrightarrow> x \<le> y \<Longrightarrow> Digamma x \<le> Digamma y"  | 
|
1758  | 
by (rule Polygamma_real_mono) simp_all  | 
|
1759  | 
||
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1760  | 
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
 | 
1761  | 
assumes "x \<ge> 3/2"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1762  | 
shows "Digamma (x :: real) > 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1763  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1764  | 
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
 | 
1765  | 
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
 | 
1766  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1767  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1768  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1769  | 
lemma ln_Gamma_real_strict_mono:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1770  | 
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
 | 
1771  | 
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
 | 
1772  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1773  | 
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
 | 
1774  | 
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
 | 
1775  | 
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
 | 
1776  | 
note \<xi>(3)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1777  | 
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
 | 
1778  | 
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
 | 
1779  | 
finally show ?thesis by simp  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1780  | 
qed  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1781  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1782  | 
lemma Gamma_real_strict_mono:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1783  | 
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
 | 
1784  | 
shows "Gamma (x :: real) < Gamma y"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1785  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1786  | 
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
 | 
1787  | 
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
 | 
1788  | 
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
 | 
1789  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1790  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1791  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1792  | 
theorem log_convex_Gamma_real: "convex_on {0<..} (ln \<circ> Gamma :: real \<Rightarrow> real)"
 | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1793  | 
by (rule convex_on_realI[of _ _ Digamma])  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1794  | 
(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
 | 
1795  | 
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
 | 
1796  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1797  | 
|
| 
63725
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1798  | 
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
 | 
1799  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1800  | 
text \<open>  | 
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1801  | 
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
 | 
1802  | 
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
 | 
1803  | 
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
 | 
1804  | 
Gamma function.  | 
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1805  | 
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
 | 
1806  | 
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
 | 
1807  | 
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
 | 
1808  | 
\<close>  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1809  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1810  | 
context%unimportant  | 
| 
63725
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1811  | 
fixes G :: "real \<Rightarrow> real"  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1812  | 
assumes G_1: "G 1 = 1"  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1813  | 
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
 | 
1814  | 
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
 | 
1815  | 
  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
 | 
1816  | 
begin  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1817  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1818  | 
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
 | 
1819  | 
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
 | 
1820  | 
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
 | 
1821  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1822  | 
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
 | 
1823  | 
"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
 | 
1824  | 
|
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
1825  | 
private lemma S_eq:  | 
| 
63725
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1826  | 
"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
 | 
1827  | 
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
 | 
1828  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1829  | 
private lemma G_lower:  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1830  | 
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
 | 
1831  | 
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
 | 
1832  | 
proof -  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1833  | 
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
 | 
1834  | 
(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
 | 
1835  | 
(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
 | 
1836  | 
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
 | 
1837  | 
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
 | 
1838  | 
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
 | 
1839  | 
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
 | 
1840  | 
unfolding S_def using n  | 
| 
63725
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1841  | 
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
 | 
1842  | 
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
 | 
1843  | 
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
 | 
1844  | 
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
 | 
1845  | 
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
 | 
1846  | 
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
 | 
1847  | 
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
 | 
1848  | 
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
 | 
1849  | 
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
 | 
1850  | 
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
 | 
1851  | 
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
 | 
1852  | 
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
 | 
1853  | 
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
 | 
1854  | 
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
 | 
1855  | 
qed  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1856  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1857  | 
private lemma G_upper:  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1858  | 
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
 | 
1859  | 
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
 | 
1860  | 
proof -  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1861  | 
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
 | 
1862  | 
(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
 | 
1863  | 
((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
 | 
1864  | 
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
 | 
1865  | 
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
 | 
1866  | 
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
 | 
1867  | 
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
 | 
1868  | 
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
 | 
1869  | 
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
 | 
1870  | 
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
 | 
1871  | 
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
 | 
1872  | 
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
 | 
1873  | 
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
 | 
1874  | 
by (simp add: ln_mult)  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1875  | 
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
 | 
1876  | 
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
 | 
1877  | 
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
 | 
1878  | 
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
 | 
1879  | 
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
 | 
1880  | 
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
 | 
1881  | 
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
 | 
1882  | 
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
 | 
1883  | 
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
 | 
1884  | 
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
 | 
1885  | 
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
 | 
1886  | 
finally show ?thesis .  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1887  | 
qed  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1888  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1889  | 
private lemma G_eq_Gamma_aux:  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1890  | 
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
 | 
1891  | 
shows "G x = Gamma x"  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1892  | 
proof (rule antisym)  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1893  | 
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
 | 
1894  | 
proof (rule tendsto_upperbound)  | 
| 
63725
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1895  | 
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
 | 
1896  | 
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
 | 
1897  | 
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
 | 
1898  | 
next  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1899  | 
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
 | 
1900  | 
proof (rule tendsto_lowerbound)  | 
| 
63725
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1901  | 
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
 | 
1902  | 
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
 | 
1903  | 
Gamma_series_LIMSEQ filterlim_real_sequentially)+  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1904  | 
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
 | 
1905  | 
next  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1906  | 
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
 | 
1907  | 
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
 | 
1908  | 
qed simp_all  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1909  | 
qed  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1910  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1911  | 
theorem Gamma_pos_real_unique:  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1912  | 
assumes x: "x > 0"  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1913  | 
shows "G x = Gamma x"  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1914  | 
proof -  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1915  | 
  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
 | 
1916  | 
proof (induction n)  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1917  | 
case (Suc n)  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1918  | 
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
 | 
1919  | 
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
 | 
1920  | 
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
 | 
1921  | 
by (auto simp: add_ac)  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1922  | 
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
 | 
1923  | 
|
| 
63725
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1924  | 
show ?thesis  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1925  | 
proof (cases "frac x = 0")  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1926  | 
case True  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1927  | 
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
 | 
1928  | 
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
 | 
1929  | 
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
 | 
1930  | 
next  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1931  | 
case False  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1932  | 
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
 | 
1933  | 
by (simp add: frac_def)  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1934  | 
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
 | 
1935  | 
show ?thesis  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1936  | 
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
 | 
1937  | 
qed  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1938  | 
qed  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1939  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1940  | 
end  | 
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1941  | 
|
| 
 
4c00ba1ad11a
Bohr-Mollerup theorem for the Gamma function
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
63721 
diff
changeset
 | 
1942  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1943  | 
subsection \<open>The Beta function\<close>  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1944  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1945  | 
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
 | 
1946  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1947  | 
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
 | 
1948  | 
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
 | 
1949  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1950  | 
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
 | 
1951  | 
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
 | 
1952  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1953  | 
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
 | 
1954  | 
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
 | 
1955  | 
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
 | 
1956  | 
(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
 | 
1957  | 
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
 | 
1958  | 
|
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
1959  | 
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
 | 
1960  | 
by (auto simp add: Beta_def elim!: nonpos_Ints_cases')  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
1961  | 
|
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
1962  | 
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
 | 
1963  | 
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
 | 
1964  | 
|
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
1965  | 
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
 | 
1966  | 
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
 | 
1967  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1968  | 
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
 | 
1969  | 
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
 | 
1970  | 
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
 | 
1971  | 
(at y within A)"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1972  | 
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
 | 
1973  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1974  | 
theorem Beta_plus1_plus1:  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1975  | 
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
 | 
1976  | 
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
 | 
1977  | 
proof -  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
1978  | 
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
 | 
1979  | 
(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
 | 
1980  | 
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
 | 
1981  | 
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
 | 
1982  | 
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
 | 
1983  | 
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
 | 
1984  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1985  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1986  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1987  | 
theorem Beta_plus1_left:  | 
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
1988  | 
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
 | 
1989  | 
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
 | 
1990  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1991  | 
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
 | 
1992  | 
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
 | 
1993  | 
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
 | 
1994  | 
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
 | 
1995  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1996  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
1997  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
1998  | 
theorem Beta_plus1_right:  | 
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
1999  | 
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
 | 
2000  | 
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
 | 
2001  | 
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
 | 
2002  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2003  | 
lemma Gamma_Gamma_Beta:  | 
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2004  | 
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
 | 
2005  | 
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
 | 
2006  | 
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
 | 
2007  | 
by (simp add: rGamma_inverse_Gamma)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2008  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2009  | 
|
| 
 
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  | 
subsection \<open>Legendre duplication theorem\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2012  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2013  | 
context  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2014  | 
begin  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2015  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2016  | 
private lemma Gamma_legendre_duplication_aux:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2017  | 
fixes z :: "'a :: Gamma"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2018  | 
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
 | 
2019  | 
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
 | 
2020  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2021  | 
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
 | 
2022  | 
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
 | 
2023  | 
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
 | 
2024  | 
  {
 | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2025  | 
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
 | 
2026  | 
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
 | 
2027  | 
Gamma_series' (2*z) (2*n)"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2028  | 
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
 | 
2029  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2030  | 
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
 | 
2031  | 
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
 | 
2032  | 
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
 | 
2033  | 
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
 | 
2034  | 
(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
 | 
2035  | 
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
 | 
2036  | 
have B: "Gamma_series' (2*z) (2*n) =  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2037  | 
?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
 | 
2038  | 
(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
 | 
2039  | 
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
 | 
2040  | 
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
 | 
2041  | 
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
 | 
2042  | 
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
 | 
2043  | 
?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
 | 
2044  | 
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
 | 
2045  | 
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
 | 
2046  | 
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
 | 
2047  | 
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
 | 
2048  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2049  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2050  | 
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
 | 
2051  | 
hence "?g \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"  | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
68721 
diff
changeset
 | 
2052  | 
using LIMSEQ_subseq_LIMSEQ[OF Gamma_series'_LIMSEQ, of "(*)2" "2*z"]  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2053  | 
by (intro tendsto_intros Gamma_series'_LIMSEQ)  | 
| 
66447
 
a1f5c5c26fa6
Replaced subseq with strict_mono
 
eberlm <eberlm@in.tum.de> 
parents: 
66286 
diff
changeset
 | 
2054  | 
(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
 | 
2055  | 
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
 | 
2056  | 
by (rule Lim_transform_eventually)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2057  | 
} note lim = this  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2058  | 
|
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2059  | 
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
 | 
2060  | 
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
 | 
2061  | 
by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2062  | 
with lim[of "1/2 :: 'a"] have "?h \<longlonglongrightarrow> 2 * Gamma (1/2 :: 'a)" by (simp add: exp_of_real)  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2063  | 
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
 | 
2064  | 
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
 | 
2065  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2066  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2067  | 
theorem Gamma_reflection_complex:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2068  | 
fixes z :: complex  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2069  | 
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
 | 
2070  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2071  | 
let ?g = "\<lambda>z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)"  | 
| 63040 | 2072  | 
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
 | 
2073  | 
let ?h = "\<lambda>z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"  | 
| 63040 | 2074  | 
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
 | 
2075  | 
|
| 62072 | 2076  | 
  \<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
 | 
2077  | 
interpret g: periodic_fun_simple' g  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2078  | 
proof  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2079  | 
fix z :: complex  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2080  | 
show "g (z + 1) = g z"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2081  | 
proof (cases "z \<in> \<int>")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2082  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2083  | 
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
 | 
2084  | 
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
 | 
2085  | 
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
 | 
2086  | 
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
 | 
2087  | 
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
 | 
2088  | 
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
 | 
2089  | 
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
 | 
2090  | 
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
 | 
2091  | 
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
 | 
2092  | 
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
 | 
2093  | 
qed (simp add: g_def)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2094  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2095  | 
|
| 62072 | 2096  | 
  \<comment> \<open>@{term g} is entire.\<close>
 | 
| 64969 | 2097  | 
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
 | 
2098  | 
proof (cases "z \<in> \<int>")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2099  | 
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
 | 
2100  | 
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
 | 
2101  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2102  | 
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
 | 
2103  | 
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
 | 
2104  | 
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
 | 
2105  | 
    moreover {
 | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2106  | 
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
 | 
2107  | 
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
 | 
2108  | 
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
 | 
2109  | 
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
 | 
2110  | 
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
 | 
2111  | 
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
 | 
2112  | 
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
 | 
2113  | 
}  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2114  | 
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
 | 
2115  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2116  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2117  | 
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
 | 
2118  | 
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
 | 
2119  | 
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
 | 
2120  | 
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
 | 
2121  | 
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
 | 
2122  | 
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
 | 
2123  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2124  | 
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
 | 
2125  | 
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
 | 
2126  | 
proof (cases "z = 0")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2127  | 
assume z': "z \<noteq> 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2128  | 
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
 | 
2129  | 
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
 | 
2130  | 
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
 | 
2131  | 
qed (simp add: g_def)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2132  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2133  | 
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
 | 
2134  | 
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
 | 
2135  | 
by (intro DERIV_chain) simp_all  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2136  | 
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
 | 
2137  | 
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
 | 
2138  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2139  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2140  | 
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
 | 
2141  | 
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
 | 
2142  | 
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
 | 
2143  | 
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
 | 
2144  | 
qed  | 
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
2145  | 
|
| 64969 | 2146  | 
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
 | 
2147  | 
by (rule holomorphic_on_subset[of _ UNIV])  | 
| 64969 | 2148  | 
(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
 | 
2149  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2150  | 
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
 | 
2151  | 
proof (cases "z \<in> \<int>")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2152  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2153  | 
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
 | 
2154  | 
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
 | 
2155  | 
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
 | 
2156  | 
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
 | 
2157  | 
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
 | 
2158  | 
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
 | 
2159  | 
ultimately show ?thesis by force  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2160  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2161  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2162  | 
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
 | 
2163  | 
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
 | 
2164  | 
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
 | 
2165  | 
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
 | 
2166  | 
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
 | 
2167  | 
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
 | 
2168  | 
(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
 | 
2169  | 
(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
 | 
2170  | 
by (simp add: g_def)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2171  | 
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
 | 
2172  | 
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
 | 
2173  | 
by (simp add: add_divide_distrib)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2174  | 
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
 | 
2175  | 
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
 | 
2176  | 
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
 | 
2177  | 
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
 | 
2178  | 
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
 | 
2179  | 
(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
 | 
2180  | 
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
 | 
2181  | 
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
 | 
2182  | 
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
 | 
2183  | 
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
 | 
2184  | 
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
 | 
2185  | 
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
 | 
2186  | 
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
 | 
2187  | 
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
 | 
2188  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2189  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2190  | 
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
 | 
2191  | 
proof -  | 
| 63040 | 2192  | 
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
 | 
2193  | 
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
 | 
2194  | 
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
 | 
2195  | 
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
 | 
2196  | 
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
 | 
2197  | 
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
 | 
2198  | 
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
 | 
2199  | 
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
 | 
2200  | 
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
 | 
2201  | 
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
 | 
2202  | 
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
 | 
2203  | 
finally show ?thesis ..  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2204  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2205  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2206  | 
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
 | 
2207  | 
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
 | 
2208  | 
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
 | 
2209  | 
|
| 64969 | 2210  | 
have h_altdef: "h z = deriv g z / g z" for z :: complex  | 
2211  | 
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
 | 
2212  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2213  | 
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
 | 
2214  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2215  | 
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
 | 
2216  | 
(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
 | 
2217  | 
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
 | 
2218  | 
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
 | 
2219  | 
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
 | 
2220  | 
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
 | 
2221  | 
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
 | 
2222  | 
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
 | 
2223  | 
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
 | 
2224  | 
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
 | 
2225  | 
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
 | 
2226  | 
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
 | 
2227  | 
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
 | 
2228  | 
by (intro DERIV_unique)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2229  | 
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
 | 
2230  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2231  | 
|
| 64969 | 2232  | 
have h_holo [holomorphic_intros]: "h holomorphic_on A" for A  | 
2233  | 
unfolding h_altdef [abs_def]  | 
|
2234  | 
by (rule holomorphic_on_subset[of _ UNIV]) (auto intro!: holomorphic_intros)  | 
|
2235  | 
define h' where "h' = deriv h"  | 
|
2236  | 
have h_h': "(h has_field_derivative h' z) (at z)" for z unfolding h'_def  | 
|
2237  | 
by (auto intro!: holomorphic_derivI[of _ UNIV] holomorphic_intros)  | 
|
2238  | 
have h'_holo [holomorphic_intros]: "h' holomorphic_on A" for A unfolding h'_def  | 
|
2239  | 
by (rule holomorphic_on_subset[of _ UNIV]) (auto intro!: holomorphic_intros)  | 
|
2240  | 
have h'_cont: "continuous_on UNIV h'"  | 
|
2241  | 
by (intro holomorphic_on_imp_continuous_on holomorphic_intros)  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2242  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2243  | 
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
 | 
2244  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2245  | 
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
 | 
2246  | 
((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
 | 
2247  | 
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
 | 
2248  | 
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
 | 
2249  | 
by (subst (asm) h_eq[symmetric])  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2250  | 
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
 | 
2251  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2252  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2253  | 
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
 | 
2254  | 
proof -  | 
| 63040 | 2255  | 
define m where "m = max 1 \<bar>Re z\<bar>"  | 
2256  | 
    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
 | 
2257  | 
    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
 | 
2258  | 
                  {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
 | 
2259  | 
(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
 | 
2260  | 
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
 | 
2261  | 
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
 | 
2262  | 
finally have "closed B" .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2263  | 
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
 | 
2264  | 
proof (intro ballI exI)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2265  | 
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
 | 
2266  | 
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
 | 
2267  | 
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
 | 
2268  | 
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
 | 
2269  | 
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
 | 
2270  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2271  | 
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
 | 
2272  | 
|
| 
69260
 
0a9688695a1b
removed relics of ASCII syntax for indexed big operators
 
haftmann 
parents: 
69064 
diff
changeset
 | 
2273  | 
define M where "M = (SUP z\<in>B. norm (h' z))"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2274  | 
have "compact (h' ` B)"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2275  | 
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
 | 
2276  | 
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
 | 
2277  | 
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
 | 
2278  | 
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
 | 
2279  | 
also have "M \<le> M/2"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2280  | 
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
 | 
2281  | 
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
 | 
2282  | 
      thus "B \<noteq> {}" by auto
 | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2283  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2284  | 
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
 | 
2285  | 
proof  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2286  | 
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
 | 
2287  | 
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
 | 
2288  | 
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
 | 
2289  | 
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
 | 
2290  | 
by (rule norm_triangle_ineq)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2291  | 
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
 | 
2292  | 
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
 | 
2293  | 
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
 | 
2294  | 
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
 | 
2295  | 
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
 | 
2296  | 
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
 | 
2297  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2298  | 
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
 | 
2299  | 
hence "M \<le> 0" by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2300  | 
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
 | 
2301  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2302  | 
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
 | 
2303  | 
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
 | 
2304  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2305  | 
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
 | 
2306  | 
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
 | 
2307  | 
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
 | 
2308  | 
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
 | 
2309  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2310  | 
from h'_zero h_h'_2 have "\<exists>c. \<forall>z\<in>UNIV. h z = c"  | 
| 66936 | 2311  | 
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
 | 
2312  | 
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
 | 
2313  | 
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
 | 
2314  | 
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
 | 
2315  | 
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
 | 
2316  | 
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
 | 
2317  | 
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
 | 
2318  | 
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
 | 
2319  | 
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
 | 
2320  | 
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
 | 
2321  | 
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
 | 
2322  | 
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
 | 
2323  | 
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
 | 
2324  | 
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
 | 
2325  | 
then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm)  | 
| 63539 | 2326  | 
from this[of 0] have "c' = pi" unfolding g_def by simp  | 
2327  | 
with c have "g z = pi" by simp  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2328  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2329  | 
show ?thesis  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2330  | 
proof (cases "z \<in> \<int>")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2331  | 
case False  | 
| 62072 | 2332  | 
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
 | 
2333  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2334  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2335  | 
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
 | 
2336  | 
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
 | 
2337  | 
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
 | 
2338  | 
ultimately show ?thesis using n  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2339  | 
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
 | 
2340  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2341  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2342  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2343  | 
lemma rGamma_reflection_complex:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2344  | 
"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
 | 
2345  | 
using Gamma_reflection_complex[of z]  | 
| 62390 | 2346  | 
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
 | 
2347  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2348  | 
lemma rGamma_reflection_complex':  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2349  | 
"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
 | 
2350  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2351  | 
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
 | 
2352  | 
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
 | 
2353  | 
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
 | 
2354  | 
by (rule rGamma_reflection_complex)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2355  | 
finally show ?thesis by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2356  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2357  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2358  | 
lemma Gamma_reflection_complex':  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2359  | 
"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
 | 
2360  | 
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
 | 
2361  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2362  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2363  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2364  | 
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
 | 
2365  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2366  | 
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
 | 
2367  | 
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
 | 
2368  | 
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
 | 
2369  | 
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
 | 
2370  | 
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
 | 
2371  | 
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
 | 
2372  | 
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
 | 
2373  | 
qed  | 
| 
 
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  | 
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
 | 
2376  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2377  | 
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
 | 
2378  | 
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
 | 
2379  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2380  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2381  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2382  | 
theorem Gamma_legendre_duplication:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2383  | 
fixes z :: complex  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2384  | 
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
 | 
2385  | 
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
 | 
2386  | 
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
 | 
2387  | 
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
 | 
2388  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2389  | 
end  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2390  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2391  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2392  | 
subsection%unimportant \<open>Limits and residues\<close>  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2393  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2394  | 
text \<open>  | 
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2395  | 
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
 | 
2396  | 
\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2397  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2398  | 
lemma rGamma_zeros:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2399  | 
"(\<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
 | 
2400  | 
proof (subst tendsto_cong)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2401  | 
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
 | 
2402  | 
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
 | 
2403  | 
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
 | 
2404  | 
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
 | 
2405  | 
(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
 | 
2406  | 
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
 | 
2407  | 
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
 | 
2408  | 
by (simp add: pochhammer_same)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2409  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2410  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2411  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2412  | 
text \<open>  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2413  | 
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
 | 
2414  | 
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
 | 
2415  | 
\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2416  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2417  | 
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
 | 
2418  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2419  | 
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
 | 
2420  | 
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
 | 
2421  | 
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
 | 
2422  | 
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
 | 
2423  | 
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
 | 
2424  | 
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
 | 
2425  | 
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
 | 
2426  | 
(simp_all add: filterlim_at)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2427  | 
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
 | 
2428  | 
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
 | 
2429  | 
ultimately show ?thesis by (simp only: )  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2430  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2431  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2432  | 
lemma Gamma_residues:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2433  | 
"(\<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
 | 
2434  | 
proof (subst tendsto_cong)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2435  | 
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
 | 
2436  | 
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
 | 
2437  | 
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
 | 
2438  | 
(at (- of_nat n))"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2439  | 
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
 | 
2440  | 
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
 | 
2441  | 
inverse ((- 1) ^ n * fact n :: 'a)"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2442  | 
by (intro tendsto_intros rGamma_zeros) simp_all  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2443  | 
also have "inverse ((- 1) ^ n * fact n) = ?c"  | 
| 68403 | 2444  | 
by (simp_all add: field_simps flip: power_mult_distrib)  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2445  | 
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
 | 
2446  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2447  | 
|
| 
66286
 
1c977b13414f
poles and residues of the Gamma function
 
eberlm <eberlm@in.tum.de> 
parents: 
65587 
diff
changeset
 | 
2448  | 
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
 | 
2449  | 
unfolding is_pole_def using Gamma_poles .  | 
| 
 
1c977b13414f
poles and residues of the Gamma function
 
eberlm <eberlm@in.tum.de> 
parents: 
65587 
diff
changeset
 | 
2450  | 
|
| 
 
1c977b13414f
poles and residues of the Gamma function
 
eberlm <eberlm@in.tum.de> 
parents: 
65587 
diff
changeset
 | 
2451  | 
lemma Gamme_residue:  | 
| 
 
1c977b13414f
poles and residues of the Gamma function
 
eberlm <eberlm@in.tum.de> 
parents: 
65587 
diff
changeset
 | 
2452  | 
"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
 | 
2453  | 
proof (rule residue_simple')  | 
| 
 
1c977b13414f
poles and residues of the Gamma function
 
eberlm <eberlm@in.tum.de> 
parents: 
65587 
diff
changeset
 | 
2454  | 
  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
 | 
2455  | 
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
 | 
2456  | 
  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
 | 
2457  | 
by (rule holomorphic_Gamma) auto  | 
| 
 
1c977b13414f
poles and residues of the Gamma function
 
eberlm <eberlm@in.tum.de> 
parents: 
65587 
diff
changeset
 | 
2458  | 
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
 | 
2459  | 
using Gamma_residues[of n] by simp  | 
| 
 
1c977b13414f
poles and residues of the Gamma function
 
eberlm <eberlm@in.tum.de> 
parents: 
65587 
diff
changeset
 | 
2460  | 
qed auto  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2461  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2462  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2463  | 
subsection \<open>Alternative definitions\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2464  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2465  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2466  | 
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
 | 
2467  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2468  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2469  | 
definition Gamma_series_euler' where  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2470  | 
"Gamma_series_euler' z n =  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2471  | 
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
 | 
2472  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2473  | 
context  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2474  | 
begin  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2475  | 
private lemma Gamma_euler'_aux1:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2476  | 
  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
 | 
2477  | 
assumes n: "n > 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2478  | 
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
 | 
2479  | 
proof -  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2480  | 
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
 | 
2481  | 
exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))"  | 
| 64267 | 2482  | 
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
 | 
2483  | 
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 | 2484  | 
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
 | 
2485  | 
also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)"  | 
| 64272 | 2486  | 
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
 | 
2487  | 
also have "(\<Prod>k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1"  | 
| 64272 | 2488  | 
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
 | 
2489  | 
finally show ?thesis ..  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2490  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2491  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2492  | 
theorem Gamma_series_euler':  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2493  | 
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
 | 
2494  | 
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
 | 
2495  | 
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
 | 
2496  | 
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
 | 
2497  | 
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
 | 
2498  | 
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
 | 
2499  | 
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
 | 
2500  | 
|
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
2501  | 
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
 | 
2502  | 
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
 | 
2503  | 
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
 | 
2504  | 
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
 | 
2505  | 
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
 | 
2506  | 
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
 | 
2507  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2508  | 
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
 | 
2509  | 
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
 | 
2510  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2511  | 
fix n :: nat assume n: "n > 0"  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2512  | 
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
 | 
2513  | 
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
 | 
2514  | 
by (subst Gamma_euler'_aux1)  | 
| 64272 | 2515  | 
(simp_all add: Gamma_series_euler'_def prod.distrib  | 
2516  | 
prod_inversef[symmetric] divide_inverse)  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2517  | 
also have "(\<Prod>k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n"  | 
| 64272 | 2518  | 
by (cases n) (simp_all add: pochhammer_prod fact_prod atLeastLessThanSuc_atLeastAtMost  | 
2519  | 
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
 | 
2520  | 
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
 | 
2521  | 
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
 | 
2522  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2523  | 
qed  | 
| 
 
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  | 
end  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2526  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2527  | 
|
| 
 
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  | 
subsubsection \<open>Weierstrass form\<close>  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2530  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2531  | 
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
 | 
2532  | 
"Gamma_series_weierstrass z n =  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2533  | 
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
 | 
2534  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2535  | 
definition%unimportant  | 
| 
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2536  | 
  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
 | 
2537  | 
"rGamma_series_weierstrass z n =  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2538  | 
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
 | 
2539  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2540  | 
lemma Gamma_series_weierstrass_nonpos_Ints:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2541  | 
"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
 | 
2542  | 
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
 | 
2543  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2544  | 
lemma rGamma_series_weierstrass_nonpos_Ints:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2545  | 
"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
 | 
2546  | 
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
 | 
2547  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2548  | 
theorem Gamma_weierstrass_complex: "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma (z :: complex)"  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2549  | 
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
 | 
2550  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2551  | 
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
 | 
2552  | 
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
 | 
2553  | 
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
 | 
2554  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2555  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2556  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2557  | 
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
 | 
2558  | 
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
 | 
2559  | 
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
 | 
2560  | 
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
 | 
2561  | 
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
 | 
2562  | 
using ln_Gamma_series'_aux[OF False]  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2563  | 
by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def  | 
| 64267 | 2564  | 
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
 | 
2565  | 
from tendsto_exp[OF this] False z have "?f \<longlonglongrightarrow> z * exp (euler_mascheroni * z) * Gamma z"  | 
| 64267 | 2566  | 
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
 | 
2567  | 
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
 | 
2568  | 
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
 | 
2569  | 
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
 | 
2570  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2571  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2572  | 
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
 | 
2573  | 
by (rule tendsto_of_real_iff)  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2574  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2575  | 
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
 | 
2576  | 
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
 | 
2577  | 
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
 | 
2578  | 
(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
 | 
2579  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2580  | 
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
 | 
2581  | 
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
 | 
2582  | 
case True  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2583  | 
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
 | 
2584  | 
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
 | 
2585  | 
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
 | 
2586  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2587  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2588  | 
case False  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2589  | 
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
 | 
2590  | 
by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def  | 
| 64272 | 2591  | 
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
 | 
2592  | 
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
 | 
2593  | 
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
 | 
2594  | 
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
 | 
2595  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
2596  | 
|
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2597  | 
subsubsection \<open>Binomial coefficient form\<close>  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2598  | 
|
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2599  | 
lemma Gamma_gbinomial:  | 
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2600  | 
"(\<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
 | 
2601  | 
proof (cases "z = 0")  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2602  | 
case False  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2603  | 
show ?thesis  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2604  | 
proof (rule Lim_transform_eventually)  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2605  | 
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
 | 
2606  | 
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
 | 
2607  | 
((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
 | 
2608  | 
proof (intro always_eventually allI)  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2609  | 
fix n :: nat  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2610  | 
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
 | 
2611  | 
by (simp add: gbinomial_pochhammer' pochhammer_rec)  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2612  | 
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
 | 
2613  | 
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
 | 
2614  | 
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
 | 
2615  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2616  | 
|
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2617  | 
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
 | 
2618  | 
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
 | 
2619  | 
by (simp add: field_simps)  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2620  | 
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
 | 
2621  | 
qed  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2622  | 
qed (simp_all add: binomial_gbinomial [symmetric])  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2623  | 
|
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2624  | 
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
 | 
2625  | 
by (subst gbinomial_minus) (simp add: power_mult_distrib [symmetric])  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2626  | 
|
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2627  | 
lemma gbinomial_asymptotic:  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2628  | 
fixes z :: "'a :: Gamma"  | 
| 
63594
 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 
hoelzl 
parents: 
63539 
diff
changeset
 | 
2629  | 
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
 | 
2630  | 
inverse (Gamma (- z))"  | 
| 
63594
 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 
hoelzl 
parents: 
63539 
diff
changeset
 | 
2631  | 
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
 | 
2632  | 
by (subst (asm) gbinomial_minus')  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2633  | 
(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
 | 
2634  | 
|
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
2635  | 
lemma fact_binomial_limit:  | 
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2636  | 
"(\<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
 | 
2637  | 
proof (rule Lim_transform_eventually)  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2638  | 
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
 | 
2639  | 
\<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
 | 
2640  | 
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
 | 
2641  | 
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
 | 
2642  | 
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
 | 
2643  | 
finally show "?f \<longlonglongrightarrow> 1 / fact k" .  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2644  | 
|
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2645  | 
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
 | 
2646  | 
using eventually_gt_at_top[of "0::nat"]  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2647  | 
proof eventually_elim  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2648  | 
fix n :: nat assume n: "n > 0"  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2649  | 
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
 | 
2650  | 
by (simp add: exp_of_nat_mult)  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2651  | 
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
 | 
2652  | 
qed  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2653  | 
qed  | 
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2654  | 
|
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2655  | 
lemma binomial_asymptotic':  | 
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2656  | 
"(\<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
 | 
2657  | 
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
 | 
2658  | 
|
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2659  | 
lemma gbinomial_Beta:  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2660  | 
assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2661  | 
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
 | 
2662  | 
using assms  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2663  | 
proof (induction n arbitrary: z)  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2664  | 
case 0  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2665  | 
hence "z + 2 \<notin> \<int>\<^sub>\<le>\<^sub>0"  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2666  | 
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
 | 
2667  | 
with 0 show ?case  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2668  | 
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
 | 
2669  | 
next  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2670  | 
case (Suc n z)  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2671  | 
show ?case  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2672  | 
proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2673  | 
case True  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2674  | 
with Suc.prems have "z = 0"  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2675  | 
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
 | 
2676  | 
show ?thesis  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2677  | 
proof (cases "n = 0")  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2678  | 
case True  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2679  | 
with \<open>z = 0\<close> show ?thesis  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2680  | 
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
 | 
2681  | 
next  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2682  | 
case False  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2683  | 
with \<open>z = 0\<close> show ?thesis  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2684  | 
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
 | 
2685  | 
qed  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2686  | 
next  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2687  | 
case False  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2688  | 
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
 | 
2689  | 
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
 | 
2690  | 
by (subst gbinomial_factors) (simp add: field_simps)  | 
| 
63594
 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 
hoelzl 
parents: 
63539 
diff
changeset
 | 
2691  | 
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
 | 
2692  | 
(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
 | 
2693  | 
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
 | 
2694  | 
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
 | 
2695  | 
by (subst Beta_plus1_right [symmetric]) simp_all  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2696  | 
finally show ?thesis .  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2697  | 
qed  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2698  | 
qed  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2699  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2700  | 
theorem gbinomial_Gamma:  | 
| 
63295
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2701  | 
assumes "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0"  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2702  | 
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
 | 
2703  | 
proof -  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2704  | 
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
 | 
2705  | 
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
 | 
2706  | 
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
 | 
2707  | 
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
 | 
2708  | 
finally show ?thesis .  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2709  | 
qed  | 
| 
 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 
eberlm 
parents: 
63040 
diff
changeset
 | 
2710  | 
|
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
2711  | 
|
| 63296 | 2712  | 
subsubsection \<open>Integral form\<close>  | 
2713  | 
||
| 66526 | 2714  | 
lemma integrable_on_powr_from_0':  | 
2715  | 
assumes a: "a > (-1::real)" and c: "c \<ge> 0"  | 
|
2716  | 
  shows   "(\<lambda>x. x powr a) integrable_on {0<..c}"
 | 
|
2717  | 
proof -  | 
|
2718  | 
  from c have *: "{0<..c} - {0..c} = {}" "{0..c} - {0<..c} = {0}" by auto
 | 
|
2719  | 
show ?thesis  | 
|
2720  | 
by (rule integrable_spike_set [OF integrable_on_powr_from_0[OF a c]]) (simp_all add: *)  | 
|
2721  | 
qed  | 
|
2722  | 
||
2723  | 
lemma absolutely_integrable_Gamma_integral:  | 
|
2724  | 
assumes "Re z > 0" "a > 0"  | 
|
2725  | 
shows "(\<lambda>t. complex_of_real t powr (z - 1) / of_real (exp (a * t)))  | 
|
2726  | 
             absolutely_integrable_on {0<..}" (is "?f absolutely_integrable_on _")
 | 
|
2727  | 
proof -  | 
|
2728  | 
have "((\<lambda>x. (Re z - 1) * (ln x / x)) \<longlongrightarrow> (Re z - 1) * 0) at_top"  | 
|
2729  | 
by (intro tendsto_intros ln_x_over_x_tendsto_0)  | 
|
2730  | 
hence "((\<lambda>x. ((Re z - 1) * ln x) / x) \<longlongrightarrow> 0) at_top" by simp  | 
|
2731  | 
from order_tendstoD(2)[OF this, of "a/2"] and \<open>a > 0\<close>  | 
|
2732  | 
have "eventually (\<lambda>x. (Re z - 1) * ln x / x < a/2) at_top" by simp  | 
|
2733  | 
from eventually_conj[OF this eventually_gt_at_top[of 0]]  | 
|
2734  | 
obtain x0 where "\<forall>x\<ge>x0. (Re z - 1) * ln x / x < a/2 \<and> x > 0"  | 
|
2735  | 
by (auto simp: eventually_at_top_linorder)  | 
|
2736  | 
hence "x0 > 0" by simp  | 
|
2737  | 
have "x powr (Re z - 1) / exp (a * x) < exp (-(a/2) * x)" if "x \<ge> x0" for x  | 
|
2738  | 
proof -  | 
|
2739  | 
from that and \<open>\<forall>x\<ge>x0. _\<close> have x: "(Re z - 1) * ln x / x < a / 2" "x > 0" by auto  | 
|
2740  | 
have "x powr (Re z - 1) = exp ((Re z - 1) * ln x)"  | 
|
2741  | 
using \<open>x > 0\<close> by (simp add: powr_def)  | 
|
2742  | 
also from x have "(Re z - 1) * ln x < (a * x) / 2" by (simp add: field_simps)  | 
|
2743  | 
finally show ?thesis by (simp add: field_simps exp_add [symmetric])  | 
|
2744  | 
qed  | 
|
2745  | 
note x0 = \<open>x0 > 0\<close> this  | 
|
2746  | 
||
2747  | 
  have "?f absolutely_integrable_on ({0<..x0} \<union> {x0..})"
 | 
|
2748  | 
proof (rule set_integrable_Un)  | 
|
2749  | 
    show "?f absolutely_integrable_on {0<..x0}"
 | 
|
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2750  | 
unfolding set_integrable_def  | 
| 66526 | 2751  | 
proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2])  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2752  | 
      show "integrable lebesgue (\<lambda>x. indicat_real {0<..x0} x *\<^sub>R x powr (Re z - 1))"         
 | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2753  | 
using x0(1) assms  | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2754  | 
by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_powr_from_0') auto  | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2755  | 
      show "(\<lambda>x. indicat_real {0<..x0} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \<in> borel_measurable lebesgue"
 | 
| 66526 | 2756  | 
by (intro measurable_completion)  | 
2757  | 
(auto intro!: borel_measurable_continuous_on_indicator continuous_intros)  | 
|
2758  | 
fix x :: real  | 
|
2759  | 
have "x powr (Re z - 1) / exp (a * x) \<le> x powr (Re z - 1) / 1" if "x \<ge> 0"  | 
|
2760  | 
using that assms by (intro divide_left_mono) auto  | 
|
2761  | 
      thus "norm (indicator {0<..x0} x *\<^sub>R ?f x) \<le> 
 | 
|
2762  | 
               norm (indicator {0<..x0} x *\<^sub>R x powr (Re z - 1))"
 | 
|
2763  | 
by (simp_all add: norm_divide norm_powr_real_powr indicator_def)  | 
|
2764  | 
qed  | 
|
2765  | 
next  | 
|
2766  | 
    show "?f absolutely_integrable_on {x0..}"
 | 
|
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2767  | 
unfolding set_integrable_def  | 
| 66526 | 2768  | 
proof (rule Bochner_Integration.integrable_bound [OF _ _ AE_I2])  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2769  | 
      show "integrable lebesgue (\<lambda>x. indicat_real {x0..} x *\<^sub>R exp (- (a / 2) * x))" using assms
 | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2770  | 
by (intro nonnegative_absolutely_integrable_1 [unfolded set_integrable_def] integrable_on_exp_minus_to_infinity) auto  | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
2771  | 
      show "(\<lambda>x. indicat_real {x0..} x *\<^sub>R (x powr (z - 1) / exp (a * x))) \<in> borel_measurable lebesgue" using x0(1)
 | 
| 66526 | 2772  | 
by (intro measurable_completion)  | 
2773  | 
(auto intro!: borel_measurable_continuous_on_indicator continuous_intros)  | 
|
2774  | 
fix x :: real  | 
|
2775  | 
      show "norm (indicator {x0..} x *\<^sub>R ?f x) \<le> 
 | 
|
2776  | 
               norm (indicator {x0..} x *\<^sub>R exp (-(a/2) * x))" using x0
 | 
|
2777  | 
by (auto simp: norm_divide norm_powr_real_powr indicator_def less_imp_le)  | 
|
2778  | 
qed  | 
|
2779  | 
qed auto  | 
|
2780  | 
  also have "{0<..x0} \<union> {x0..} = {0<..}" using x0(1) by auto
 | 
|
2781  | 
finally show ?thesis .  | 
|
2782  | 
qed  | 
|
2783  | 
||
| 63296 | 2784  | 
lemma integrable_Gamma_integral_bound:  | 
2785  | 
fixes a c :: real  | 
|
2786  | 
assumes a: "a > -1" and c: "c \<ge> 0"  | 
|
2787  | 
  defines "f \<equiv> \<lambda>x. if x \<in> {0..c} then x powr a else exp (-x/2)"
 | 
|
2788  | 
  shows   "f integrable_on {0..}"
 | 
|
2789  | 
proof -  | 
|
2790  | 
  have "f integrable_on {0..c}"
 | 
|
2791  | 
    by (rule integrable_spike_finite[of "{}", OF _ _ integrable_on_powr_from_0[of a c]])
 | 
|
2792  | 
(insert a c, simp_all add: f_def)  | 
|
2793  | 
  moreover have A: "(\<lambda>x. exp (-x/2)) integrable_on {c..}"
 | 
|
2794  | 
using integrable_on_exp_minus_to_infinity[of "1/2"] by simp  | 
|
2795  | 
  have "f integrable_on {c..}"
 | 
|
2796  | 
    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
 | 
2797  | 
  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
 | 
2798  | 
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
 | 
2799  | 
qed  | 
| 63296 | 2800  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
2801  | 
theorem Gamma_integral_complex:  | 
| 63296 | 2802  | 
assumes z: "Re z > 0"  | 
2803  | 
  shows   "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}"
 | 
|
2804  | 
proof -  | 
|
2805  | 
have A: "((\<lambda>t. (of_real t) powr (z - 1) * of_real ((1 - t) ^ n))  | 
|
2806  | 
          has_integral (fact n / pochhammer z (n+1))) {0..1}"
 | 
|
2807  | 
if "Re z > 0" for n z using that  | 
|
2808  | 
proof (induction n arbitrary: z)  | 
|
2809  | 
case 0  | 
|
2810  | 
have "((\<lambda>t. complex_of_real t powr (z - 1)) has_integral  | 
|
2811  | 
            (of_real 1 powr z / z - of_real 0 powr z / z)) {0..1}" using 0
 | 
|
2812  | 
by (intro fundamental_theorem_of_calculus_interior)  | 
|
2813  | 
(auto intro!: continuous_intros derivative_eq_intros has_vector_derivative_real_complex)  | 
|
2814  | 
thus ?case by simp  | 
|
2815  | 
next  | 
|
2816  | 
case (Suc n)  | 
|
2817  | 
let ?f = "\<lambda>t. complex_of_real t powr z / z"  | 
|
2818  | 
let ?f' = "\<lambda>t. complex_of_real t powr (z - 1)"  | 
|
2819  | 
let ?g = "\<lambda>t. (1 - complex_of_real t) ^ Suc n"  | 
|
2820  | 
let ?g' = "\<lambda>t. - ((1 - complex_of_real t) ^ n) * of_nat (Suc n)"  | 
|
2821  | 
have "((\<lambda>t. ?f' t * ?g t) has_integral  | 
|
2822  | 
            (of_nat (Suc n)) * fact n / pochhammer z (n+2)) {0..1}"
 | 
|
2823  | 
(is "(_ has_integral ?I) _")  | 
|
2824  | 
proof (rule integration_by_parts_interior[where f' = ?f' and g = ?g])  | 
|
2825  | 
      from Suc.prems show "continuous_on {0..1} ?f" "continuous_on {0..1} ?g"
 | 
|
2826  | 
by (auto intro!: continuous_intros)  | 
|
2827  | 
next  | 
|
2828  | 
      fix t :: real assume t: "t \<in> {0<..<1}"
 | 
|
2829  | 
show "(?f has_vector_derivative ?f' t) (at t)" using t Suc.prems  | 
|
2830  | 
by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex)  | 
|
2831  | 
show "(?g has_vector_derivative ?g' t) (at t)"  | 
|
2832  | 
by (rule has_vector_derivative_real_complex derivative_eq_intros refl)+ simp_all  | 
|
2833  | 
next  | 
|
2834  | 
from Suc.prems have [simp]: "z \<noteq> 0" by auto  | 
|
2835  | 
from Suc.prems have A: "Re (z + of_nat n) > 0" for n by simp  | 
|
2836  | 
have [simp]: "z + of_nat n \<noteq> 0" "z + 1 + of_nat n \<noteq> 0" for n  | 
|
2837  | 
using A[of n] A[of "Suc n"] by (auto simp add: add.assoc simp del: plus_complex.sel)  | 
|
2838  | 
have "((\<lambda>x. of_real x powr z * of_real ((1 - x) ^ n) * (- of_nat (Suc n) / z)) has_integral  | 
|
2839  | 
              fact n / pochhammer (z+1) (n+1) * (- of_nat (Suc n) / z)) {0..1}"
 | 
|
2840  | 
(is "(?A has_integral ?B) _")  | 
|
2841  | 
using Suc.IH[of "z+1"] Suc.prems by (intro has_integral_mult_left) (simp_all add: add_ac pochhammer_rec)  | 
|
2842  | 
also have "?A = (\<lambda>t. ?f t * ?g' t)" by (intro ext) (simp_all add: field_simps)  | 
|
2843  | 
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
 | 
2844  | 
by (simp add: divide_simps pochhammer_rec  | 
| 64272 | 2845  | 
prod_shift_bounds_cl_Suc_ivl del: of_nat_Suc)  | 
| 63296 | 2846  | 
      finally show "((\<lambda>t. ?f t * ?g' t) has_integral (?f 1 * ?g 1 - ?f 0 * ?g 0 - ?I)) {0..1}"
 | 
2847  | 
by simp  | 
|
2848  | 
qed (simp_all add: bounded_bilinear_mult)  | 
|
2849  | 
thus ?case by simp  | 
|
2850  | 
qed  | 
|
2851  | 
||
2852  | 
  have B: "((\<lambda>t. if t \<in> {0..of_nat n} then
 | 
|
2853  | 
of_real t powr (z - 1) * (1 - of_real t / of_nat n) ^ n else 0)  | 
|
2854  | 
           has_integral (of_nat n powr z * fact n / pochhammer z (n+1))) {0..}" for n
 | 
|
2855  | 
proof (cases "n > 0")  | 
|
2856  | 
case [simp]: True  | 
|
2857  | 
hence [simp]: "n \<noteq> 0" by auto  | 
|
2858  | 
with has_integral_affinity01[OF A[OF z, of n], of "inverse (of_nat n)" 0]  | 
|
2859  | 
have "((\<lambda>x. (of_nat n - of_real x) ^ n * (of_real x / of_nat n) powr (z - 1) / of_nat n ^ n)  | 
|
2860  | 
              has_integral fact n * of_nat n / pochhammer z (n+1)) ((\<lambda>x. real n * x)`{0..1})"
 | 
|
2861  | 
(is "(?f has_integral ?I) ?ivl") by (simp add: field_simps scaleR_conv_of_real)  | 
|
2862  | 
    also from True have "((\<lambda>x. real n*x)`{0..1}) = {0..real n}"
 | 
|
2863  | 
by (subst image_mult_atLeastAtMost) simp_all  | 
|
2864  | 
also have "?f = (\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)"  | 
|
2865  | 
using True by (intro ext) (simp add: field_simps)  | 
|
2866  | 
finally have "((\<lambda>x. (of_real x / of_nat n) powr (z - 1) * (1 - of_real x / of_nat n) ^ n)  | 
|
2867  | 
                    has_integral ?I) {0..real n}" (is ?P) .
 | 
|
2868  | 
also have "?P \<longleftrightarrow> ((\<lambda>x. exp ((z - 1) * of_real (ln (x / of_nat n))) * (1 - of_real x / of_nat n) ^ n)  | 
|
2869  | 
                        has_integral ?I) {0..real n}"
 | 
|
2870  | 
      by (intro has_integral_spike_finite_eq[of "{0}"]) (auto simp: powr_def Ln_of_real [symmetric])
 | 
|
2871  | 
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)  | 
|
2872  | 
                        has_integral ?I) {0..real n}"
 | 
|
2873  | 
      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: ln_div)
 | 
|
2874  | 
finally have \<dots> .  | 
|
2875  | 
note B = has_integral_mult_right[OF this, of "exp ((z - 1) * ln (of_nat n))"]  | 
|
2876  | 
have "((\<lambda>x. exp ((z - 1) * of_real (ln x)) * (1 - of_real x / of_nat n) ^ n)  | 
|
2877  | 
            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}" (is ?P)
 | 
|
2878  | 
by (insert B, subst (asm) mult.assoc [symmetric], subst (asm) exp_add [symmetric])  | 
|
| 66936 | 2879  | 
(simp add: algebra_simps)  | 
| 63296 | 2880  | 
also have "?P \<longleftrightarrow> ((\<lambda>x. of_real x powr (z - 1) * (1 - of_real x / of_nat n) ^ n)  | 
2881  | 
            has_integral (?I * exp ((z - 1) * ln (of_nat n)))) {0..real n}"
 | 
|
2882  | 
      by (intro has_integral_spike_finite_eq[of "{0}"]) (simp_all add: powr_def Ln_of_real)
 | 
|
2883  | 
also have "fact n * of_nat n / pochhammer z (n+1) * exp ((z - 1) * Ln (of_nat n)) =  | 
|
2884  | 
(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
 | 
2885  | 
by (auto simp add: powr_def algebra_simps exp_diff exp_of_real)  | 
| 63296 | 2886  | 
finally show ?thesis by (subst has_integral_restrict) simp_all  | 
2887  | 
next  | 
|
2888  | 
case False  | 
|
2889  | 
thus ?thesis by (subst has_integral_restrict) (simp_all add: has_integral_refl)  | 
|
2890  | 
qed  | 
|
2891  | 
||
2892  | 
have "eventually (\<lambda>n. Gamma_series z n =  | 
|
2893  | 
of_nat n powr z * fact n / pochhammer z (n+1)) sequentially"  | 
|
2894  | 
using eventually_gt_at_top[of "0::nat"]  | 
|
2895  | 
by eventually_elim (simp add: powr_def algebra_simps Ln_of_nat Gamma_series_def)  | 
|
2896  | 
from this and Gamma_series_LIMSEQ[of z]  | 
|
2897  | 
have C: "(\<lambda>k. of_nat k powr z * fact k / pochhammer z (k+1)) \<longlonglongrightarrow> Gamma z"  | 
|
2898  | 
by (rule Lim_transform_eventually)  | 
|
2899  | 
||
2900  | 
  {
 | 
|
2901  | 
fix x :: real assume x: "x \<ge> 0"  | 
|
2902  | 
have lim_exp: "(\<lambda>k. (1 - x / real k) ^ k) \<longlonglongrightarrow> exp (-x)"  | 
|
2903  | 
using tendsto_exp_limit_sequentially[of "-x"] by simp  | 
|
2904  | 
have "(\<lambda>k. of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k))  | 
|
2905  | 
\<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))" (is ?P)  | 
|
2906  | 
by (intro tendsto_intros lim_exp)  | 
|
2907  | 
also from eventually_gt_at_top[of "nat \<lceil>x\<rceil>"]  | 
|
2908  | 
have "eventually (\<lambda>k. of_nat k > x) sequentially" by eventually_elim linarith  | 
|
2909  | 
hence "?P \<longleftrightarrow> (\<lambda>k. if x \<le> of_nat k then  | 
|
2910  | 
of_real x powr (z - 1) * of_real ((1 - x / of_nat k) ^ k) else 0)  | 
|
2911  | 
\<longlonglongrightarrow> of_real x powr (z - 1) * of_real (exp (-x))"  | 
|
2912  | 
by (intro tendsto_cong) (auto elim!: eventually_mono)  | 
|
2913  | 
finally have \<dots> .  | 
|
2914  | 
}  | 
|
2915  | 
  hence D: "\<forall>x\<in>{0..}. (\<lambda>k. if x \<in> {0..real k} then
 | 
|
2916  | 
of_real x powr (z - 1) * (1 - of_real x / of_nat k) ^ k else 0)  | 
|
2917  | 
\<longlonglongrightarrow> of_real x powr (z - 1) / of_real (exp x)"  | 
|
2918  | 
by (simp add: exp_minus field_simps cong: if_cong)  | 
|
2919  | 
||
2920  | 
have "((\<lambda>x. (Re z - 1) * (ln x / x)) \<longlongrightarrow> (Re z - 1) * 0) at_top"  | 
|
2921  | 
by (intro tendsto_intros ln_x_over_x_tendsto_0)  | 
|
2922  | 
hence "((\<lambda>x. ((Re z - 1) * ln x) / x) \<longlongrightarrow> 0) at_top" by simp  | 
|
2923  | 
from order_tendstoD(2)[OF this, of "1/2"]  | 
|
2924  | 
have "eventually (\<lambda>x. (Re z - 1) * ln x / x < 1/2) at_top" by simp  | 
|
2925  | 
from eventually_conj[OF this eventually_gt_at_top[of 0]]  | 
|
2926  | 
obtain x0 where "\<forall>x\<ge>x0. (Re z - 1) * ln x / x < 1/2 \<and> x > 0"  | 
|
2927  | 
by (auto simp: eventually_at_top_linorder)  | 
|
2928  | 
hence x0: "x0 > 0" "\<And>x. x \<ge> x0 \<Longrightarrow> (Re z - 1) * ln x < x / 2" by auto  | 
|
2929  | 
||
2930  | 
  define h where "h = (\<lambda>x. if x \<in> {0..x0} then x powr (Re z - 1) else exp (-x/2))"
 | 
|
2931  | 
have le_h: "x powr (Re z - 1) * exp (-x) \<le> h x" if x: "x \<ge> 0" for x  | 
|
2932  | 
proof (cases "x > x0")  | 
|
2933  | 
case True  | 
|
2934  | 
from True x0(1) have "x powr (Re z - 1) * exp (-x) = exp ((Re z - 1) * ln x - x)"  | 
|
2935  | 
by (simp add: powr_def exp_diff exp_minus field_simps exp_add)  | 
|
2936  | 
also from x0(2)[of x] True have "\<dots> < exp (-x/2)"  | 
|
2937  | 
by (simp add: field_simps)  | 
|
2938  | 
finally show ?thesis using True by (auto simp add: h_def)  | 
|
2939  | 
next  | 
|
2940  | 
case False  | 
|
2941  | 
from x have "x powr (Re z - 1) * exp (- x) \<le> x powr (Re z - 1) * 1"  | 
|
2942  | 
by (intro mult_left_mono) simp_all  | 
|
2943  | 
with False show ?thesis by (auto simp add: h_def)  | 
|
2944  | 
qed  | 
|
2945  | 
||
2946  | 
  have E: "\<forall>x\<in>{0..}. cmod (if x \<in> {0..real k} then of_real x powr (z - 1) *
 | 
|
2947  | 
(1 - complex_of_real x / of_nat k) ^ k else 0) \<le> h x"  | 
|
2948  | 
(is "\<forall>x\<in>_. ?f x \<le> _") for k  | 
|
2949  | 
proof safe  | 
|
2950  | 
fix x :: real assume x: "x \<ge> 0"  | 
|
2951  | 
    {
 | 
|
2952  | 
fix x :: real and n :: nat assume x: "x \<le> of_nat n"  | 
|
2953  | 
have "(1 - complex_of_real x / of_nat n) = complex_of_real ((1 - x / of_nat n))" by simp  | 
|
2954  | 
also have "norm \<dots> = \<bar>(1 - x / real n)\<bar>" by (subst norm_of_real) (rule refl)  | 
|
2955  | 
also from x have "\<dots> = (1 - x / real n)" by (intro abs_of_nonneg) (simp_all add: divide_simps)  | 
|
2956  | 
finally have "cmod (1 - complex_of_real x / of_nat n) = 1 - x / real n" .  | 
|
2957  | 
} note D = this  | 
|
2958  | 
from D[of x k] x  | 
|
2959  | 
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)"  | 
|
2960  | 
by (auto simp: norm_mult norm_powr_real_powr norm_power intro!: mult_nonneg_nonneg)  | 
|
2961  | 
also have "\<dots> \<le> x powr (Re z - 1) * exp (-x)"  | 
|
2962  | 
by (auto intro!: mult_left_mono exp_ge_one_minus_x_over_n_power_n)  | 
|
2963  | 
also from x have "\<dots> \<le> h x" by (rule le_h)  | 
|
2964  | 
finally show "?f x \<le> h x" .  | 
|
2965  | 
qed  | 
|
| 
63594
 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 
hoelzl 
parents: 
63539 
diff
changeset
 | 
2966  | 
|
| 63296 | 2967  | 
  have F: "h integrable_on {0..}" unfolding h_def
 | 
2968  | 
by (rule integrable_Gamma_integral_bound) (insert assms x0(1), simp_all)  | 
|
2969  | 
show ?thesis  | 
|
2970  | 
by (rule has_integral_dominated_convergence[OF B F E D C])  | 
|
2971  | 
qed  | 
|
2972  | 
||
2973  | 
lemma Gamma_integral_real:  | 
|
2974  | 
assumes x: "x > (0 :: real)"  | 
|
2975  | 
  shows   "((\<lambda>t. t powr (x - 1) / exp t) has_integral Gamma x) {0..}"
 | 
|
2976  | 
proof -  | 
|
2977  | 
have A: "((\<lambda>t. complex_of_real t powr (complex_of_real x - 1) /  | 
|
2978  | 
          complex_of_real (exp t)) has_integral complex_of_real (Gamma x)) {0..}"
 | 
|
2979  | 
using Gamma_integral_complex[of x] assms by (simp_all add: Gamma_complex_of_real powr_of_real)  | 
|
2980  | 
  have "((\<lambda>t. complex_of_real (t powr (x - 1) / exp t)) has_integral of_real (Gamma x)) {0..}"
 | 
|
2981  | 
by (rule has_integral_eq[OF _ A]) (simp_all add: powr_of_real [symmetric])  | 
|
2982  | 
from has_integral_linear[OF this bounded_linear_Re] show ?thesis by (simp add: o_def)  | 
|
2983  | 
qed  | 
|
2984  | 
||
| 66526 | 2985  | 
lemma absolutely_integrable_Gamma_integral':  | 
2986  | 
assumes "Re z > 0"  | 
|
2987  | 
  shows   "(\<lambda>t. complex_of_real t powr (z - 1) / of_real (exp t)) absolutely_integrable_on {0<..}"
 | 
|
2988  | 
using absolutely_integrable_Gamma_integral [OF assms zero_less_one] by simp  | 
|
2989  | 
||
2990  | 
lemma Gamma_integral_complex':  | 
|
2991  | 
assumes z: "Re z > 0"  | 
|
2992  | 
  shows   "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0<..}"
 | 
|
2993  | 
proof -  | 
|
2994  | 
  have "((\<lambda>t. of_real t powr (z - 1) / of_real (exp t)) has_integral Gamma z) {0..}"
 | 
|
2995  | 
by (rule Gamma_integral_complex) fact+  | 
|
2996  | 
  hence "((\<lambda>t. if t \<in> {0<..} then of_real t powr (z - 1) / of_real (exp t) else 0) 
 | 
|
2997  | 
           has_integral Gamma z) {0..}"
 | 
|
2998  | 
    by (rule has_integral_spike [of "{0}", rotated 2]) auto
 | 
|
2999  | 
also have "?this = ?thesis"  | 
|
3000  | 
by (subst has_integral_restrict) auto  | 
|
3001  | 
finally show ?thesis .  | 
|
3002  | 
qed  | 
|
3003  | 
||
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3004  | 
lemma Gamma_conv_nn_integral_real:  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3005  | 
assumes "s > (0::real)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3006  | 
  shows   "Gamma s = nn_integral lborel (\<lambda>t. ennreal (indicator {0..} t * t powr (s - 1) / exp t))"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3007  | 
using nn_integral_has_integral_lebesgue[OF _ Gamma_integral_real[OF assms]] by simp  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3008  | 
|
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3009  | 
lemma integrable_Beta:  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3010  | 
assumes "a > 0" "b > (0::real)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3011  | 
  shows   "set_integrable lborel {0..1} (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3012  | 
proof -  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3013  | 
define C where "C = max 1 ((1/2) powr (b - 1))"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3014  | 
define D where "D = max 1 ((1/2) powr (a - 1))"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3015  | 
  have C: "(1 - x) powr (b - 1) \<le> C" if "x \<in> {0..1/2}" for x
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3016  | 
proof (cases "b < 1")  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3017  | 
case False  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3018  | 
with that have "(1 - x) powr (b - 1) \<le> (1 powr (b - 1))" by (intro powr_mono2) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3019  | 
thus ?thesis by (auto simp: C_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3020  | 
qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 C_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3021  | 
  have D: "x powr (a - 1) \<le> D" if "x \<in> {1/2..1}" for x
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3022  | 
proof (cases "a < 1")  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3023  | 
case False  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3024  | 
with that have "x powr (a - 1) \<le> (1 powr (a - 1))" by (intro powr_mono2) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3025  | 
thus ?thesis by (auto simp: D_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3026  | 
next  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3027  | 
case True  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3028  | 
qed (insert that, auto simp: max.coboundedI1 max.coboundedI2 powr_mono2' powr_mono2 D_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3029  | 
have [simp]: "C \<ge> 0" "D \<ge> 0" by (simp_all add: C_def D_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3030  | 
|
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3031  | 
  have I1: "set_integrable lborel {0..1/2} (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
 | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3032  | 
unfolding set_integrable_def  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3033  | 
proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2])  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3034  | 
    have "(\<lambda>t. t powr (a - 1)) integrable_on {0..1/2}"
 | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3035  | 
by (rule integrable_on_powr_from_0) (use assms in auto)  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3036  | 
    hence "(\<lambda>t. t powr (a - 1)) absolutely_integrable_on {0..1/2}"
 | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3037  | 
by (subst absolutely_integrable_on_iff_nonneg) auto  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3038  | 
from integrable_mult_right[OF this [unfolded set_integrable_def], of C]  | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3039  | 
    show "integrable lborel (\<lambda>x. indicat_real {0..1/2} x *\<^sub>R (C * x powr (a - 1)))"
 | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3040  | 
by (subst (asm) integrable_completion) (auto simp: mult_ac)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3041  | 
next  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3042  | 
fix x :: real  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3043  | 
    have "x powr (a - 1) * (1 - x) powr (b - 1) \<le> x powr (a - 1) * C" if "x \<in> {0..1/2}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3044  | 
using that by (intro mult_left_mono powr_mono2 C) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3045  | 
    thus "norm (indicator {0..1/2} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \<le>
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3046  | 
            norm (indicator {0..1/2} x *\<^sub>R (C * x powr (a - 1)))"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3047  | 
by (auto simp: indicator_def abs_mult mult_ac)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3048  | 
qed (auto intro!: AE_I2 simp: indicator_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3049  | 
|
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3050  | 
  have I2: "set_integrable lborel {1/2..1} (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
 | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3051  | 
unfolding set_integrable_def  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3052  | 
proof (rule Bochner_Integration.integrable_bound[OF _ _ AE_I2])  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3053  | 
    have "(\<lambda>t. t powr (b - 1)) integrable_on {0..1/2}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3054  | 
by (rule integrable_on_powr_from_0) (use assms in auto)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3055  | 
hence "(\<lambda>t. t powr (b - 1)) integrable_on (cbox 0 (1/2))" by simp  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3056  | 
from integrable_affinity[OF this, of "-1" 1]  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3057  | 
      have "(\<lambda>t. (1 - t) powr (b - 1)) integrable_on {1/2..1}" by simp
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3058  | 
    hence "(\<lambda>t. (1 - t) powr (b - 1)) absolutely_integrable_on {1/2..1}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3059  | 
by (subst absolutely_integrable_on_iff_nonneg) auto  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3060  | 
from integrable_mult_right[OF this [unfolded set_integrable_def], of D]  | 
| 
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3061  | 
    show "integrable lborel (\<lambda>x. indicat_real {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))"
 | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3062  | 
by (subst (asm) integrable_completion) (auto simp: mult_ac)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3063  | 
next  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3064  | 
fix x :: real  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3065  | 
    have "x powr (a - 1) * (1 - x) powr (b - 1) \<le> D * (1 - x) powr (b - 1)" if "x \<in> {1/2..1}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3066  | 
using that by (intro mult_right_mono powr_mono2 D) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3067  | 
    thus "norm (indicator {1/2..1} x *\<^sub>R (x powr (a - 1) * (1 - x) powr (b - 1))) \<le>
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3068  | 
            norm (indicator {1/2..1} x *\<^sub>R (D * (1 - x) powr (b - 1)))"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3069  | 
by (auto simp: indicator_def abs_mult mult_ac)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3070  | 
qed (auto intro!: AE_I2 simp: indicator_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3071  | 
|
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3072  | 
  have "set_integrable lborel ({0..1/2} \<union> {1/2..1}) (\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1))"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3073  | 
by (intro set_integrable_Un I1 I2) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3074  | 
  also have "{0..1/2} \<union> {1/2..1} = {0..(1::real)}" by auto
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3075  | 
finally show ?thesis .  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3076  | 
qed  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3077  | 
|
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3078  | 
lemma integrable_Beta':  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3079  | 
assumes "a > 0" "b > (0::real)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3080  | 
  shows   "(\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3081  | 
using integrable_Beta[OF assms] by (rule set_borel_integral_eq_integral)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3082  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
3083  | 
theorem has_integral_Beta_real:  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3084  | 
assumes a: "a > 0" and b: "b > (0 :: real)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3085  | 
  shows "((\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1)) has_integral Beta a b) {0..1}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3086  | 
proof -  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3087  | 
  define B where "B = integral {0..1} (\<lambda>x. x powr (a - 1) * (1 - x) powr (b - 1))"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3088  | 
have [simp]: "B \<ge> 0" unfolding B_def using a b  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3089  | 
by (intro integral_nonneg integrable_Beta') auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3090  | 
from a b have "ennreal (Gamma a * Gamma b) =  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3091  | 
    (\<integral>\<^sup>+ t. ennreal (indicator {0..} t * t powr (a - 1) / exp t) \<partial>lborel) *
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3092  | 
    (\<integral>\<^sup>+ t. ennreal (indicator {0..} t * t powr (b - 1) / exp t) \<partial>lborel)"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3093  | 
by (subst ennreal_mult') (simp_all add: Gamma_conv_nn_integral_real)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3094  | 
  also have "\<dots> = (\<integral>\<^sup>+t. \<integral>\<^sup>+u. ennreal (indicator {0..} t * t powr (a - 1) / exp t) *
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3095  | 
                            ennreal (indicator {0..} u * u powr (b - 1) / exp u) \<partial>lborel \<partial>lborel)"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3096  | 
by (simp add: nn_integral_cmult nn_integral_multc)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3097  | 
  also have "\<dots> = (\<integral>\<^sup>+t. \<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{0..}) (t,u) * t powr (a - 1) * u powr (b - 1)
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3098  | 
/ exp (t + u)) \<partial>lborel \<partial>lborel)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3099  | 
by (intro nn_integral_cong)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3100  | 
(auto simp: indicator_def divide_ennreal ennreal_mult' [symmetric] exp_add)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3101  | 
  also have "\<dots> = (\<integral>\<^sup>+t. \<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{t..}) (t,u) * t powr (a - 1) * 
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3102  | 
(u - t) powr (b - 1) / exp u) \<partial>lborel \<partial>lborel)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3103  | 
proof (rule nn_integral_cong, goal_cases)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3104  | 
case (1 t)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3105  | 
    have "(\<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{0..}) (t,u) * t powr (a - 1) * 
 | 
| 67399 | 3106  | 
u powr (b - 1) / exp (t + u)) \<partial>distr lborel borel ((+) (-t))) =  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3107  | 
               (\<integral>\<^sup>+u. ennreal (indicator ({0..}\<times>{t..}) (t,u) * t powr (a - 1) * 
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3108  | 
(u - t) powr (b - 1) / exp u) \<partial>lborel)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3109  | 
by (subst nn_integral_distr) (auto intro!: nn_integral_cong simp: indicator_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3110  | 
thus ?case by (subst (asm) lborel_distr_plus)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3111  | 
qed  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3112  | 
  also have "\<dots> = (\<integral>\<^sup>+u. \<integral>\<^sup>+t. ennreal (indicator ({0..}\<times>{t..}) (t,u) * t powr (a - 1) * 
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3113  | 
(u - t) powr (b - 1) / exp u) \<partial>lborel \<partial>lborel)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3114  | 
by (subst lborel_pair.Fubini')  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3115  | 
(auto simp: case_prod_unfold indicator_def cong: measurable_cong_sets)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3116  | 
  also have "\<dots> = (\<integral>\<^sup>+u. \<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) *
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3117  | 
                              ennreal (indicator {0..} u / exp u) \<partial>lborel \<partial>lborel)"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3118  | 
by (intro nn_integral_cong) (auto simp: indicator_def ennreal_mult' [symmetric])  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3119  | 
  also have "\<dots> = (\<integral>\<^sup>+u. (\<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) 
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3120  | 
                          \<partial>lborel) * ennreal (indicator {0..} u / exp u) \<partial>lborel)"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3121  | 
by (subst nn_integral_multc [symmetric]) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3122  | 
  also have "\<dots> = (\<integral>\<^sup>+u. (\<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) 
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3123  | 
                          \<partial>lborel) * ennreal (indicator {0<..} u / exp u) \<partial>lborel)"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3124  | 
by (intro nn_integral_cong_AE eventually_mono[OF AE_lborel_singleton[of 0]])  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3125  | 
(auto simp: indicator_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3126  | 
  also have "\<dots> = (\<integral>\<^sup>+u. ennreal B * ennreal (indicator {0..} u / exp u * u powr (a + b - 1)) \<partial>lborel)"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3127  | 
proof (intro nn_integral_cong, goal_cases)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3128  | 
case (1 u)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3129  | 
show ?case  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3130  | 
proof (cases "u > 0")  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3131  | 
case True  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3132  | 
      have "(\<integral>\<^sup>+t. ennreal (indicator {0..u} t * t powr (a - 1) * (u - t) powr (b - 1)) \<partial>lborel) = 
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3133  | 
              (\<integral>\<^sup>+t. ennreal (indicator {0..1} t * (u * t) powr (a - 1) * (u - u * t) powr (b - 1)) 
 | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
68721 
diff
changeset
 | 
3134  | 
\<partial>distr lborel borel ((*) (1 / u)))" (is "_ = nn_integral _ ?f")  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3135  | 
using True  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3136  | 
by (subst nn_integral_distr) (auto simp: indicator_def field_simps intro!: nn_integral_cong)  | 
| 
69064
 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 
nipkow 
parents: 
68721 
diff
changeset
 | 
3137  | 
also have "distr lborel borel ((*) (1 / u)) = density lborel (\<lambda>_. u)"  | 
| 
67278
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3138  | 
using \<open>u > 0\<close> by (subst lborel_distr_mult) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3139  | 
      also have "nn_integral \<dots> ?f = (\<integral>\<^sup>+x. ennreal (indicator {0..1} x * (u * (u * x) powr (a - 1) *
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3140  | 
(u * (1 - x)) powr (b - 1))) \<partial>lborel)" using \<open>u > 0\<close>  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3141  | 
by (subst nn_integral_density) (auto simp: ennreal_mult' [symmetric] algebra_simps)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3142  | 
also have "\<dots> = (\<integral>\<^sup>+x. ennreal (u powr (a + b - 1)) *  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3143  | 
                            ennreal (indicator {0..1} x * x powr (a - 1) *
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3144  | 
(1 - x) powr (b - 1)) \<partial>lborel)" using \<open>u > 0\<close> a b  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3145  | 
by (intro nn_integral_cong)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3146  | 
(auto simp: indicator_def powr_mult powr_add powr_diff mult_ac ennreal_mult' [symmetric])  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3147  | 
also have "\<dots> = ennreal (u powr (a + b - 1)) *  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3148  | 
                        (\<integral>\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) *
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3149  | 
(1 - x) powr (b - 1)) \<partial>lborel)"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3150  | 
by (subst nn_integral_cmult) auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3151  | 
also have "((\<lambda>x. x powr (a - 1) * (1 - x) powr (b - 1)) has_integral  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3152  | 
                   integral {0..1} (\<lambda>x. x powr (a - 1) * (1 - x) powr (b - 1))) {0..1}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3153  | 
using a b by (intro integrable_integral integrable_Beta')  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3154  | 
from nn_integral_has_integral_lebesgue[OF _ this] a b  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3155  | 
        have "(\<integral>\<^sup>+x. ennreal (indicator {0..1} x * x powr (a - 1) *
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3156  | 
(1 - x) powr (b - 1)) \<partial>lborel) = B" by (simp add: mult_ac B_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3157  | 
finally show ?thesis using \<open>u > 0\<close> by (simp add: ennreal_mult' [symmetric] mult_ac)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3158  | 
qed auto  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3159  | 
qed  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3160  | 
also have "\<dots> = ennreal B * ennreal (Gamma (a + b))"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3161  | 
using a b by (subst nn_integral_cmult) (auto simp: Gamma_conv_nn_integral_real)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3162  | 
also have "\<dots> = ennreal (B * Gamma (a + b))"  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3163  | 
by (subst (1 2) mult.commute, intro ennreal_mult' [symmetric]) (use a b in auto)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3164  | 
finally have "B = Beta a b" using a b Gamma_real_pos[of "a + b"]  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3165  | 
by (subst (asm) ennreal_inj) (auto simp: field_simps Beta_def Gamma_eq_zero_iff)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3166  | 
  moreover have "(\<lambda>t. t powr (a - 1) * (1 - t) powr (b - 1)) integrable_on {0..1}"
 | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3167  | 
by (intro integrable_Beta' a b)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3168  | 
ultimately show ?thesis by (simp add: has_integral_iff B_def)  | 
| 
 
c60e3d615b8c
Removed Analysis/ex/Circle_Area; replaced by more general Analysis/Ball_Volume
 
eberlm <eberlm@in.tum.de> 
parents: 
66936 
diff
changeset
 | 
3169  | 
qed  | 
| 63296 | 3170  | 
|
3171  | 
||
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
3172  | 
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
 | 
3173  | 
|
| 
68624
 
205d352ed727
Tagged Ball_Volume and Gamma_Function in HOL-Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
68403 
diff
changeset
 | 
3174  | 
theorem sin_product_formula_complex:  | 
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3175  | 
fixes z :: complex  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3176  | 
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
 | 
3177  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3178  | 
let ?f = "rGamma_series_weierstrass"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3179  | 
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
 | 
3180  | 
\<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
 | 
3181  | 
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
 | 
3182  | 
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
 | 
3183  | 
(\<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
 | 
3184  | 
proof  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3185  | 
fix n :: nat  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3186  | 
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
 | 
3187  | 
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
 | 
3188  | 
by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus  | 
| 64272 | 3189  | 
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
 | 
3190  | 
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
 | 
3191  | 
(\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2)"  | 
| 64272 | 3192  | 
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
 | 
3193  | 
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
 | 
3194  | 
by (simp add: divide_simps)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3195  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3196  | 
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
 | 
3197  | 
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
 | 
3198  | 
finally show ?thesis .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3199  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3200  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3201  | 
lemma sin_product_formula_real:  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3202  | 
"(\<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
 | 
3203  | 
proof -  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3204  | 
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
 | 
3205  | 
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
 | 
3206  | 
\<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
 | 
3207  | 
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
 | 
3208  | 
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
 | 
3209  | 
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
 | 
3210  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3211  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3212  | 
lemma sin_product_formula_real':  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3213  | 
assumes "x \<noteq> (0::real)"  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3214  | 
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
 | 
3215  | 
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
 | 
3216  | 
by simp  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3217  | 
|
| 63721 | 3218  | 
theorem wallis: "(\<lambda>n. \<Prod>k=1..n. (4*real k^2) / (4*real k^2 - 1)) \<longlonglongrightarrow> pi / 2"  | 
3219  | 
proof -  | 
|
| 
66512
 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 
paulson <lp15@cam.ac.uk> 
parents: 
66453 
diff
changeset
 | 
3220  | 
from tendsto_inverse[OF tendsto_mult[OF  | 
| 63721 | 3221  | 
sin_product_formula_real[of "1/2"] tendsto_const[of "2/pi"]]]  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3222  | 
have "(\<lambda>n. (\<Prod>k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) \<longlonglongrightarrow> pi/2"  | 
| 64272 | 3223  | 
by (simp add: prod_inversef [symmetric])  | 
| 
67976
 
75b94eb58c3d
Analysis builds using set_borel_measurable_def, etc.
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
3224  | 
also have "(\<lambda>n. (\<Prod>k=1..n. inverse (1 - (1/2)\<^sup>2 / (real k)\<^sup>2))) =  | 
| 63721 | 3225  | 
(\<lambda>n. (\<Prod>k=1..n. (4*real k^2)/(4*real k^2 - 1)))"  | 
| 64272 | 3226  | 
by (intro ext prod.cong refl) (simp add: divide_simps)  | 
| 63721 | 3227  | 
finally show ?thesis .  | 
3228  | 
qed  | 
|
3229  | 
||
| 
62085
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
3230  | 
|
| 
 
5b7758af429e
Tuned approximations in Multivariate_Analysis
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
62072 
diff
changeset
 | 
3231  | 
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
 | 
3232  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3233  | 
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
 | 
3234  | 
proof -  | 
| 63040 | 3235  | 
define P where "P x n = (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)" for x :: real and n  | 
3236  | 
define K where "K = (\<Sum>n. inverse (real_of_nat (Suc n))^2)"  | 
|
3237  | 
define f where [abs_def]: "f x = (\<Sum>n. P x n / of_nat (Suc n)^2)" for x  | 
|
3238  | 
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
 | 
3239  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3240  | 
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
 | 
3241  | 
proof (cases "x = 0")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3242  | 
assume x: "x = 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3243  | 
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
 | 
3244  | 
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
 | 
3245  | 
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
 | 
3246  | 
next  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3247  | 
assume x: "x \<noteq> 0"  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3248  | 
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
 | 
3249  | 
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
 | 
3250  | 
also have "(\<lambda>n. P x n - P x (Suc n)) = (\<lambda>n. (x^2 / of_nat (Suc n)^2) * P x n)"  | 
| 64272 | 3251  | 
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
 | 
3252  | 
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
 | 
3253  | 
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
 | 
3254  | 
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
 | 
3255  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3256  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3257  | 
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
 | 
3258  | 
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
 | 
3259  | 
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
 | 
3260  | 
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
 | 
3261  | 
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
 | 
3262  | 
      {
 | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3263  | 
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
 | 
3264  | 
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
 | 
3265  | 
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
 | 
3266  | 
        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
 | 
3267  | 
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
 | 
3268  | 
}  | 
| 64272 | 3269  | 
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
 | 
3270  | 
thus "norm (P x n / (of_nat (Suc n)^2)) \<le> 1 / of_nat (Suc n)^2"  | 
| 64272 | 3271  | 
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
 | 
3272  | 
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
 | 
3273  | 
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
 | 
3274  | 
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
 | 
3275  | 
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
 | 
3276  | 
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
 | 
3277  | 
finally have "f \<midarrow> 0 \<rightarrow> K" .  | 
| 
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3278  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3279  | 
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
 | 
3280  | 
proof (rule Lim_transform_eventually)  | 
| 63040 | 3281  | 
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
 | 
3282  | 
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
 | 
3283  | 
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
 | 
3284  | 
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
 | 
3285  | 
proof eventually_elim  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3286  | 
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
 | 
3287  | 
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
 | 
3288  | 
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
 | 
3289  | 
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
 | 
3290  | 
by (simp add: eval_nat_numeral)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3291  | 
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
 | 
3292  | 
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
 | 
3293  | 
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
 | 
3294  | 
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
 | 
3295  | 
by (simp add: g_def)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3296  | 
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
 | 
3297  | 
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
 | 
3298  | 
finally show "f' x = f x" .  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3299  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3300  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3301  | 
have "isCont f' 0" unfolding f'_def  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3302  | 
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
 | 
3303  | 
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
 | 
3304  | 
proof (cases "x = 0")  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3305  | 
assume x: "x \<noteq> 0"  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3306  | 
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
 | 
3307  | 
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
 | 
3308  | 
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
 | 
3309  | 
qed (simp only: summable_0_powser)  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3310  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3311  | 
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
 | 
3312  | 
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
 | 
3313  | 
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
 | 
3314  | 
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
 | 
3315  | 
qed  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3316  | 
|
| 
62049
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3317  | 
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
 | 
3318  | 
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
 | 
3319  | 
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
 | 
3320  | 
by (subst summable_Suc_iff) (simp add: power_inverse)  | 
| 
62131
 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 
paulson 
parents: 
62085 
diff
changeset
 | 
3321  | 
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
 | 
3322  | 
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
 | 
3323  | 
qed  | 
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3324  | 
|
| 
 
b0f941e207cf
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
 
eberlm 
parents:  
diff
changeset
 | 
3325  | 
end  |