Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
authoreberlm
Mon, 04 Jan 2016 17:45:36 +0100
changeset 62049 b0f941e207cf
parent 62048 fefd79f6b232
child 62054 cff7114e4572
Added lots of material on infinite sums, convergence radii, harmonic numbers, Gamma function
src/HOL/Library/Extended_Real.thy
src/HOL/Library/Liminf_Limsup.thy
src/HOL/Library/Nonpos_Ints.thy
src/HOL/Library/Periodic_Fun.thy
src/HOL/Multivariate_Analysis/Complex_Transcendental.thy
src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy
src/HOL/Multivariate_Analysis/Gamma.thy
src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy
src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy
src/HOL/Multivariate_Analysis/Integral_Test.thy
src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy
src/HOL/Multivariate_Analysis/Summation.thy
src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy
src/HOL/Real_Vector_Spaces.thy
src/HOL/Series.thy
src/HOL/Topological_Spaces.thy
src/HOL/Transcendental.thy
--- a/src/HOL/Library/Extended_Real.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Library/Extended_Real.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -3,6 +3,7 @@
     Author:     Robert Himmelmann, TU München
     Author:     Armin Heller, TU München
     Author:     Bogdan Grechuk, University of Edinburgh
+    Author:     Manuel Eberl, TU München
 *)
 
 section \<open>Extended real number line\<close>
@@ -1629,6 +1630,14 @@
   by (cases a b c rule: ereal3_cases)
      (auto simp: field_simps zero_less_mult_iff)
 
+lemma ereal_inverse_real: "\<bar>z\<bar> \<noteq> \<infinity> \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> ereal (inverse (real_of_ereal z)) = inverse z"
+  by (cases z) simp_all
+
+lemma ereal_inverse_mult:
+  "a \<noteq> 0 \<Longrightarrow> b \<noteq> 0 \<Longrightarrow> inverse (a * (b::ereal)) = inverse a * inverse b"
+  by (cases a; cases b) auto
+
+     
 subsection "Complete lattice"
 
 instantiation ereal :: lattice
@@ -2472,6 +2481,21 @@
     by (auto simp: order_tendsto_iff)
 qed
 
+lemma tendsto_PInfty': "(f \<longlongrightarrow> \<infinity>) F = (\<forall>r>c. eventually (\<lambda>x. ereal r < f x) F)"
+proof (subst tendsto_PInfty, intro iffI allI impI)
+  assume A: "\<forall>r>c. eventually (\<lambda>x. ereal r < f x) F"
+  fix r :: real
+  from A have A: "eventually (\<lambda>x. ereal r < f x) F" if "r > c" for r using that by blast
+  show "eventually (\<lambda>x. ereal r < f x) F"
+  proof (cases "r > c")
+    case False
+    hence B: "ereal r \<le> ereal (c + 1)" by simp
+    have "c < c + 1" by simp
+    from A[OF this] show "eventually (\<lambda>x. ereal r < f x) F"
+      by eventually_elim (rule le_less_trans[OF B])
+  qed (simp add: A)
+qed simp
+
 lemma tendsto_PInfty_eq_at_top:
   "((\<lambda>z. ereal (f z)) \<longlongrightarrow> \<infinity>) F \<longleftrightarrow> (LIM z F. f z :> at_top)"
   unfolding tendsto_PInfty filterlim_at_top_dense by simp
@@ -2495,6 +2519,21 @@
     by auto
 qed
 
+lemma tendsto_MInfty': "(f \<longlongrightarrow> -\<infinity>) F = (\<forall>r<c. eventually (\<lambda>x. ereal r > f x) F)"
+proof (subst tendsto_MInfty, intro iffI allI impI)
+  assume A: "\<forall>r<c. eventually (\<lambda>x. ereal r > f x) F"
+  fix r :: real
+  from A have A: "eventually (\<lambda>x. ereal r > f x) F" if "r < c" for r using that by blast
+  show "eventually (\<lambda>x. ereal r > f x) F"
+  proof (cases "r < c")
+    case False
+    hence B: "ereal r \<ge> ereal (c - 1)" by simp
+    have "c > c - 1" by simp
+    from A[OF this] show "eventually (\<lambda>x. ereal r > f x) F"
+      by eventually_elim (erule less_le_trans[OF _ B])
+  qed (simp add: A)
+qed simp
+
 lemma Lim_PInfty: "f \<longlonglongrightarrow> \<infinity> \<longleftrightarrow> (\<forall>B. \<exists>N. \<forall>n\<ge>N. f n \<ge> ereal B)"
   unfolding tendsto_PInfty eventually_sequentially
 proof safe
@@ -2527,6 +2566,19 @@
 lemma Lim_bounded_MInfty: "f \<longlonglongrightarrow> l \<Longrightarrow> (\<And>n. ereal B \<le> f n) \<Longrightarrow> l \<noteq> -\<infinity>"
   using LIMSEQ_le_const[of f l "ereal B"] by auto
 
+lemma tendsto_zero_erealI:
+  assumes "\<And>e. e > 0 \<Longrightarrow> eventually (\<lambda>x. \<bar>f x\<bar> < ereal e) F"
+  shows   "(f \<longlongrightarrow> 0) F"
+proof (subst filterlim_cong[OF refl refl])
+  from assms[OF zero_less_one] show "eventually (\<lambda>x. f x = ereal (real_of_ereal (f x))) F"
+    by eventually_elim (auto simp: ereal_real)
+  hence "eventually (\<lambda>x. abs (real_of_ereal (f x)) < e) F" if "e > 0" for e using assms[OF that]
+    by eventually_elim (simp add: real_less_ereal_iff that)
+  hence "((\<lambda>x. real_of_ereal (f x)) \<longlongrightarrow> 0) F" unfolding tendsto_iff 
+    by (auto simp: tendsto_iff dist_real_def)
+  thus "((\<lambda>x. ereal (real_of_ereal (f x))) \<longlongrightarrow> 0) F" by (simp add: zero_ereal_def)
+qed
+
 lemma tendsto_explicit:
   "f \<longlonglongrightarrow> f0 \<longleftrightarrow> (\<forall>S. open S \<longrightarrow> f0 \<in> S \<longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> S))"
   unfolding tendsto_def eventually_sequentially by auto
@@ -3678,6 +3730,126 @@
     by auto
 qed
 
+lemma continuous_uminus_ereal [continuous_intros]: "continuous_on (A :: ereal set) uminus"
+  unfolding continuous_on_def
+  by (intro ballI tendsto_uminus_ereal[of "\<lambda>x. x::ereal"]) simp
+
+lemma ereal_uminus_atMost [simp]: "uminus ` {..(a::ereal)} = {-a..}"
+proof (intro equalityI subsetI)
+  fix x :: ereal assume "x \<in> {-a..}"
+  hence "-(-x) \<in> uminus ` {..a}" by (intro imageI) (simp add: ereal_uminus_le_reorder)
+  thus "x \<in> uminus ` {..a}" by simp
+qed auto
+
+lemma continuous_on_inverse_ereal [continuous_intros]: 
+  "continuous_on {0::ereal ..} inverse"
+  unfolding continuous_on_def
+proof clarsimp
+  fix x :: ereal assume "0 \<le> x"
+  moreover have "at 0 within {0 ..} = at_right (0::ereal)"
+    by (auto simp: filter_eq_iff eventually_at_filter le_less)
+  moreover have "at x within {0 ..} = at x" if "0 < x"
+    using that by (intro at_within_nhd[of _ "{0<..}"]) auto
+  ultimately show "(inverse \<longlongrightarrow> inverse x) (at x within {0..})"
+    by (auto simp: le_less inverse_ereal_tendsto_at_right_0 inverse_ereal_tendsto_pos)
+qed
+  
+lemma continuous_inverse_ereal_nonpos: "continuous_on ({..<0} :: ereal set) inverse"
+proof (subst continuous_on_cong[OF refl]) 
+  have "continuous_on {(0::ereal)<..} inverse"
+    by (rule continuous_on_subset[OF continuous_on_inverse_ereal]) auto
+  thus "continuous_on {..<(0::ereal)} (uminus \<circ> inverse \<circ> uminus)"
+    by (intro continuous_intros) simp_all
+qed simp
+
+lemma tendsto_inverse_ereal:
+  assumes "(f \<longlongrightarrow> (c :: ereal)) F"
+  assumes "eventually (\<lambda>x. f x \<ge> 0) F"
+  shows   "((\<lambda>x. inverse (f x)) \<longlongrightarrow> inverse c) F"
+  by (cases "F = bot") 
+     (auto intro!: tendsto_le_const[of F] assms 
+                   continuous_on_tendsto_compose[OF continuous_on_inverse_ereal])
+
+
+subsubsection \<open>liminf and limsup\<close>
+
+lemma Limsup_ereal_mult_right: 
+  assumes "F \<noteq> bot" "(c::real) \<ge> 0"
+  shows   "Limsup F (\<lambda>n. f n * ereal c) = Limsup F f * ereal c"
+proof (rule Limsup_compose_continuous_mono)
+  from assms show "continuous_on UNIV (\<lambda>a. a * ereal c)"
+    using tendsto_cmult_ereal[of "ereal c" "\<lambda>x. x" ]
+    by (force simp: continuous_on_def mult_ac)
+qed (insert assms, auto simp: mono_def ereal_mult_right_mono)
+
+lemma Liminf_ereal_mult_right: 
+  assumes "F \<noteq> bot" "(c::real) \<ge> 0"
+  shows   "Liminf F (\<lambda>n. f n * ereal c) = Liminf F f * ereal c"
+proof (rule Liminf_compose_continuous_mono)
+  from assms show "continuous_on UNIV (\<lambda>a. a * ereal c)"
+    using tendsto_cmult_ereal[of "ereal c" "\<lambda>x. x" ]
+    by (force simp: continuous_on_def mult_ac)
+qed (insert assms, auto simp: mono_def ereal_mult_right_mono)
+
+lemma Limsup_ereal_mult_left: 
+  assumes "F \<noteq> bot" "(c::real) \<ge> 0"
+  shows   "Limsup F (\<lambda>n. ereal c * f n) = ereal c * Limsup F f"
+  using Limsup_ereal_mult_right[OF assms] by (subst (1 2) mult.commute)
+
+lemma limsup_ereal_mult_right: 
+  "(c::real) \<ge> 0 \<Longrightarrow> limsup (\<lambda>n. f n * ereal c) = limsup f * ereal c"
+  by (rule Limsup_ereal_mult_right) simp_all
+
+lemma limsup_ereal_mult_left: 
+  "(c::real) \<ge> 0 \<Longrightarrow> limsup (\<lambda>n. ereal c * f n) = ereal c * limsup f"
+  by (subst (1 2) mult.commute, rule limsup_ereal_mult_right) simp_all
+
+lemma Limsup_add_ereal_right: 
+  "F \<noteq> bot \<Longrightarrow> abs c \<noteq> \<infinity> \<Longrightarrow> Limsup F (\<lambda>n. g n + (c :: ereal)) = Limsup F g + c"
+  by (rule Limsup_compose_continuous_mono) (auto simp: mono_def ereal_add_mono continuous_on_def)
+
+lemma Limsup_add_ereal_left: 
+  "F \<noteq> bot \<Longrightarrow> abs c \<noteq> \<infinity> \<Longrightarrow> Limsup F (\<lambda>n. (c :: ereal) + g n) = c + Limsup F g"
+  by (subst (1 2) add.commute) (rule Limsup_add_ereal_right)
+
+lemma Liminf_add_ereal_right: 
+  "F \<noteq> bot \<Longrightarrow> abs c \<noteq> \<infinity> \<Longrightarrow> Liminf F (\<lambda>n. g n + (c :: ereal)) = Liminf F g + c"
+  by (rule Liminf_compose_continuous_mono) (auto simp: mono_def ereal_add_mono continuous_on_def)
+
+lemma Liminf_add_ereal_left: 
+  "F \<noteq> bot \<Longrightarrow> abs c \<noteq> \<infinity> \<Longrightarrow> Liminf F (\<lambda>n. (c :: ereal) + g n) = c + Liminf F g"
+  by (subst (1 2) add.commute) (rule Liminf_add_ereal_right)
+
+lemma 
+  assumes "F \<noteq> bot"
+  assumes nonneg: "eventually (\<lambda>x. f x \<ge> (0::ereal)) F"
+  shows   Liminf_inverse_ereal: "Liminf F (\<lambda>x. inverse (f x)) = inverse (Limsup F f)"
+  and     Limsup_inverse_ereal: "Limsup F (\<lambda>x. inverse (f x)) = inverse (Liminf F f)"
+proof -
+  def inv \<equiv> "\<lambda>x. if x \<le> 0 then \<infinity> else inverse x :: ereal"
+  have "continuous_on ({..0} \<union> {0..}) inv" unfolding inv_def
+    by (intro continuous_on_If) (auto intro!: continuous_intros)
+  also have "{..0} \<union> {0..} = (UNIV :: ereal set)" by auto
+  finally have cont: "continuous_on UNIV inv" .
+  have antimono: "antimono inv" unfolding inv_def antimono_def
+    by (auto intro!: ereal_inverse_antimono)
+  
+  have "Liminf F (\<lambda>x. inverse (f x)) = Liminf F (\<lambda>x. inv (f x))" using nonneg
+    by (auto intro!: Liminf_eq elim!: eventually_mono simp: inv_def)
+  also have "... = inv (Limsup F f)"
+    by (simp add: assms(1) Liminf_compose_continuous_antimono[OF cont antimono])
+  also from assms have "Limsup F f \<ge> 0" by (intro le_Limsup) simp_all
+  hence "inv (Limsup F f) = inverse (Limsup F f)" by (simp add: inv_def)
+  finally show "Liminf F (\<lambda>x. inverse (f x)) = inverse (Limsup F f)" .
+
+  have "Limsup F (\<lambda>x. inverse (f x)) = Limsup F (\<lambda>x. inv (f x))" using nonneg
+    by (auto intro!: Limsup_eq elim!: eventually_mono simp: inv_def)
+  also have "... = inv (Liminf F f)"
+    by (simp add: assms(1) Limsup_compose_continuous_antimono[OF cont antimono])
+  also from assms have "Liminf F f \<ge> 0" by (intro Liminf_bounded) simp_all
+  hence "inv (Liminf F f) = inverse (Liminf F f)" by (simp add: inv_def)
+  finally show "Limsup F (\<lambda>x. inverse (f x)) = inverse (Liminf F f)" .
+qed
 
 subsubsection \<open>Tests for code generator\<close>
 
--- a/src/HOL/Library/Liminf_Limsup.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Library/Liminf_Limsup.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -1,5 +1,6 @@
 (*  Title:      HOL/Library/Liminf_Limsup.thy
     Author:     Johannes Hölzl, TU München
+    Author:     Manuel Eberl, TU München
 *)
 
 section \<open>Liminf and Limsup on complete lattices\<close>
@@ -189,6 +190,42 @@
     unfolding Liminf_def le_SUP_iff by auto
 qed
 
+lemma Limsup_le_iff:
+  fixes X :: "_ \<Rightarrow> _ :: complete_linorder"
+  shows "C \<ge> Limsup F X \<longleftrightarrow> (\<forall>y>C. eventually (\<lambda>x. y > X x) F)"
+proof -
+  { fix y P assume "eventually P F" "y > SUPREMUM (Collect P) X"
+    then have "eventually (\<lambda>x. y > X x) F"
+      by (auto elim!: eventually_mono dest: SUP_lessD) }
+  moreover
+  { fix y P assume "y > C" and y: "\<forall>y>C. eventually (\<lambda>x. y > X x) F"
+    have "\<exists>P. eventually P F \<and> y > SUPREMUM (Collect P) X"
+    proof (cases "\<exists>z. C < z \<and> z < y")
+      case True
+      then obtain z where z: "C < z \<and> z < y" ..
+      moreover from z have "z \<ge> SUPREMUM {x. z > X x} X"
+        by (auto intro!: SUP_least)
+      ultimately show ?thesis
+        using y by (intro exI[of _ "\<lambda>x. z > X x"]) auto
+    next
+      case False
+      then have "C \<ge> SUPREMUM {x. y > X x} X"
+        by (intro SUP_least) (auto simp: not_less)
+      with \<open>y > C\<close> show ?thesis
+        using y by (intro exI[of _ "\<lambda>x. y > X x"]) auto
+    qed }
+  ultimately show ?thesis
+    unfolding Limsup_def INF_le_iff by auto
+qed
+
+lemma less_LiminfD:
+  "y < Liminf F (f :: _ \<Rightarrow> 'a :: complete_linorder) \<Longrightarrow> eventually (\<lambda>x. f x > y) F"
+  using le_Liminf_iff[of "Liminf F f" F f] by simp
+
+lemma Limsup_lessD:
+  "y > Limsup F (f :: _ \<Rightarrow> 'a :: complete_linorder) \<Longrightarrow> eventually (\<lambda>x. f x < y) F"
+  using Limsup_le_iff[of F f "Limsup F f"] by simp
+
 lemma lim_imp_Liminf:
   fixes f :: "'a \<Rightarrow> _ :: {complete_linorder,linorder_topology}"
   assumes ntriv: "\<not> trivial_limit F"
@@ -327,6 +364,56 @@
   unfolding continuous_on_eq_continuous_within
   by (auto simp: continuous_within intro: tendsto_within_subset)
 
+lemma Liminf_compose_continuous_mono:
+  fixes f :: "'a::{complete_linorder, linorder_topology} \<Rightarrow> 'b::{complete_linorder, linorder_topology}"
+  assumes c: "continuous_on UNIV f" and am: "mono f" and F: "F \<noteq> bot"
+  shows "Liminf F (\<lambda>n. f (g n)) = f (Liminf F g)"
+proof -
+  { fix P assume "eventually P F"
+    have "\<exists>x. P x"
+    proof (rule ccontr)
+      assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
+        by auto
+      with \<open>eventually P F\<close> F show False
+        by auto
+    qed }
+  note * = this
+
+  have "f (Liminf F g) = (SUP P : {P. eventually P F}. f (Inf (g ` Collect P)))"
+    unfolding Liminf_def SUP_def
+    by (subst continuous_at_Sup_mono[OF am continuous_on_imp_continuous_within[OF c]])
+       (auto intro: eventually_True)
+  also have "\<dots> = (SUP P : {P. eventually P F}. INFIMUM (g ` Collect P) f)"
+    by (intro SUP_cong refl continuous_at_Inf_mono[OF am continuous_on_imp_continuous_within[OF c]])
+       (auto dest!: eventually_happens simp: F)
+  finally show ?thesis by (auto simp: Liminf_def)
+qed
+
+lemma Limsup_compose_continuous_mono:
+  fixes f :: "'a::{complete_linorder, linorder_topology} \<Rightarrow> 'b::{complete_linorder, linorder_topology}"
+  assumes c: "continuous_on UNIV f" and am: "mono f" and F: "F \<noteq> bot"
+  shows "Limsup F (\<lambda>n. f (g n)) = f (Limsup F g)"
+proof -
+  { fix P assume "eventually P F"
+    have "\<exists>x. P x"
+    proof (rule ccontr)
+      assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
+        by auto
+      with \<open>eventually P F\<close> F show False
+        by auto
+    qed }
+  note * = this
+
+  have "f (Limsup F g) = (INF P : {P. eventually P F}. f (Sup (g ` Collect P)))"
+    unfolding Limsup_def INF_def
+    by (subst continuous_at_Inf_mono[OF am continuous_on_imp_continuous_within[OF c]])
+       (auto intro: eventually_True)
+  also have "\<dots> = (INF P : {P. eventually P F}. SUPREMUM (g ` Collect P) f)"
+    by (intro INF_cong refl continuous_at_Sup_mono[OF am continuous_on_imp_continuous_within[OF c]])
+       (auto dest!: eventually_happens simp: F)
+  finally show ?thesis by (auto simp: Limsup_def)
+qed
+
 lemma Liminf_compose_continuous_antimono:
   fixes f :: "'a::{complete_linorder,linorder_topology} \<Rightarrow> 'b::{complete_linorder,linorder_topology}"
   assumes c: "continuous_on UNIV f"
@@ -351,6 +438,34 @@
   finally show ?thesis
     by (auto simp: Liminf_def)
 qed
+
+lemma Limsup_compose_continuous_antimono:
+  fixes f :: "'a::{complete_linorder, linorder_topology} \<Rightarrow> 'b::{complete_linorder, linorder_topology}"
+  assumes c: "continuous_on UNIV f" and am: "antimono f" and F: "F \<noteq> bot"
+  shows "Limsup F (\<lambda>n. f (g n)) = f (Liminf F g)"
+proof -
+  { fix P assume "eventually P F"
+    have "\<exists>x. P x"
+    proof (rule ccontr)
+      assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
+        by auto
+      with \<open>eventually P F\<close> F show False
+        by auto
+    qed }
+  note * = this
+
+  have "f (Liminf F g) = (INF P : {P. eventually P F}. f (Inf (g ` Collect P)))"
+    unfolding Liminf_def SUP_def
+    by (subst continuous_at_Sup_antimono[OF am continuous_on_imp_continuous_within[OF c]])
+       (auto intro: eventually_True)
+  also have "\<dots> = (INF P : {P. eventually P F}. SUPREMUM (g ` Collect P) f)"
+    by (intro INF_cong refl continuous_at_Inf_antimono[OF am continuous_on_imp_continuous_within[OF c]])
+       (auto dest!: eventually_happens simp: F)
+  finally show ?thesis
+    by (auto simp: Limsup_def)
+qed
+
+
 subsection \<open>More Limits\<close>
 
 lemma convergent_limsup_cl:
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Nonpos_Ints.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -0,0 +1,141 @@
+(*
+  Title:    HOL/Library/Nonpos_Ints.thy
+  Author:   Manuel Eberl, TU München
+  
+  The set of non-positive integers on a ring. (in analogy to the set of non-negative
+  integers @{term "\<nat>"}) This is useful e.g. for the Gamma function.
+*)
+theory Nonpos_Ints
+imports Complex_Main
+begin
+
+subsection \<open>Non-positive integers\<close>
+
+definition nonpos_Ints ("\<int>\<^sub>\<le>\<^sub>0") where "\<int>\<^sub>\<le>\<^sub>0 = {of_int n |n. n \<le> 0}"
+
+lemma zero_in_nonpos_Ints [simp,intro]: "0 \<in> \<int>\<^sub>\<le>\<^sub>0"
+  unfolding nonpos_Ints_def by (auto intro!: exI[of _ "0::int"])
+
+lemma neg_one_in_nonpos_Ints [simp,intro]: "-1 \<in> \<int>\<^sub>\<le>\<^sub>0"
+  unfolding nonpos_Ints_def by (auto intro!: exI[of _ "-1::int"])
+
+lemma neg_numeral_in_nonpos_Ints [simp,intro]: "-numeral n \<in> \<int>\<^sub>\<le>\<^sub>0"
+  unfolding nonpos_Ints_def by (auto intro!: exI[of _ "-numeral n::int"])
+
+lemma one_notin_nonpos_Ints [simp]: "(1 :: 'a :: ring_char_0) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  by (auto simp: nonpos_Ints_def)
+
+lemma numeral_notin_nonpos_Ints [simp]: "(numeral n :: 'a :: ring_char_0) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  by (auto simp: nonpos_Ints_def)
+
+
+lemma minus_of_nat_in_nonpos_Ints [simp, intro]: "- of_nat n \<in> \<int>\<^sub>\<le>\<^sub>0"
+proof -
+  have "- of_nat n = of_int (-int n)" by simp
+  also have "-int n \<le> 0" by simp
+  hence "of_int (-int n) \<in> \<int>\<^sub>\<le>\<^sub>0" unfolding nonpos_Ints_def by blast
+  finally show ?thesis .
+qed
+
+lemma of_nat_in_nonpos_Ints_iff: "(of_nat n :: 'a :: {ring_1,ring_char_0}) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> n = 0"
+proof
+  assume "(of_nat n :: 'a) \<in> \<int>\<^sub>\<le>\<^sub>0"
+  then obtain m where "of_nat n = (of_int m :: 'a)" "m \<le> 0" by (auto simp: nonpos_Ints_def)
+  hence "(of_int m :: 'a) = of_nat n" by simp
+  also have "... = of_int (int n)" by simp
+  finally have "m = int n" by (subst (asm) of_int_eq_iff)
+  with `m \<le> 0` show "n = 0" by auto
+qed simp
+
+lemma nonpos_Ints_of_int: "n \<le> 0 \<Longrightarrow> of_int n \<in> \<int>\<^sub>\<le>\<^sub>0"
+  unfolding nonpos_Ints_def by blast
+
+lemma nonpos_IntsI: 
+  "x \<in> \<int> \<Longrightarrow> x \<le> 0 \<Longrightarrow> (x :: 'a :: linordered_idom) \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using assms unfolding nonpos_Ints_def Ints_def by auto
+
+lemma nonpos_Ints_subset_Ints: "\<int>\<^sub>\<le>\<^sub>0 \<subseteq> \<int>"
+  unfolding nonpos_Ints_def Ints_def by blast
+
+lemma nonpos_Ints_nonpos [dest]: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x \<le> (0 :: 'a :: linordered_idom)"
+  unfolding nonpos_Ints_def by auto
+
+lemma nonpos_Ints_Int [dest]: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x \<in> \<int>"
+  unfolding nonpos_Ints_def Ints_def by blast
+
+lemma nonpos_Ints_cases:
+  assumes "x \<in> \<int>\<^sub>\<le>\<^sub>0"
+  obtains n where "x = of_int n" "n \<le> 0"
+  using assms unfolding nonpos_Ints_def by (auto elim!: Ints_cases)
+
+lemma nonpos_Ints_cases':
+  assumes "x \<in> \<int>\<^sub>\<le>\<^sub>0"
+  obtains n where "x = -of_nat n"
+proof -
+  from assms obtain m where "x = of_int m" and m: "m \<le> 0" by (auto elim!: nonpos_Ints_cases)
+  hence "x = - of_int (-m)" by auto
+  also from m have "(of_int (-m) :: 'a) = of_nat (nat (-m))" by simp_all
+  finally show ?thesis by (rule that)
+qed
+
+lemma of_real_in_nonpos_Ints_iff: "(of_real x :: 'a :: real_algebra_1) \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0"
+proof
+  assume "of_real x \<in> (\<int>\<^sub>\<le>\<^sub>0 :: 'a set)"
+  then obtain n where "(of_real x :: 'a) = of_int n" "n \<le> 0" by (erule nonpos_Ints_cases)
+  note `of_real x = of_int n`
+  also have "of_int n = of_real (of_int n)" by (rule of_real_of_int_eq [symmetric])
+  finally have "x = of_int n" by (subst (asm) of_real_eq_iff)
+  with `n \<le> 0` show "x \<in> \<int>\<^sub>\<le>\<^sub>0" by (simp add: nonpos_Ints_of_int)
+qed (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int)
+
+lemma nonpos_Ints_altdef: "\<int>\<^sub>\<le>\<^sub>0 = {n \<in> \<int>. (n :: 'a :: linordered_idom) \<le> 0}"
+  by (auto intro!: nonpos_IntsI elim!: nonpos_Ints_cases)
+
+lemma uminus_in_Nats_iff: "-x \<in> \<nat> \<longleftrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0"
+proof
+  assume "-x \<in> \<nat>"
+  then obtain n where "n \<ge> 0" "-x = of_int n" by (auto simp: Nats_altdef1)
+  hence "-n \<le> 0" "x = of_int (-n)" by (simp_all add: eq_commute minus_equation_iff[of x])
+  thus "x \<in> \<int>\<^sub>\<le>\<^sub>0" unfolding nonpos_Ints_def by blast
+next
+  assume "x \<in> \<int>\<^sub>\<le>\<^sub>0"
+  then obtain n where "n \<le> 0" "x = of_int n" by (auto simp: nonpos_Ints_def)
+  hence "-n \<ge> 0" "-x = of_int (-n)" by (simp_all add: eq_commute minus_equation_iff[of x])
+  thus "-x \<in> \<nat>" unfolding Nats_altdef1 by blast
+qed
+
+lemma uminus_in_nonpos_Ints_iff: "-x \<in> \<int>\<^sub>\<le>\<^sub>0 \<longleftrightarrow> x \<in> \<nat>"
+  using uminus_in_Nats_iff[of "-x"] by simp
+
+lemma nonpos_Ints_mult: "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x * y \<in> \<nat>"
+  using Nats_mult[of "-x" "-y"] by (simp add: uminus_in_Nats_iff)
+
+lemma Nats_mult_nonpos_Ints: "x \<in> \<nat> \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x * y \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using Nats_mult[of x "-y"] by (simp add: uminus_in_Nats_iff)
+
+lemma nonpos_Ints_mult_Nats:
+  "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<nat> \<Longrightarrow> x * y \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using Nats_mult[of "-x" y] by (simp add: uminus_in_Nats_iff)
+
+lemma nonpos_Ints_add:
+  "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x + y \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using Nats_add[of "-x" "-y"] uminus_in_Nats_iff[of "y+x", simplified minus_add] 
+  by (simp add: uminus_in_Nats_iff add.commute)
+
+lemma nonpos_Ints_diff_Nats:
+  "x \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> y \<in> \<nat> \<Longrightarrow> x - y \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using Nats_add[of "-x" "y"] uminus_in_Nats_iff[of "x-y", simplified minus_add] 
+  by (simp add: uminus_in_Nats_iff add.commute)
+
+lemma Nats_diff_nonpos_Ints:
+  "x \<in> \<nat> \<Longrightarrow> y \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> x - y \<in> \<nat>"
+  using Nats_add[of "x" "-y"] by (simp add: uminus_in_Nats_iff add.commute)
+
+lemma plus_of_nat_eq_0_imp: "z + of_nat n = 0 \<Longrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
+proof -
+  assume "z + of_nat n = 0"
+  hence A: "z = - of_nat n" by (simp add: eq_neg_iff_add_eq_0)
+  show "z \<in> \<int>\<^sub>\<le>\<^sub>0" by (subst A) simp
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Library/Periodic_Fun.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -0,0 +1,130 @@
+(*
+  Title:    HOL/Library/Periodic_Fun.thy
+  Author:   Manuel Eberl, TU München
+  
+  A locale for periodic functions. The idea is that one proves $f(x + p) = f(x)$
+  for some period $p$ and gets derived results like $f(x - p) = f(x)$ and $f(x + 2p) = f(x)$
+  for free.
+*)
+theory Periodic_Fun
+imports Complex_Main
+begin
+
+text \<open>
+  @{term g} and @{term gm} are ``plus/minus k periods'' functions. 
+  @{term g1} and @{term gn1} are ``plus/minus one period'' functions.
+  This is useful e.g. if the period is one; the lemmas one gets are then 
+  @{term "f (x + 1) = f x"} instead of @{term "f (x + 1 * 1) = f x"} etc.
+\<close>
+locale periodic_fun = 
+  fixes f :: "('a :: {ring_1}) \<Rightarrow> 'b" and g gm :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" and g1 gn1 :: "'a \<Rightarrow> 'a"
+  assumes plus_1: "f (g1 x) = f x"
+  assumes periodic_arg_plus_0: "g x 0 = x"
+  assumes periodic_arg_plus_distrib: "g x (of_int (m + n)) = g (g x (of_int n)) (of_int m)"
+  assumes plus_1_eq: "g x 1 = g1 x" and minus_1_eq: "g x (-1) = gn1 x" 
+          and minus_eq: "g x (-y) = gm x y"
+begin
+
+lemma plus_of_nat: "f (g x (of_nat n)) = f x"
+  by (induction n) (insert periodic_arg_plus_distrib[of _ 1 "int n" for n], 
+                    simp_all add: plus_1 periodic_arg_plus_0 plus_1_eq)
+
+lemma minus_of_nat: "f (gm x (of_nat n)) = f x"
+proof -
+  have "f (g x (- of_nat n)) = f (g (g x (- of_nat n)) (of_nat n))"
+    by (rule plus_of_nat[symmetric])
+  also have "\<dots> = f (g (g x (of_int (- of_nat n))) (of_int (of_nat n)))" by simp
+  also have "\<dots> = f x" 
+    by (subst periodic_arg_plus_distrib [symmetric]) (simp add: periodic_arg_plus_0)
+  finally show ?thesis by (simp add: minus_eq)
+qed
+
+lemma plus_of_int: "f (g x (of_int n)) = f x"
+  by (induction n) (simp_all add: plus_of_nat minus_of_nat minus_eq del: of_nat_Suc)
+
+lemma minus_of_int: "f (gm x (of_int n)) = f x"
+  using plus_of_int[of x "of_int (-n)"] by (simp add: minus_eq)
+
+lemma plus_numeral: "f (g x (numeral n)) = f x"
+  by (subst of_nat_numeral[symmetric], subst plus_of_nat) (rule refl)
+
+lemma minus_numeral: "f (gm x (numeral n)) = f x"
+  by (subst of_nat_numeral[symmetric], subst minus_of_nat) (rule refl)
+
+lemma minus_1: "f (gn1 x) = f x"
+  using minus_of_nat[of x 1] by (simp add: minus_1_eq minus_eq[symmetric])
+
+lemmas periodic_simps = plus_of_nat minus_of_nat plus_of_int minus_of_int 
+                        plus_numeral minus_numeral plus_1 minus_1
+
+end
+
+
+text \<open>
+  Specialised case of the @{term periodic_fun} locale for periods that are not 1.
+  Gives lemmas @{term "f (x - period) = f x"} etc.
+\<close>
+locale periodic_fun_simple = 
+  fixes f :: "('a :: {ring_1}) \<Rightarrow> 'b" and period :: 'a
+  assumes plus_period: "f (x + period) = f x"
+begin
+sublocale periodic_fun f "\<lambda>z x. z + x * period" "\<lambda>z x. z - x * period" 
+  "\<lambda>z. z + period" "\<lambda>z. z - period"
+  by standard (simp_all add: ring_distribs plus_period)
+end
+
+
+text \<open>
+  Specialised case of the @{term periodic_fun} locale for period 1.
+  Gives lemmas @{term "f (x - 1) = f x"} etc.
+\<close>
+locale periodic_fun_simple' = 
+  fixes f :: "('a :: {ring_1}) \<Rightarrow> 'b"
+  assumes plus_period: "f (x + 1) = f x"
+begin
+sublocale periodic_fun f "\<lambda>z x. z + x" "\<lambda>z x. z - x" "\<lambda>z. z + 1" "\<lambda>z. z - 1"
+  by standard (simp_all add: ring_distribs plus_period)
+
+lemma of_nat: "f (of_nat n) = f 0" using plus_of_nat[of 0 n] by simp
+lemma uminus_of_nat: "f (-of_nat n) = f 0" using minus_of_nat[of 0 n] by simp
+lemma of_int: "f (of_int n) = f 0" using plus_of_int[of 0 n] by simp
+lemma uminus_of_int: "f (-of_int n) = f 0" using minus_of_int[of 0 n] by simp
+lemma of_numeral: "f (numeral n) = f 0" using plus_numeral[of 0 n] by simp
+lemma of_neg_numeral: "f (-numeral n) = f 0" using minus_numeral[of 0 n] by simp
+lemma of_1: "f 1 = f 0" using plus_of_nat[of 0 1] by simp
+lemma of_neg_1: "f (-1) = f 0" using minus_of_nat[of 0 1] by simp
+
+lemmas periodic_simps' = 
+  of_nat uminus_of_nat of_int uminus_of_int of_numeral of_neg_numeral of_1 of_neg_1
+
+end
+
+lemma sin_plus_pi: "sin ((z :: 'a :: {real_normed_field,banach}) + of_real pi) = - sin z"
+  by (simp add: sin_add)
+  
+lemma cos_plus_pi: "cos ((z :: 'a :: {real_normed_field,banach}) + of_real pi) = - cos z"
+  by (simp add: cos_add)
+
+interpretation sin: periodic_fun_simple sin "2 * of_real pi :: 'a :: {real_normed_field,banach}"
+proof
+  fix z :: 'a
+  have "sin (z + 2 * of_real pi) = sin (z + of_real pi + of_real pi)" by (simp add: ac_simps)
+  also have "\<dots> = sin z" by (simp only: sin_plus_pi) simp
+  finally show "sin (z + 2 * of_real pi) = sin z" .
+qed
+
+interpretation cos: periodic_fun_simple cos "2 * of_real pi :: 'a :: {real_normed_field,banach}"
+proof
+  fix z :: 'a
+  have "cos (z + 2 * of_real pi) = cos (z + of_real pi + of_real pi)" by (simp add: ac_simps)
+  also have "\<dots> = cos z" by (simp only: cos_plus_pi) simp
+  finally show "cos (z + 2 * of_real pi) = cos z" .
+qed
+
+interpretation tan: periodic_fun_simple tan "2 * of_real pi :: 'a :: {real_normed_field,banach}"
+  by standard (simp only: tan_def [abs_def] sin.plus_1 cos.plus_1)
+
+interpretation cot: periodic_fun_simple cot "2 * of_real pi :: 'a :: {real_normed_field,banach}"
+  by standard (simp only: cot_def [abs_def] sin.plus_1 cos.plus_1)
+  
+end
\ No newline at end of file
--- a/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Multivariate_Analysis/Complex_Transcendental.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -3,9 +3,32 @@
 text\<open>By John Harrison et al.  Ported from HOL Light by L C Paulson (2015)\<close>
 
 theory Complex_Transcendental
-imports Complex_Analysis_Basics
+imports 
+  Complex_Analysis_Basics
+  Summation
 begin
 
+(* TODO: Figure out what to do with Möbius transformations *)
+definition "moebius a b c d = (\<lambda>z. (a*z+b) / (c*z+d :: 'a :: field))"
+
+lemma moebius_inverse: 
+  assumes "a * d \<noteq> b * c" "c * z + d \<noteq> 0"
+  shows   "moebius d (-b) (-c) a (moebius a b c d z) = z"
+proof -
+  from assms have "(-c) * moebius a b c d z + a \<noteq> 0" unfolding moebius_def
+    by (simp add: field_simps)
+  with assms show ?thesis
+    unfolding moebius_def by (simp add: moebius_def divide_simps) (simp add: algebra_simps)?
+qed
+
+lemma moebius_inverse': 
+  assumes "a * d \<noteq> b * c" "c * z - a \<noteq> 0"
+  shows   "moebius a b c d (moebius d (-b) (-c) a z) = z"
+  using assms moebius_inverse[of d a "-b" "-c" z]
+  by (auto simp: algebra_simps)
+
+
+
 lemma cmod_add_real_less:
   assumes "Im z \<noteq> 0" "r\<noteq>0"
     shows "cmod (z + r) < cmod z + \<bar>r\<bar>"
@@ -1328,6 +1351,88 @@
     done
 qed
 
+lemma Ln_series:
+  fixes z :: complex
+  assumes "norm z < 1"
+  shows   "(\<lambda>n. (-1)^Suc n / of_nat n * z^n) sums ln (1 + z)" (is "(\<lambda>n. ?f n * z^n) sums _")
+proof -
+  let ?F = "\<lambda>z. \<Sum>n. ?f n * z^n" and ?F' = "\<lambda>z. \<Sum>n. diffs ?f n * z^n"
+  have r: "conv_radius ?f = 1"
+    by (intro conv_radius_ratio_limit_nonzero[of _ 1])
+       (simp_all add: norm_divide LIMSEQ_Suc_n_over_n del: of_nat_Suc)
+
+  have "\<exists>c. \<forall>z\<in>ball 0 1. ln (1 + z) - ?F z = c"
+  proof (rule has_field_derivative_zero_constant)
+    fix z :: complex assume z': "z \<in> ball 0 1"
+    hence z: "norm z < 1" by (simp add: dist_0_norm)
+    def t \<equiv> "of_real (1 + norm z) / 2 :: complex"
+    from z have t: "norm z < norm t" "norm t < 1" unfolding t_def
+      by (simp_all add: field_simps norm_divide del: of_real_add)
+
+    have "Re (-z) \<le> norm (-z)" by (rule complex_Re_le_cmod)
+    also from z have "... < 1" by simp
+    finally have "((\<lambda>z. ln (1 + z)) has_field_derivative inverse (1+z)) (at z)"
+      by (auto intro!: derivative_eq_intros)
+    moreover have "(?F has_field_derivative ?F' z) (at z)" using t r
+      by (intro termdiffs_strong[of _ t] summable_in_conv_radius) simp_all
+    ultimately have "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative (inverse (1 + z) - ?F' z)) 
+                       (at z within ball 0 1)"
+      by (intro derivative_intros) (simp_all add: at_within_open[OF z'])
+    also have "(\<lambda>n. of_nat n * ?f n * z ^ (n - Suc 0)) sums ?F' z" using t r
+      by (intro diffs_equiv termdiff_converges[OF t(1)] summable_in_conv_radius) simp_all
+    from sums_split_initial_segment[OF this, of 1]
+      have "(\<lambda>i. (-z) ^ i) sums ?F' z" by (simp add: power_minus[of z] del: of_nat_Suc)
+    hence "?F' z = inverse (1 + z)" using z by (simp add: sums_iff suminf_geometric divide_inverse)
+    also have "inverse (1 + z) - inverse (1 + z) = 0" by simp
+    finally show "((\<lambda>z. ln (1 + z) - ?F z) has_field_derivative 0) (at z within ball 0 1)" .
+  qed simp_all
+  then obtain c where c: "\<And>z. z \<in> ball 0 1 \<Longrightarrow> ln (1 + z) - ?F z = c" by blast
+  from c[of 0] have "c = 0" by (simp only: powser_zero) simp
+  with c[of z] assms have "ln (1 + z) = ?F z" by (simp add: dist_0_norm)
+  moreover have "summable (\<lambda>n. ?f n * z^n)" using assms r
+    by (intro summable_in_conv_radius) simp_all
+  ultimately show ?thesis by (simp add: sums_iff)
+qed
+
+lemma Ln_approx_linear:
+  fixes z :: complex
+  assumes "norm z < 1"
+  shows   "norm (ln (1 + z) - z) \<le> norm z^2 / (1 - norm z)"
+proof -
+  let ?f = "\<lambda>n. (-1)^Suc n / of_nat n"
+  from assms have "(\<lambda>n. ?f n * z^n) sums ln (1 + z)" using Ln_series by simp
+  moreover have "(\<lambda>n. (if n = 1 then 1 else 0) * z^n) sums z" using powser_sums_if[of 1] by simp
+  ultimately have "(\<lambda>n. (?f n - (if n = 1 then 1 else 0)) * z^n) sums (ln (1 + z) - z)"
+    by (subst left_diff_distrib, intro sums_diff) simp_all
+  from sums_split_initial_segment[OF this, of "Suc 1"]
+    have "(\<lambda>i. (-(z^2)) * inverse (2 + of_nat i) * (- z)^i) sums (Ln (1 + z) - z)"
+    by (simp add: power2_eq_square mult_ac power_minus[of z] divide_inverse)
+  hence "(Ln (1 + z) - z) = (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i)"
+    by (simp add: sums_iff)
+  also have A: "summable (\<lambda>n. norm z^2 * (inverse (real_of_nat (Suc (Suc n))) * cmod z ^ n))"
+    by (rule summable_mult, rule summable_comparison_test_ev[OF _ summable_geometric[of "norm z"]])
+       (auto simp: assms field_simps intro!: always_eventually)
+  hence "norm (\<Sum>i. (-(z^2)) * inverse (of_nat (i+2)) * (-z)^i) \<le> 
+             (\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i))"
+    by (intro summable_norm)
+       (auto simp: norm_power norm_inverse norm_mult mult_ac simp del: of_nat_add of_nat_Suc)
+  also have "norm ((-z)^2 * (-z)^i) * inverse (of_nat (i+2)) \<le> norm ((-z)^2 * (-z)^i) * 1" for i
+    by (intro mult_left_mono) (simp_all add: divide_simps)
+  hence "(\<Sum>i. norm (-(z^2) * inverse (of_nat (i+2)) * (-z)^i)) \<le> 
+           (\<Sum>i. norm (-(z^2) * (-z)^i))" using A assms
+    apply (simp_all only: norm_power norm_inverse norm_divide norm_mult)
+    apply (intro suminf_le summable_mult summable_geometric)
+    apply (auto simp: norm_power field_simps simp del: of_nat_add of_nat_Suc)
+    done
+  also have "... = norm z^2 * (\<Sum>i. norm z^i)" using assms
+    by (subst suminf_mult [symmetric]) (auto intro!: summable_geometric simp: norm_mult norm_power)
+  also have "(\<Sum>i. norm z^i) = inverse (1 - norm z)" using assms
+    by (subst suminf_geometric) (simp_all add: divide_inverse)
+  also have "norm z^2 * ... = norm z^2 / (1 - norm z)" by (simp add: divide_inverse)
+  finally show ?thesis .
+qed
+
+
 text\<open>Relation between Arg and arctangent in upper halfplane\<close>
 lemma Arg_arctan_upperhalf:
   assumes "0 < Im z"
@@ -1806,6 +1911,23 @@
 definition Arctan :: "complex \<Rightarrow> complex" where
     "Arctan \<equiv> \<lambda>z. (\<i>/2) * Ln((1 - \<i>*z) / (1 + \<i>*z))"
 
+lemma Arctan_def_moebius: "Arctan z = \<i>/2 * Ln (moebius (-\<i>) 1 \<i> 1 z)"
+  by (simp add: Arctan_def moebius_def add_ac)
+
+lemma Ln_conv_Arctan:
+  assumes "z \<noteq> -1"
+  shows   "Ln z = -2*\<i> * Arctan (moebius 1 (- 1) (- \<i>) (- \<i>) z)"
+proof -
+  have "Arctan (moebius 1 (- 1) (- \<i>) (- \<i>) z) =
+             \<i>/2 * Ln (moebius (- \<i>) 1 \<i> 1 (moebius 1 (- 1) (- \<i>) (- \<i>) z))"
+    by (simp add: Arctan_def_moebius)
+  also from assms have "\<i> * z \<noteq> \<i> * (-1)" by (subst mult_left_cancel) simp
+  hence "\<i> * z - -\<i> \<noteq> 0" by (simp add: eq_neg_iff_add_eq_0)
+  from moebius_inverse'[OF _ this, of 1 1]
+    have "moebius (- \<i>) 1 \<i> 1 (moebius 1 (- 1) (- \<i>) (- \<i>) z) = z" by simp
+  finally show ?thesis by (simp add: field_simps)
+qed
+
 lemma Arctan_0 [simp]: "Arctan 0 = 0"
   by (simp add: Arctan_def)
 
@@ -1930,6 +2052,108 @@
     "(\<And>z. z \<in> s \<Longrightarrow> Re z = 0 \<Longrightarrow> \<bar>Im z\<bar> < 1) \<Longrightarrow> Arctan holomorphic_on s"
   by (simp add: complex_differentiable_within_Arctan holomorphic_on_def)
 
+lemma Arctan_series:
+  assumes z: "norm (z :: complex) < 1"
+  defines "g \<equiv> \<lambda>n. if odd n then -\<i>*\<i>^n / n else 0"
+  defines "h \<equiv> \<lambda>z n. (-1)^n / of_nat (2*n+1) * (z::complex)^(2*n+1)"
+  shows   "(\<lambda>n. g n * z^n) sums Arctan z"
+  and     "h z sums Arctan z"
+proof -
+  def G \<equiv> "\<lambda>z. (\<Sum>n. g n * z^n)"
+  have summable: "summable (\<lambda>n. g n * u^n)" if "norm u < 1" for u
+  proof (cases "u = 0")
+    assume u: "u \<noteq> 0"
+    have "(\<lambda>n. ereal (norm (h u n) / norm (h u (Suc n)))) = (\<lambda>n. ereal (inverse (norm u)^2) * 
+              ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n)))))"
+    proof
+      fix n
+      have "ereal (norm (h u n) / norm (h u (Suc n))) = 
+             ereal (inverse (norm u)^2) * ereal ((of_nat (2*Suc n+1) / of_nat (Suc n)) / 
+                 (of_nat (2*Suc n-1) / of_nat (Suc n)))"
+      by (simp add: h_def norm_mult norm_power norm_divide divide_simps 
+                    power2_eq_square eval_nat_numeral del: of_nat_add of_nat_Suc)
+      also have "of_nat (2*Suc n+1) / of_nat (Suc n) = (2::real) + inverse (real (Suc n))"
+        by (auto simp: divide_simps simp del: of_nat_Suc) simp_all?
+      also have "of_nat (2*Suc n-1) / of_nat (Suc n) = (2::real) - inverse (real (Suc n))"
+        by (auto simp: divide_simps simp del: of_nat_Suc) simp_all?      
+      finally show "ereal (norm (h u n) / norm (h u (Suc n))) = ereal (inverse (norm u)^2) * 
+              ereal ((2 + inverse (real (Suc n))) / (2 - inverse (real (Suc n))))" .
+    qed
+    also have "\<dots> \<longlonglongrightarrow> ereal (inverse (norm u)^2) * ereal ((2 + 0) / (2 - 0))"
+      by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) simp_all
+    finally have "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) = inverse (norm u)^2"
+      by (intro lim_imp_Liminf) simp_all
+    moreover from power_strict_mono[OF that, of 2] u have "inverse (norm u)^2 > 1"
+      by (simp add: divide_simps)
+    ultimately have A: "liminf (\<lambda>n. ereal (cmod (h u n) / cmod (h u (Suc n)))) > 1" by simp
+    from u have "summable (h u)"
+      by (intro summable_norm_cancel[OF ratio_test_convergence[OF _ A]])
+         (auto simp: h_def norm_divide norm_mult norm_power simp del: of_nat_Suc 
+               intro!: mult_pos_pos divide_pos_pos always_eventually)
+    thus "summable (\<lambda>n. g n * u^n)"
+      by (subst summable_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
+         (auto simp: power_mult subseq_def g_def h_def elim!: oddE)
+  qed (simp add: h_def)
+
+  have "\<exists>c. \<forall>u\<in>ball 0 1. Arctan u - G u = c"
+  proof (rule has_field_derivative_zero_constant)
+    fix u :: complex assume "u \<in> ball 0 1"
+    hence u: "norm u < 1" by (simp add: dist_0_norm)
+    def K \<equiv> "(norm u + 1) / 2"
+    from u and abs_Im_le_cmod[of u] have Im_u: "\<bar>Im u\<bar> < 1" by linarith
+    from u have K: "0 \<le> K" "norm u < K" "K < 1" by (simp_all add: K_def)
+    hence "(G has_field_derivative (\<Sum>n. diffs g n * u ^ n)) (at u)" unfolding G_def
+      by (intro termdiffs_strong[of _ "of_real K"] summable) simp_all
+    also have "(\<lambda>n. diffs g n * u^n) = (\<lambda>n. if even n then (\<i>*u)^n else 0)"
+      by (intro ext) (simp_all del: of_nat_Suc add: g_def diffs_def power_mult_distrib)
+    also have "suminf \<dots> = (\<Sum>n. (-(u^2))^n)"
+      by (subst suminf_mono_reindex[of "\<lambda>n. 2*n", symmetric]) 
+         (auto elim!: evenE simp: subseq_def power_mult power_mult_distrib)
+    also from u have "norm u^2 < 1^2" by (intro power_strict_mono) simp_all
+    hence "(\<Sum>n. (-(u^2))^n) = inverse (1 + u^2)" 
+      by (subst suminf_geometric) (simp_all add: norm_power inverse_eq_divide)
+    finally have "(G has_field_derivative inverse (1 + u\<^sup>2)) (at u)" .
+    from DERIV_diff[OF has_field_derivative_Arctan this] Im_u u
+      show "((\<lambda>u. Arctan u - G u) has_field_derivative 0) (at u within ball 0 1)"
+      by (simp_all add: dist_0_norm at_within_open[OF _ open_ball])
+  qed simp_all
+  then obtain c where c: "\<And>u. norm u < 1 \<Longrightarrow> Arctan u - G u = c" by (auto simp: dist_0_norm)
+  from this[of 0] have "c = 0" by (simp add: G_def g_def powser_zero)
+  with c z have "Arctan z = G z" by simp
+  with summable[OF z] show "(\<lambda>n. g n * z^n) sums Arctan z" unfolding G_def by (simp add: sums_iff)
+  thus "h z sums Arctan z" by (subst (asm) sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
+                              (auto elim!: oddE simp: subseq_def power_mult g_def h_def)
+qed
+
+text \<open>A quickly-converging series for the logarithm, based on the arctangent.\<close>
+lemma ln_series_quadratic:
+  assumes x: "x > (0::real)"
+  shows "(\<lambda>n. (2*((x - 1) / (x + 1)) ^ (2*n+1) / of_nat (2*n+1))) sums ln x"
+proof -
+  def y \<equiv> "of_real ((x-1)/(x+1)) :: complex"
+  from x have x': "complex_of_real x \<noteq> of_real (-1)"  by (subst of_real_eq_iff) auto
+  from x have "\<bar>x - 1\<bar> < \<bar>x + 1\<bar>" by linarith
+  hence "norm (complex_of_real (x - 1) / complex_of_real (x + 1)) < 1"
+    by (simp add: norm_divide del: of_real_add of_real_diff)
+  hence "norm (\<i> * y) < 1" unfolding y_def by (subst norm_mult) simp
+  hence "(\<lambda>n. (-2*\<i>) * ((-1)^n / of_nat (2*n+1) * (\<i>*y)^(2*n+1))) sums ((-2*\<i>) * Arctan (\<i>*y))"
+    by (intro Arctan_series sums_mult) simp_all
+  also have "(\<lambda>n. (-2*\<i>) * ((-1)^n / of_nat (2*n+1) * (\<i>*y)^(2*n+1))) = 
+                 (\<lambda>n. (-2*\<i>) * ((-1)^n * (\<i>*y*(-y\<^sup>2)^n)/of_nat (2*n+1)))"
+    by (intro ext) (simp_all add: power_mult power_mult_distrib)
+  also have "\<dots> = (\<lambda>n. 2*y* ((-1) * (-y\<^sup>2))^n/of_nat (2*n+1))"
+    by (intro ext, subst power_mult_distrib) (simp add: algebra_simps power_mult)
+  also have "\<dots> = (\<lambda>n. 2*y^(2*n+1) / of_nat (2*n+1))" 
+    by (subst power_add, subst power_mult) (simp add: mult_ac)
+  also have "\<dots> = (\<lambda>n. of_real (2*((x-1)/(x+1))^(2*n+1) / of_nat (2*n+1)))"
+    by (intro ext) (simp add: y_def)
+  also have "\<i> * y = (of_real x - 1) / (-\<i> * (of_real x + 1))"
+    by (subst divide_divide_eq_left [symmetric]) (simp add: y_def)
+  also have "\<dots> = moebius 1 (-1) (-\<i>) (-\<i>) (of_real x)" by (simp add: moebius_def algebra_simps)
+  also from x' have "-2*\<i>*Arctan \<dots> = Ln (of_real x)" by (intro Ln_conv_Arctan [symmetric]) simp_all
+  also from x have "\<dots> = ln x" by (rule Ln_of_real)
+  finally show ?thesis by (subst (asm) sums_of_real_iff)
+qed
 
 subsection \<open>Real arctangent\<close>
 
--- a/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Multivariate_Analysis/Extended_Real_Limits.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -316,46 +316,6 @@
   apply (metis INF_absorb centre_in_ball)
   done
 
-lemma continuous_on_inverse_ereal: "continuous_on {0::ereal ..} inverse"
-  unfolding continuous_on_def
-proof clarsimp
-  fix x :: ereal assume "0 \<le> x"
-  moreover have "at 0 within {0 ..} = at_right (0::ereal)"
-    by (auto simp: filter_eq_iff eventually_at_filter le_less)
-  moreover have "0 < x \<Longrightarrow> at x within {0 ..} = at x"
-    using at_within_interior[of x "{0 ..}"] by (simp add: interior_Ici[of "- \<infinity>"])
-  ultimately show "(inverse \<longlongrightarrow> inverse x) (at x within {0..})"
-    by (auto simp: le_less inverse_ereal_tendsto_at_right_0 inverse_ereal_tendsto_pos)
-qed
-
-
-lemma Liminf_inverse_ereal:
-  assumes nneg: "\<forall>\<^sub>F x in F. f x \<ge> (0 :: ereal)" and "F \<noteq> bot"
-  shows "Liminf F (\<lambda>n. inverse (f n)) = inverse (Limsup F f)"
-proof -
-  def I \<equiv> "\<lambda>x::ereal. if x \<le> 0 then \<infinity> else inverse x"
-  have "Liminf F (\<lambda>n. I (f n)) = I (Limsup F f)"
-  proof (rule Liminf_compose_continuous_antimono)
-    have "continuous_on ({.. 0} \<union> {0 ..}) I"
-      unfolding I_def by (intro continuous_on_cases) (auto intro: continuous_on_const continuous_on_inverse_ereal)
-    also have "{.. 0} \<union> {0::ereal ..} = UNIV"
-      by auto
-    finally show "continuous_on UNIV I" .
-    show "antimono I"
-      unfolding antimono_def I_def by (auto intro: ereal_inverse_antimono)
-  qed fact
-  also have "Liminf F (\<lambda>n. I (f n)) = Liminf F (\<lambda>n. inverse (f n))"
-  proof (rule Liminf_eq)
-    show "\<forall>\<^sub>F x in F. I (f x) = inverse (f x)"
-      using nneg by eventually_elim (auto simp: I_def)
-  qed
-  also have "0 \<le> Limsup F f"
-    by (intro le_Limsup) fact+
-  then have "I (Limsup F f) = inverse (Limsup F f)"
-    by (auto simp: I_def)
-  finally show ?thesis .
-qed
-
 subsection \<open>monoset\<close>
 
 definition (in order) mono_set:
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Gamma.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -0,0 +1,2640 @@
+(*
+  Title:    HOL/Multivariate_Analysis/Gamma.thy
+  Author:   Manuel Eberl, TU München
+  
+  Several equivalent definitions of the Gamma function and its 
+  most important properties. Also contains the definition and some properties
+  of the log-Gamma function and the Digamma function and the other Polygamma functions.
+  
+  Based on the Gamma function, we also prove the Weierstraß product form of the
+  sin function and, based on this, the solution of the Basel problem (the 
+  sum over all @{term "1 / (n::nat)^2"}.
+*)
+theory Gamma
+imports 
+  Complex_Transcendental
+  Summation
+  Harmonic_Numbers
+  "~~/src/HOL/Library/Nonpos_Ints"
+  "~~/src/HOL/Library/Periodic_Fun"
+begin
+
+lemma pochhammer_eq_0_imp_nonpos_Int: 
+  "pochhammer (x::'a::field_char_0) n = 0 \<Longrightarrow> x \<in> \<int>\<^sub>\<le>\<^sub>0"
+  by (auto simp: pochhammer_eq_0_iff)
+  
+lemma closed_nonpos_Ints [simp]: "closed (\<int>\<^sub>\<le>\<^sub>0 :: 'a :: real_normed_algebra_1 set)"
+proof -
+  have "\<int>\<^sub>\<le>\<^sub>0 = (of_int ` {n. n \<le> 0} :: 'a set)" 
+    by (auto elim!: nonpos_Ints_cases intro!: nonpos_Ints_of_int)
+  also have "closed \<dots>" by (rule closed_of_int_image)
+  finally show ?thesis .
+qed
+
+lemma plus_one_in_nonpos_Ints_imp: "z + 1 \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using nonpos_Ints_diff_Nats[of "z+1" "1"] by simp_all
+
+lemma fraction_not_in_ints:
+  assumes "\<not>(n dvd m)" "n \<noteq> 0"
+  shows   "of_int m / of_int n \<notin> (\<int> :: 'a :: {division_ring,ring_char_0} set)"
+proof
+  assume "of_int m / (of_int n :: 'a) \<in> \<int>"
+  then obtain k where "of_int m / of_int n = (of_int k :: 'a)" by (elim Ints_cases)
+  with assms have "of_int m = (of_int (k * n) :: 'a)" by (auto simp add: divide_simps)
+  hence "m = k * n" by (subst (asm) of_int_eq_iff)
+  hence "n dvd m" by simp
+  with assms(1) show False by contradiction
+qed
+
+lemma not_in_Ints_imp_not_in_nonpos_Ints: "z \<notin> \<int> \<Longrightarrow> z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  by (auto simp: Ints_def nonpos_Ints_def)
+
+lemma double_in_nonpos_Ints_imp:
+  assumes "2 * (z :: 'a :: field_char_0) \<in> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<or> z + 1/2 \<in> \<int>\<^sub>\<le>\<^sub>0"
+proof-
+  from assms obtain k where k: "2 * z = - of_nat k" by (elim nonpos_Ints_cases')
+  thus ?thesis by (cases "even k") (auto elim!: evenE oddE simp: field_simps)
+qed
+
+
+lemma sin_series: "(\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
+proof -
+  from sin_converges[of z] have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z" .
+  also have "(\<lambda>n. sin_coeff n *\<^sub>R z^n) sums sin z \<longleftrightarrow> 
+                 (\<lambda>n. ((-1)^n / fact (2*n+1)) *\<^sub>R z^(2*n+1)) sums sin z"
+    by (subst sums_mono_reindex[of "\<lambda>n. 2*n+1", symmetric])
+       (auto simp: sin_coeff_def subseq_def ac_simps elim!: oddE)
+  finally show ?thesis .
+qed
+
+lemma cos_series: "(\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
+proof -
+  from cos_converges[of z] have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z" .
+  also have "(\<lambda>n. cos_coeff n *\<^sub>R z^n) sums cos z \<longleftrightarrow> 
+                 (\<lambda>n. ((-1)^n / fact (2*n)) *\<^sub>R z^(2*n)) sums cos z"
+    by (subst sums_mono_reindex[of "\<lambda>n. 2*n", symmetric])
+       (auto simp: cos_coeff_def subseq_def ac_simps elim!: evenE)
+  finally show ?thesis .
+qed
+
+lemma sin_z_over_z_series:
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes "z \<noteq> 0"
+  shows   "(\<lambda>n. (-1)^n / fact (2*n+1) * z^(2*n)) sums (sin z / z)"
+proof -
+  from sin_series[of z] have "(\<lambda>n. z * ((-1)^n / fact (2*n+1)) * z^(2*n)) sums sin z"
+    by (simp add: field_simps scaleR_conv_of_real)
+  from sums_mult[OF this, of "inverse z"] and assms show ?thesis
+    by (simp add: field_simps)
+qed
+
+lemma sin_z_over_z_series':
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes "z \<noteq> 0"
+  shows   "(\<lambda>n. sin_coeff (n+1) *\<^sub>R z^n) sums (sin z / z)"
+proof -
+  from sums_split_initial_segment[OF sin_converges[of z], of 1] 
+    have "(\<lambda>n. z * (sin_coeff (n+1) *\<^sub>R z ^ n)) sums sin z" by simp
+  from sums_mult[OF this, of "inverse z"] assms show ?thesis by (simp add: field_simps)
+qed
+
+lemma has_field_derivative_sin_z_over_z:
+  fixes A :: "'a :: {real_normed_field,banach} set"
+  shows "((\<lambda>z. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0 within A)"
+      (is "(?f has_field_derivative ?f') _")
+proof (rule has_field_derivative_at_within)
+  have "((\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n) 
+            has_field_derivative (\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n+1))) n * 0^n)) (at 0)"
+  proof (rule termdiffs_strong)
+    from summable_ignore_initial_segment[OF sums_summable[OF sin_converges[of "1::'a"]], of 1]
+      show "summable (\<lambda>n. of_real (sin_coeff (n+1)) * (1::'a)^n)" by (simp add: of_real_def)
+  qed simp
+  also have "(\<lambda>z::'a. \<Sum>n. of_real (sin_coeff (n+1)) * z^n) = ?f"
+  proof
+    fix z
+    show "(\<Sum>n. of_real (sin_coeff (n+1)) * z^n)  = ?f z"
+      by (cases "z = 0") (insert sin_z_over_z_series'[of z], 
+            simp_all add: scaleR_conv_of_real sums_iff powser_zero sin_coeff_def)
+  qed
+  also have "(\<Sum>n. diffs (\<lambda>n. of_real (sin_coeff (n + 1))) n * (0::'a) ^ n) = 
+                 diffs (\<lambda>n. of_real (sin_coeff (Suc n))) 0" by (simp add: powser_zero)
+  also have "\<dots> = 0" by (simp add: sin_coeff_def diffs_def)
+  finally show "((\<lambda>z::'a. if z = 0 then 1 else sin z / z) has_field_derivative 0) (at 0)" .
+qed
+
+ 
+lemma tendsto_complex_of_real_iff: "((\<lambda>x. complex_of_real (f x)) \<longlongrightarrow> of_real c) F = (f \<longlongrightarrow> c) F"
+  by (rule tendsto_of_real_iff)
+
+lemma round_Re_minimises_norm:
+  "norm ((z::complex) - of_int m) \<ge> norm (z - of_int (round (Re z)))"
+proof -
+  let ?n = "round (Re z)"
+  have "norm (z - of_int ?n) = sqrt ((Re z - of_int ?n)\<^sup>2 + (Im z)\<^sup>2)"
+    by (simp add: cmod_def)
+  also have "\<bar>Re z - of_int ?n\<bar> \<le> \<bar>Re z - of_int m\<bar>" by (rule round_diff_minimal)
+  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)"
+    by (intro real_sqrt_le_mono add_mono) (simp_all add: abs_le_square_iff)
+  also have "\<dots> = norm (z - of_int m)" by (simp add: cmod_def)
+  finally show ?thesis .
+qed
+
+lemma Re_pos_in_ball:
+  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
+  shows   "Re t > 0"
+proof -
+  have "Re (z - t) \<le> norm (z - t)" by (rule complex_Re_le_cmod)
+  also from assms have "\<dots> < Re z / 2" by (simp add: dist_complex_def)
+  finally show "Re t > 0" using assms by simp
+qed
+
+lemma no_nonpos_Int_in_ball_complex:
+  assumes "Re z > 0" "t \<in> ball z (Re z/2)"
+  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  using Re_pos_in_ball[OF assms] by (force elim!: nonpos_Ints_cases)
+
+lemma no_nonpos_Int_in_ball: 
+  assumes "t \<in> ball z (dist z (round (Re z)))"
+  shows   "t \<notin> \<int>\<^sub>\<le>\<^sub>0"
+proof
+  assume "t \<in> \<int>\<^sub>\<le>\<^sub>0"
+  then obtain n where "t = of_int n" by (auto elim!: nonpos_Ints_cases)
+  have "dist z (of_int n) \<le> dist z t + dist t (of_int n)" by (rule dist_triangle)
+  also from assms have "dist z t < dist z (round (Re z))" by simp
+  also have "\<dots> \<le> dist z (of_int n)"
+    using round_Re_minimises_norm[of z] by (simp add: dist_complex_def)
+  finally have "dist t (of_int n) > 0" by simp
+  with `t = of_int n` show False by simp
+qed
+
+lemma no_nonpos_Int_in_ball':
+  assumes "(z :: 'a :: {euclidean_space,real_normed_algebra_1}) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  obtains d where "d > 0" "\<And>t. t \<in> ball z d \<Longrightarrow> t \<notin> \<int>\<^sub>\<le>\<^sub>0"
+proof (rule that)
+  from assms show "setdist {z} \<int>\<^sub>\<le>\<^sub>0 > 0" by (subst setdist_gt_0_compact_closed) auto
+next
+  fix t assume "t \<in> ball z (setdist {z} \<int>\<^sub>\<le>\<^sub>0)"
+  thus "t \<notin> \<int>\<^sub>\<le>\<^sub>0" using setdist_le_dist[of z "{z}" t "\<int>\<^sub>\<le>\<^sub>0"] by force
+qed
+
+lemma Re_pos_or_Im_nz_in_ball:
+  assumes "Re z > 0 \<or> Im z \<noteq> 0" "t \<in> ball z (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
+  shows   "Re t > 0 \<or> Im t \<noteq> 0"
+using assms(1)
+proof (cases "Im z = 0")
+  assume A: "Im z = 0"
+  with assms(1) have "Re z > 0" by blast
+  with assms(2) A Re_pos_in_ball[of z t] show ?thesis by simp
+next
+  assume A: "Im z \<noteq> 0"
+  have "abs (Im z) - abs (Im t) \<le> abs (Im z - Im t)" by linarith
+  also have "\<dots> = abs (Im (z - t))" by simp
+  also have "\<dots> \<le> norm (z - t)" by (rule abs_Im_le_cmod)
+  also from A assms(2) have "\<dots> \<le> abs (Im z) / 2" by (simp add: dist_complex_def)
+  finally have "abs (Im t) > 0" using A by simp
+  thus ?thesis by force
+qed
+
+
+subsection \<open>Definitions\<close>
+
+text \<open>
+  We define the Gamma function by first defining its multiplicative inverse @{term "Gamma_inv"}. 
+  This is more convenient because @{term "Gamma_inv"} is entire, which makes proofs of its 
+  properties more convenient because one does not have to watch out for discontinuities.
+  (e.g. @{term "Gamma_inv"} fulfils @{term "rGamma z = z * rGamma (z + 1)"} everywhere,
+  whereas @{term "Gamma"} does not fulfil the analogous equation on the non-positive integers)
+  
+  We define the Gamma function (resp. its inverse) in the Euler form. This form has the advantage 
+  that it is a relatively simple limit that converges everywhere. The limit at the poles is 0 
+  (due to division by 0). The functional equation @{term "Gamma (z + 1) = z * Gamma z"} follows 
+  immediately from the definition.
+\<close>
+
+definition Gamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "Gamma_series z n = fact n * exp (z * of_real (ln (of_nat n))) / pochhammer z (n+1)"
+
+definition Gamma_series' :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "Gamma_series' z n = fact (n - 1) * exp (z * of_real (ln (of_nat n))) / pochhammer z n"
+
+definition rGamma_series :: "('a :: {banach,real_normed_field}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "rGamma_series z n = pochhammer z (n+1) / (fact n * exp (z * of_real (ln (of_nat n))))"
+
+lemma Gamma_series_altdef: "Gamma_series z n = inverse (rGamma_series z n)"
+  and rGamma_series_altdef: "rGamma_series z n = inverse (Gamma_series z n)"
+  unfolding Gamma_series_def rGamma_series_def by simp_all
+  
+lemma rGamma_series_minus_of_nat:
+  "eventually (\<lambda>n. rGamma_series (- of_nat k) n = 0) sequentially"
+  using eventually_ge_at_top[of k]
+  by eventually_elim (auto simp: rGamma_series_def pochhammer_of_nat_eq_0_iff)  
+
+lemma Gamma_series_minus_of_nat:
+  "eventually (\<lambda>n. Gamma_series (- of_nat k) n = 0) sequentially"
+  using eventually_ge_at_top[of k]
+  by eventually_elim (auto simp: Gamma_series_def pochhammer_of_nat_eq_0_iff)  
+
+lemma Gamma_series'_minus_of_nat:
+  "eventually (\<lambda>n. Gamma_series' (- of_nat k) n = 0) sequentially"
+  using eventually_gt_at_top[of k]
+  by eventually_elim (auto simp: Gamma_series'_def pochhammer_of_nat_eq_0_iff)
+  
+lemma rGamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma_series z \<longlonglongrightarrow> 0"
+  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule rGamma_series_minus_of_nat, simp)
+
+lemma Gamma_series_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series z \<longlonglongrightarrow> 0"
+  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series_minus_of_nat, simp)
+
+lemma Gamma_series'_nonpos_Ints_LIMSEQ: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma_series' z \<longlonglongrightarrow> 0"
+  by (elim nonpos_Ints_cases', hypsubst, subst tendsto_cong, rule Gamma_series'_minus_of_nat, simp)
+  
+lemma Gamma_series_Gamma_series': 
+  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(\<lambda>n. Gamma_series' z n / Gamma_series z n) \<longlonglongrightarrow> 1"
+proof (rule Lim_transform_eventually)
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    from n z have "Gamma_series' z n / Gamma_series z n = (z + of_nat n) / of_nat n"
+      by (cases n, simp)
+         (auto simp add: Gamma_series_def Gamma_series'_def pochhammer_rec'
+               dest: pochhammer_eq_0_imp_nonpos_Int plus_of_nat_eq_0_imp)
+    also from n have "\<dots> = z / of_nat n + 1" by (simp add: divide_simps)
+    finally show "z / of_nat n + 1 = Gamma_series' z n / Gamma_series z n" ..
+  qed
+  have "(\<lambda>x. z / of_nat x) \<longlonglongrightarrow> 0"
+    by (rule tendsto_norm_zero_cancel)
+       (insert tendsto_mult[OF tendsto_const[of "norm z"] lim_inverse_n], 
+        simp add:  norm_divide inverse_eq_divide)
+  from tendsto_add[OF this tendsto_const[of 1]] show "(\<lambda>n. z / of_nat n + 1) \<longlonglongrightarrow> 1" by simp
+qed
+
+
+subsection \<open>Convergence of the Euler series form\<close>
+
+text \<open>
+  We now show that the series that defines the Gamma function in the Euler form converges 
+  and that the function defined by it is continuous on the complex halfspace with positive 
+  real part.
+  
+  We do this by showing that the logarithm of the Euler series is continuous and converges 
+  locally uniformly, which means that the log-Gamma function defined by its limit is also 
+  continuous.
+  
+  This will later allow us to lift holomorphicity and continuity from the log-Gamma 
+  function to the inverse Gamma function, and from that to the Gamma function itself.
+\<close>
+
+definition ln_Gamma_series :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "ln_Gamma_series z n = z * ln (of_nat n) - ln z - (\<Sum>k=1..n. ln (z / of_nat k + 1))"
+
+definition ln_Gamma_series' :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> nat \<Rightarrow> 'a" where
+  "ln_Gamma_series' z n =
+     - euler_mascheroni*z - ln z + (\<Sum>k=1..n. z / of_nat n - ln (z / of_nat k + 1))"
+
+definition ln_Gamma :: "('a :: {banach,real_normed_field,ln}) \<Rightarrow> 'a" where
+  "ln_Gamma z = lim (ln_Gamma_series z)"
+
+
+text \<open>
+  We now show that the log-Gamma series converges locally uniformly for all complex numbers except 
+  the non-positive integers. We do this by proving that the series is locally Cauchy, adapting this
+  proof:
+  http://math.stackexchange.com/questions/887158/convergence-of-gammaz-lim-n-to-infty-fracnzn-prod-m-0nzm
+\<close>
+
+context
+begin
+
+private lemma ln_Gamma_series_complex_converges_aux:
+  fixes z :: complex and k :: nat
+  assumes z: "z \<noteq> 0" and k: "of_nat k \<ge> 2*norm z" "k \<ge> 2"
+  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"
+proof -
+  let ?k = "of_nat k :: complex" and ?z = "norm z"
+  have "z *ln (1 - 1/?k) + ln (z/?k+1) = z*(ln (1 - 1/?k :: complex) + 1/?k) + (ln (1+z/?k) - z/?k)"
+    by (simp add: algebra_simps)
+  also have "norm ... \<le> ?z * norm (ln (1-1/?k) + 1/?k :: complex) + norm (ln (1+z/?k) - z/?k)"
+    by (subst norm_mult [symmetric], rule norm_triangle_ineq)
+  also have "norm (Ln (1 + -1/?k) - (-1/?k)) \<le> (norm (-1/?k))\<^sup>2 / (1 - norm(-1/?k))"
+    using k by (intro Ln_approx_linear) (simp add: norm_divide)
+  hence "?z * norm (ln (1-1/?k) + 1/?k) \<le> ?z * ((norm (1/?k))^2 / (1 - norm (1/?k)))" 
+    by (intro mult_left_mono) simp_all
+  also have "... \<le> (?z * (of_nat k / (of_nat k - 1))) / of_nat k^2" using k
+    by (simp add: field_simps power2_eq_square norm_divide)
+  also have "... \<le> (?z * 2) / of_nat k^2" using k
+    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
+  also have "norm (ln (1+z/?k) - z/?k) \<le> norm (z/?k)^2 / (1 - norm (z/?k))" using k
+    by (intro Ln_approx_linear) (simp add: norm_divide)
+  hence "norm (ln (1+z/?k) - z/?k) \<le> ?z^2 / of_nat k^2 / (1 - ?z / of_nat k)" 
+    by (simp add: field_simps norm_divide)
+  also have "... \<le> (?z^2 * (of_nat k / (of_nat k - ?z))) / of_nat k^2" using k
+    by (simp add: field_simps power2_eq_square)
+  also have "... \<le> (?z^2 * 2) / of_nat k^2" using k
+    by (intro divide_right_mono mult_left_mono) (simp_all add: field_simps)
+  also note add_divide_distrib [symmetric]
+  finally show ?thesis by (simp only: distrib_left mult.commute)
+qed
+
+lemma ln_Gamma_series_complex_converges:
+  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  assumes d: "d > 0" "\<And>n. n \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> norm (z - of_int n) > d"
+  shows "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n :: complex)"
+proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI')
+  fix e :: real assume e: "e > 0"
+  def e'' \<equiv> "SUP t:ball z d. norm t + norm t^2"
+  def e' \<equiv> "e / (2*e'')"
+  have "bounded ((\<lambda>t. norm t + norm t^2) ` cball z d)"
+    by (intro compact_imp_bounded compact_continuous_image) (auto intro!: continuous_intros)
+  hence "bounded ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_subset) auto
+  hence bdd: "bdd_above ((\<lambda>t. norm t + norm t^2) ` ball z d)" by (rule bounded_imp_bdd_above)
+
+  with z d(1) d(2)[of "-1"] have e''_pos: "e'' > 0" unfolding e''_def
+    by (subst less_cSUP_iff) (auto intro!: add_pos_nonneg bexI[of _ z])
+  have e'': "norm t + norm t^2 \<le> e''" if "t \<in> ball z d" for t unfolding e''_def using that
+    by (rule cSUP_upper[OF _ bdd])
+  from e z e''_pos have e': "e' > 0" unfolding e'_def 
+    by (intro divide_pos_pos mult_pos_pos add_pos_pos) (simp_all add: field_simps)
+
+  have "summable (\<lambda>k. inverse ((real_of_nat k)^2))"
+    by (rule inverse_power_summable) simp
+  from summable_partial_sum_bound[OF this e'] guess M . note M = this
+
+  def N \<equiv> "max 2 (max (nat \<lceil>2 * (norm z + d)\<rceil>) M)"
+  {
+    from d have "\<lceil>2 * (cmod z + d)\<rceil> \<ge> \<lceil>0::real\<rceil>" 
+      by (intro ceiling_mono mult_nonneg_nonneg add_nonneg_nonneg) simp_all
+    hence "2 * (norm z + d) \<le> of_nat (nat \<lceil>2 * (norm z + d)\<rceil>)" unfolding N_def
+      by (simp_all add: le_of_int_ceiling)
+    also have "... \<le> of_nat N" unfolding N_def 
+      by (subst of_nat_le_iff) (rule max.coboundedI2, rule max.cobounded1)
+    finally have "of_nat N \<ge> 2 * (norm z + d)" .
+    moreover have "N \<ge> 2" "N \<ge> M" unfolding N_def by simp_all
+    moreover have "(\<Sum>k=m..n. 1/(of_nat k)\<^sup>2) < e'" if "m \<ge> N" for m n
+      using M[OF order.trans[OF `N \<ge> M` that]] unfolding real_norm_def
+      by (subst (asm) abs_of_nonneg) (auto intro: setsum_nonneg simp: divide_simps)
+    moreover note calculation
+  } note N = this
+
+  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"
+    unfolding dist_complex_def
+  proof (intro exI[of _ N] ballI allI impI)
+    fix t m n assume t: "t \<in> ball z d" and mn: "m \<ge> N" "n > m"
+    from d(2)[of 0] t have "0 < dist z 0 - dist z t" by (simp add: field_simps dist_complex_def)
+    also have "dist z 0 - dist z t \<le> dist 0 t" using dist_triangle[of 0 z t]
+      by (simp add: dist_commute)
+    finally have t_nz: "t \<noteq> 0" by auto
+
+    have "norm t \<le> norm z + norm (t - z)" by (rule norm_triangle_sub)
+    also from t have "norm (t - z) < d" by (simp add: dist_complex_def norm_minus_commute)
+    also have "2 * (norm z + d) \<le> of_nat N" by (rule N)
+    also have "N \<le> m" by (rule mn)
+    finally have norm_t: "2 * norm t < of_nat m" by simp
+
+    have "ln_Gamma_series t m - ln_Gamma_series t n = 
+              (-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m)))) +
+              ((\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)))"
+      by (simp add: ln_Gamma_series_def algebra_simps)
+    also have "(\<Sum>k=1..n. Ln (t / of_nat k + 1)) - (\<Sum>k=1..m. Ln (t / of_nat k + 1)) =
+                 (\<Sum>k\<in>{1..n}-{1..m}. Ln (t / of_nat k + 1))" using mn
+      by (simp add: setsum_diff)
+    also from mn have "{1..n}-{1..m} = {Suc m..n}" by fastforce
+    also have "-(t * Ln (of_nat n)) - (-(t * Ln (of_nat m))) = 
+                   (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1)) - t * Ln (of_nat k))" using mn
+      by (subst setsum_telescope'' [symmetric]) simp_all
+    also have "... = (\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k))" using mn N
+      by (intro setsum_cong_Suc)
+         (simp_all del: of_nat_Suc add: field_simps Ln_of_nat Ln_of_nat_over_of_nat)
+    also have "of_nat (k - 1) / of_nat k = 1 - 1 / (of_nat k :: complex)" if "k \<in> {Suc m..n}" for k
+      using that of_nat_eq_0_iff[of "Suc i" for i] by (cases k) (simp_all add: divide_simps)
+    hence "(\<Sum>k = Suc m..n. t * Ln (of_nat (k - 1) / of_nat k)) =
+             (\<Sum>k = Suc m..n. t * Ln (1 - 1 / of_nat k))" using mn N
+      by (intro setsum.cong) simp_all
+    also note setsum.distrib [symmetric]
+    also have "norm (\<Sum>k=Suc m..n. t * Ln (1 - 1/of_nat k) + Ln (t/of_nat k + 1)) \<le> 
+      (\<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
+      by (intro order.trans[OF norm_setsum setsum_mono[OF ln_Gamma_series_complex_converges_aux]]) simp_all
+    also have "... \<le> 2 * (norm t + norm t^2) * (\<Sum>k=Suc m..n. 1 / (of_nat k)\<^sup>2)"
+      by (simp add: setsum_right_distrib)
+    also have "... < 2 * (norm t + norm t^2) * e'" using mn z t_nz
+      by (intro mult_strict_left_mono N mult_pos_pos add_pos_pos) simp_all
+    also from e''_pos have "... = e * ((cmod t + (cmod t)\<^sup>2) / e'')" 
+      by (simp add: e'_def field_simps power2_eq_square)
+    also from e''[OF t] e''_pos e 
+      have "\<dots> \<le> e * 1" by (intro mult_left_mono) (simp_all add: field_simps)
+    finally show "norm (ln_Gamma_series t m - ln_Gamma_series t n) < e" by simp
+  qed
+qed
+
+end
+
+lemma ln_Gamma_series_complex_converges':
+  assumes z: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "\<exists>d>0. uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
+proof -
+  def d' \<equiv> "Re z"
+  def d \<equiv> "if d' > 0 then d' / 2 else norm (z - of_int (round d')) / 2"
+  have "of_int (round d') \<in> \<int>\<^sub>\<le>\<^sub>0" if "d' \<le> 0" using that
+    by (intro nonpos_Ints_of_int) (simp_all add: round_def)
+  with assms have d_pos: "d > 0" unfolding d_def by (force simp: not_less)
+
+  have "d < cmod (z - of_int n)" if "n \<in> \<int>\<^sub>\<le>\<^sub>0" for n
+  proof (cases "Re z > 0")
+    case True
+    from nonpos_Ints_nonpos[OF that] have n: "n \<le> 0" by simp
+    from True have "d = Re z/2" by (simp add: d_def d'_def)
+    also from n True have "\<dots> < Re (z - of_int n)" by simp
+    also have "\<dots> \<le> norm (z - of_int n)" by (rule complex_Re_le_cmod)
+    finally show ?thesis .
+  next
+    case False
+    with assms nonpos_Ints_of_int[of "round (Re z)"] 
+      have "z \<noteq> of_int (round d')" by (auto simp: not_less)
+    with False have "d < norm (z - of_int (round d'))" by (simp add: d_def d'_def)
+    also have "\<dots> \<le> norm (z - of_int n)" unfolding d'_def by (rule round_Re_minimises_norm)
+    finally show ?thesis .
+  qed
+  hence conv: "uniformly_convergent_on (ball z d) (\<lambda>n z. ln_Gamma_series z n)"
+    by (intro ln_Gamma_series_complex_converges d_pos z) simp_all
+  from d_pos conv show ?thesis by blast
+qed
+
+lemma ln_Gamma_series_complex_converges'': "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> convergent (ln_Gamma_series z)"
+  by (drule ln_Gamma_series_complex_converges') (auto intro: uniformly_convergent_imp_convergent)
+
+lemma ln_Gamma_complex_LIMSEQ: "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> ln_Gamma_series z \<longlonglongrightarrow> ln_Gamma z"
+  using ln_Gamma_series_complex_converges'' by (simp add: convergent_LIMSEQ_iff ln_Gamma_def)
+
+lemma exp_ln_Gamma_series_complex:
+  assumes "n > 0" "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "exp (ln_Gamma_series z n :: complex) = Gamma_series z n"
+proof -
+  from assms have "z \<noteq> 0" by (intro notI) auto
+  with assms have "exp (ln_Gamma_series z n) = 
+          (of_nat n) powr z / (z * (\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))))"
+    unfolding ln_Gamma_series_def powr_def by (simp add: exp_diff exp_setsum)
+  also from assms have "(\<Prod>k=1..n. exp (Ln (z / of_nat k + 1))) = (\<Prod>k=1..n. z / of_nat k + 1)"
+    by (intro setprod.cong[OF refl], subst exp_Ln) (auto simp: field_simps plus_of_nat_eq_0_imp)
+  also have "... = (\<Prod>k=1..n. z + k) / fact n" unfolding fact_altdef
+    by (subst setprod_dividef [symmetric]) (simp_all add: field_simps)
+  also from assms have "z * ... = (\<Prod>k=0..n. z + k) / fact n"
+    by (cases n) (simp_all add: setprod_nat_ivl_1_Suc)
+  also have "(\<Prod>k=0..n. z + k) = pochhammer z (Suc n)" unfolding pochhammer_def by simp
+  also have "of_nat n powr z / (pochhammer z (Suc n) / fact n) = Gamma_series z n"
+    unfolding Gamma_series_def using assms by (simp add: divide_simps powr_def Ln_of_nat)
+  finally show ?thesis .
+qed
+
+
+lemma ln_Gamma_series'_aux:
+  assumes "(z::complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(\<lambda>k. z / of_nat (Suc k) - ln (1 + z / of_nat (Suc k))) sums 
+              (ln_Gamma z + euler_mascheroni * z + ln z)" (is "?f sums ?s")
+unfolding sums_def
+proof (rule Lim_transform)
+  show "(\<lambda>n. ln_Gamma_series z n + of_real (harm n - ln (of_nat n)) * z + ln z) \<longlonglongrightarrow> ?s"
+    (is "?g \<longlonglongrightarrow> _") 
+    by (intro tendsto_intros ln_Gamma_complex_LIMSEQ euler_mascheroni_LIMSEQ_of_real assms)
+
+  have A: "eventually (\<lambda>n. (\<Sum>k<n. ?f k) - ?g n = 0) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    have "(\<Sum>k<n. ?f k) = (\<Sum>k=1..n. z / of_nat k - ln (1 + z / of_nat k))"
+      by (subst atLeast0LessThan [symmetric], subst setsum_shift_bounds_Suc_ivl [symmetric],
+          subst atLeastLessThanSuc_atLeastAtMost) simp_all
+    also have "\<dots> = z * of_real (harm n) - (\<Sum>k=1..n. ln (1 + z / of_nat k))"
+      by (simp add: harm_def setsum_subtractf setsum_right_distrib divide_inverse)
+    also from n have "\<dots> - ?g n = 0"
+      by (simp add: ln_Gamma_series_def setsum_subtractf algebra_simps Ln_of_nat) 
+    finally show "(\<Sum>k<n. ?f k) - ?g n = 0" .
+  qed
+  show "(\<lambda>n. (\<Sum>k<n. ?f k) - ?g n) \<longlonglongrightarrow> 0" by (subst tendsto_cong[OF A]) simp_all
+qed
+
+
+lemma uniformly_summable_deriv_ln_Gamma:
+  assumes z: "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0" and d: "d > 0" "d \<le> norm z/2"
+  shows "uniformly_convergent_on (ball z d)
+            (\<lambda>k z. \<Sum>i<k. inverse (of_nat (Suc i)) - inverse (z + of_nat (Suc i)))"
+           (is "uniformly_convergent_on _ (\<lambda>k z. \<Sum>i<k. ?f i z)")
+proof (rule weierstrass_m_test'_ev)
+  {
+    fix t assume t: "t \<in> ball z d"
+    have "norm z = norm (t + (z - t))" by simp
+    have "norm (t + (z - t)) \<le> norm t + norm (z - t)" by (rule norm_triangle_ineq)
+    also from t d have "norm (z - t) < norm z / 2" by (simp add: dist_norm)
+    finally have A: "norm t > norm z / 2" using z by (simp add: field_simps)
+    
+    have "norm t = norm (z + (t - z))" by simp
+    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
+    also from t d have "norm (t - z) \<le> norm z / 2" by (simp add: dist_norm norm_minus_commute)
+    also from z have "\<dots> < norm z" by simp
+    finally have B: "norm t < 2 * norm z" by simp
+    note A B
+  } note ball = this
+
+  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"
+    using eventually_gt_at_top apply eventually_elim
+  proof safe
+    fix t :: 'a assume t: "t \<in> ball z d"
+    from z ball[OF t] have t_nz: "t \<noteq> 0" by auto
+    fix n :: nat assume n: "n > nat \<lceil>4 * norm z\<rceil>"
+    from ball[OF t] t_nz have "4 * norm z > 2 * norm t" by simp
+    also from n have "\<dots>  < of_nat n" by linarith
+    finally have n: "of_nat n > 2 * norm t" .
+    hence "of_nat n > norm t" by simp
+    hence t': "t \<noteq> -of_nat (Suc n)" by (intro notI) (simp del: of_nat_Suc)
+
+    with t_nz have "?f n t = 1 / (of_nat (Suc n) * (1 + of_nat (Suc n)/t))"
+      by (simp add: divide_simps eq_neg_iff_add_eq_0 del: of_nat_Suc)
+    also have "norm \<dots> = inverse (of_nat (Suc n)) * inverse (norm (of_nat (Suc n)/t + 1))"
+      by (simp add: norm_divide norm_mult divide_simps add_ac del: of_nat_Suc)
+    also {
+      from z t_nz ball[OF t] have "of_nat (Suc n) / (4 * norm z) \<le> of_nat (Suc n) / (2 * norm t)"
+        by (intro divide_left_mono mult_pos_pos) simp_all
+      also have "\<dots> < norm (of_nat (Suc n) / t) - norm (1 :: 'a)" 
+        using t_nz n by (simp add: field_simps norm_divide del: of_nat_Suc)
+      also have "\<dots> \<le> norm (of_nat (Suc n)/t + 1)" by (rule norm_diff_ineq)
+      finally have "inverse (norm (of_nat (Suc n)/t + 1)) \<le> 4 * norm z / of_nat (Suc n)"
+        using z by (simp add: divide_simps norm_divide mult_ac del: of_nat_Suc)
+    }
+    also have "inverse (real_of_nat (Suc n)) * (4 * norm z / real_of_nat (Suc n)) =
+                 4 * norm z * inverse (of_nat (Suc n)^2)" 
+                 by (simp add: divide_simps power2_eq_square del: of_nat_Suc)
+    finally show "norm (?f n t) \<le> 4 * norm z * inverse (of_nat (Suc n)^2)" 
+      by (simp del: of_nat_Suc)
+  qed
+next
+  show "summable (\<lambda>n. 4 * norm z * inverse ((of_nat (Suc n))^2))"
+    by (subst summable_Suc_iff) (simp add: summable_mult inverse_power_summable)
+qed
+
+lemma summable_deriv_ln_Gamma:
+  "z \<noteq> (0 :: 'a :: {real_normed_field,banach}) \<Longrightarrow> 
+     summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat (Suc n)))"
+  unfolding summable_iff_convergent
+  by (rule uniformly_convergent_imp_convergent, 
+      rule uniformly_summable_deriv_ln_Gamma[of z "norm z/2"]) simp_all
+
+
+definition Polygamma :: "nat \<Rightarrow> ('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
+  "Polygamma n z = (if n = 0 then 
+      (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni else 
+      (-1)^Suc n * fact n * (\<Sum>k. inverse ((z + of_nat k)^Suc n)))"
+
+abbreviation Digamma :: "('a :: {real_normed_field,banach}) \<Rightarrow> 'a" where
+  "Digamma \<equiv> Polygamma 0"
+
+lemma Digamma_def: 
+  "Digamma z = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) - euler_mascheroni"
+  by (simp add: Polygamma_def)
+
+
+lemma summable_Digamma: 
+  assumes "(z :: 'a :: {real_normed_field,banach}) \<noteq> 0"
+  shows   "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
+proof -
+  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums 
+                       (0 - inverse (z + of_nat 0))"
+    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
+              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
+  from summable_add[OF summable_deriv_ln_Gamma[OF assms] sums_summable[OF sums]]
+    show "summable (\<lambda>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))" by simp
+qed
+
+lemma summable_offset:
+  assumes "summable (\<lambda>n. f (n + k) :: 'a :: real_normed_vector)"
+  shows   "summable f"
+proof -
+  from assms have "convergent (\<lambda>m. \<Sum>n<m. f (n + k))" by (simp add: summable_iff_convergent)
+  hence "convergent (\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)))" 
+    by (intro convergent_add convergent_const)
+  also have "(\<lambda>m. (\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k))) = (\<lambda>m. \<Sum>n<m+k. f n)"
+  proof
+    fix m :: nat
+    have "{..<m+k} = {..<k} \<union> {k..<m+k}" by auto
+    also have "(\<Sum>n\<in>\<dots>. f n) = (\<Sum>n<k. f n) + (\<Sum>n=k..<m+k. f n)" 
+      by (rule setsum.union_disjoint) auto
+    also have "(\<Sum>n=k..<m+k. f n) = (\<Sum>n=0..<m+k-k. f (n + k))"
+      by (intro setsum.reindex_cong[of "\<lambda>n. n + k"])
+         (simp, subst image_add_atLeastLessThan, auto)
+    finally show "(\<Sum>n<k. f n) + (\<Sum>n<m. f (n + k)) = (\<Sum>n<m+k. f n)" by (simp add: atLeast0LessThan)
+  qed
+  finally have "(\<lambda>a. setsum f {..<a}) \<longlonglongrightarrow> lim (\<lambda>m. setsum f {..<m + k})"
+    by (auto simp: convergent_LIMSEQ_iff dest: LIMSEQ_offset)
+  thus ?thesis by (auto simp: summable_iff_convergent convergent_def)
+qed
+
+lemma Polygamma_converges:
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
+  shows "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)^n))"
+proof (rule weierstrass_m_test'_ev)
+  def e \<equiv> "(1 + d / norm z)"
+  def m \<equiv> "nat \<lceil>norm z * e\<rceil>"
+  {
+    fix t assume t: "t \<in> ball z d"
+    have "norm t = norm (z + (t - z))" by simp
+    also have "\<dots> \<le> norm z + norm (t - z)" by (rule norm_triangle_ineq)
+    also from t have "norm (t - z) < d" by (simp add: dist_norm norm_minus_commute)
+    finally have "norm t < norm z * e" using z by (simp add: divide_simps e_def)
+  } note ball = this
+  
+  show "eventually (\<lambda>k. \<forall>t\<in>ball z d. norm (inverse ((t + of_nat k)^n)) \<le> 
+            inverse (of_nat (k - m)^n)) sequentially"
+    using eventually_gt_at_top[of m] apply eventually_elim
+  proof (intro ballI)
+    fix k :: nat and t :: 'a assume k: "k > m" and t: "t \<in> ball z d"
+    from k have "real_of_nat (k - m) = of_nat k - of_nat m" by (simp add: of_nat_diff)
+    also have "\<dots> \<le> norm (of_nat k :: 'a) - norm z * e" 
+      unfolding m_def by (subst norm_of_nat) linarith
+    also from ball[OF t] have "\<dots> \<le> norm (of_nat k :: 'a) - norm t" by simp
+    also have "\<dots> \<le> norm (of_nat k + t)" by (rule norm_diff_ineq)
+    finally have "inverse ((norm (t + of_nat k))^n) \<le> inverse (real_of_nat (k - m)^n)" using k n
+      by (intro le_imp_inverse_le power_mono) (simp_all add: add_ac del: of_nat_Suc)
+    thus "norm (inverse ((t + of_nat k)^n)) \<le> inverse (of_nat (k - m)^n)"
+      by (simp add: norm_inverse norm_power power_inverse)
+  qed
+  
+  have "summable (\<lambda>k. inverse ((real_of_nat k)^n))" 
+    using inverse_power_summable[of n] n by simp
+  hence "summable (\<lambda>k. inverse ((real_of_nat (k + m - m))^n))" by simp
+  thus "summable (\<lambda>k. inverse ((real_of_nat (k - m))^n))" by (rule summable_offset)
+qed
+
+lemma Polygamma_converges':
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes z: "z \<noteq> 0" and n: "n \<ge> 2"
+  shows "summable (\<lambda>k. inverse ((z + of_nat k)^n))"
+  using uniformly_convergent_imp_convergent[OF Polygamma_converges[OF assms, of 1], of z]
+  by (simp add: summable_iff_convergent)
+
+lemma has_field_derivative_ln_Gamma_complex [derivative_intros]:
+  assumes z: "Re z > 0 \<or> Im z \<noteq> 0"
+  shows   "(ln_Gamma has_field_derivative Digamma z) (at z)"
+proof -
+  have not_nonpos_Int [simp]: "t \<notin> \<int>\<^sub>\<le>\<^sub>0" if "Re t > 0" for t 
+    using that by (auto elim!: nonpos_Ints_cases')
+  from z have z': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" and z'': "z \<noteq> 0" by (auto elim!: nonpos_Ints_cases)
+  let ?f' = "\<lambda>z k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))"
+  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"
+  def d \<equiv> "min (norm z/2) (if Im z = 0 then Re z / 2 else abs (Im z) / 2)"
+  from z have d: "d > 0" "norm z/2 \<ge> d" by (auto simp add: d_def)
+  have ball: "Im t = 0 \<longrightarrow> Re t > 0" if "dist z t < d" for t 
+    using Re_pos_or_Im_nz_in_ball[OF z, of t] that unfolding d_def by auto
+  have sums: "(\<lambda>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n)) sums 
+                       (0 - inverse (z + of_nat 0))"
+    by (intro telescope_sums filterlim_compose[OF tendsto_inverse_0]
+              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
+
+  have "((\<lambda>z. \<Sum>n. ?f z n) has_field_derivative ?F' z) (at z)"
+    using d z ln_Gamma_series'_aux[OF z']
+    by (intro has_field_derivative_series'(2)[of "ball z d" _ _ z] uniformly_summable_deriv_ln_Gamma)
+       (auto intro!: derivative_eq_intros add_pos_pos mult_pos_pos dest!: ball
+             simp: field_simps Re_complex_div_gt_0 sums_iff Im_divide_of_nat
+             simp del: of_nat_Suc)
+  with z have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) has_field_derivative 
+                   ?F' z - euler_mascheroni - inverse z) (at z)" 
+    by (force intro!: derivative_eq_intros simp: Digamma_def)
+  also have "?F' z - euler_mascheroni - inverse z = (?F' z + -inverse z) - euler_mascheroni" by simp
+  also from sums have "-inverse z = (\<Sum>n. inverse (z + of_nat (Suc n)) - inverse (z + of_nat n))"
+    by (simp add: sums_iff)
+  also from sums summable_deriv_ln_Gamma[OF z''] 
+    have "?F' z + \<dots> =  (\<Sum>n. inverse (of_nat (Suc n)) - inverse (z + of_nat n))"
+    by (subst suminf_add) (simp_all add: add_ac sums_iff)
+  also have "\<dots> - euler_mascheroni = Digamma z" by (simp add: Digamma_def)
+  finally have "((\<lambda>z. (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) 
+                    has_field_derivative Digamma z) (at z)" .
+  moreover from eventually_nhds_ball[OF d(1), of z]
+    have "eventually (\<lambda>z. ln_Gamma z = (\<Sum>k. ?f z k) - euler_mascheroni * z - Ln z) (nhds z)"
+  proof eventually_elim
+    fix t assume "t \<in> ball z d"
+    hence "t \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto dest!: ball elim!: nonpos_Ints_cases)
+    from ln_Gamma_series'_aux[OF this] 
+      show "ln_Gamma t = (\<Sum>k. ?f t k) - euler_mascheroni * t - Ln t" by (simp add: sums_iff)
+  qed
+  ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
+qed
+
+declare has_field_derivative_ln_Gamma_complex[THEN DERIV_chain2, derivative_intros]
+
+
+lemma Digamma_1 [simp]: "Digamma (1 :: 'a :: {real_normed_field,banach}) = - euler_mascheroni"
+  by (simp add: Digamma_def)
+  
+lemma Digamma_plus1:
+  assumes "z \<noteq> 0"
+  shows   "Digamma (z+1) = Digamma z + 1/z"
+proof -
+  have sums: "(\<lambda>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) 
+                  sums (inverse (z + of_nat 0) - 0)"
+    by (intro telescope_sums'[OF filterlim_compose[OF tendsto_inverse_0]]
+              tendsto_add_filterlim_at_infinity[OF tendsto_const] tendsto_of_nat)
+  have "Digamma (z+1) = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat (Suc k))) - 
+          euler_mascheroni" (is "_ = suminf ?f - _") by (simp add: Digamma_def add_ac)
+  also have "suminf ?f = (\<Sum>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) +
+                         (\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k)))"
+    using summable_Digamma[OF assms] sums by (subst suminf_add) (simp_all add: add_ac sums_iff)
+  also have "(\<Sum>k. inverse (z + of_nat k) - inverse (z + of_nat (Suc k))) = 1/z"
+    using sums by (simp add: sums_iff inverse_eq_divide)
+  finally show ?thesis by (simp add: Digamma_def[of z])
+qed
+
+lemma Polygamma_plus1:
+  assumes "z \<noteq> 0"
+  shows   "Polygamma n (z + 1) = Polygamma n z + (-1)^n * fact n / (z ^ Suc n)"
+proof (cases "n = 0")
+  assume n: "n \<noteq> 0"
+  let ?f = "\<lambda>k. inverse ((z + of_nat k) ^ Suc n)"
+  have "Polygamma n (z + 1) = (-1) ^ Suc n * fact n * (\<Sum>k. ?f (k+1))"
+    using n by (simp add: Polygamma_def add_ac)
+  also have "(\<Sum>k. ?f (k+1)) + (\<Sum>k<1. ?f k) = (\<Sum>k. ?f k)"
+    using Polygamma_converges'[OF assms, of "Suc n"] n
+    by (subst suminf_split_initial_segment [symmetric]) simp_all
+  hence "(\<Sum>k. ?f (k+1)) = (\<Sum>k. ?f k) - inverse (z ^ Suc n)" by (simp add: algebra_simps)
+  also have "(-1) ^ Suc n * fact n * ((\<Sum>k. ?f k) - inverse (z ^ Suc n)) = 
+               Polygamma n z + (-1)^n * fact n / (z ^ Suc n)" using n
+    by (simp add: inverse_eq_divide algebra_simps Polygamma_def)
+  finally show ?thesis .
+qed (insert assms, simp add: Digamma_plus1 inverse_eq_divide)
+
+lemma Digamma_of_nat: 
+  "Digamma (of_nat (Suc n) :: 'a :: {real_normed_field,banach}) = harm n - euler_mascheroni"
+proof (induction n)
+  case (Suc n)
+  have "Digamma (of_nat (Suc (Suc n)) :: 'a) = Digamma (of_nat (Suc n) + 1)" by simp
+  also have "\<dots> = Digamma (of_nat (Suc n)) + inverse (of_nat (Suc n))" 
+    by (subst Digamma_plus1) (simp_all add: inverse_eq_divide del: of_nat_Suc)
+  also have "Digamma (of_nat (Suc n) :: 'a) = harm n - euler_mascheroni " by (rule Suc)
+  also have "\<dots> + inverse (of_nat (Suc n)) = harm (Suc n) - euler_mascheroni"
+    by (simp add: harm_Suc)
+  finally show ?case .
+qed (simp add: harm_def)
+
+lemma Digamma_numeral: "Digamma (numeral n) = harm (pred_numeral n) - euler_mascheroni"
+  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Digamma_of_nat) (rule refl)
+
+lemma Polygamma_of_real: "x \<noteq> 0 \<Longrightarrow> Polygamma n (of_real x) = of_real (Polygamma n x)"
+  unfolding Polygamma_def using summable_Digamma[of x] Polygamma_converges'[of x "Suc n"]
+  by (simp_all add: suminf_of_real)
+
+lemma Polygamma_Real: "z \<in> \<real> \<Longrightarrow> z \<noteq> 0 \<Longrightarrow> Polygamma n z \<in> \<real>"
+  by (elim Reals_cases, hypsubst, subst Polygamma_of_real) simp_all
+
+lemma Digamma_half_integer:
+  "Digamma (of_nat n + 1/2 :: 'a :: {real_normed_field,banach}) = 
+      (\<Sum>k<n. 2 / (of_nat (2*k+1))) - euler_mascheroni - of_real (2 * ln 2)"
+proof (induction n)
+  case 0
+  have "Digamma (1/2 :: 'a) = of_real (Digamma (1/2))" by (simp add: Polygamma_of_real [symmetric])
+  also have "Digamma (1/2::real) = 
+               (\<Sum>k. inverse (of_nat (Suc k)) - inverse (of_nat k + 1/2)) - euler_mascheroni" 
+    by (simp add: Digamma_def add_ac)
+  also have "(\<Sum>k. inverse (of_nat (Suc k) :: real) - inverse (of_nat k + 1/2)) =
+             (\<Sum>k. inverse (1/2) * (inverse (2 * of_nat (Suc k)) - inverse (2 * of_nat k + 1)))"
+    by (simp_all add: add_ac inverse_mult_distrib[symmetric] ring_distribs del: inverse_divide)
+  also have "\<dots> = - 2 * ln 2" using sums_minus[OF alternating_harmonic_series_sums']
+    by (subst suminf_mult) (simp_all add: algebra_simps sums_iff)
+  finally show ?case by simp
+next
+  case (Suc n)
+  have nz: "2 * of_nat n + (1:: 'a) \<noteq> 0"
+     using of_nat_neq_0[of "2*n"] by (simp only: of_nat_Suc) (simp add: add_ac)
+  hence nz': "of_nat n + (1/2::'a) \<noteq> 0" by (simp add: field_simps)
+  have "Digamma (of_nat (Suc n) + 1/2 :: 'a) = Digamma (of_nat n + 1/2 + 1)" by simp
+  also from nz' have "\<dots> = Digamma (of_nat n + 1 / 2) + 1 / (of_nat n + 1 / 2)"
+    by (rule Digamma_plus1)
+  also from nz nz' have "1 / (of_nat n + 1 / 2 :: 'a) = 2 / (2 * of_nat n + 1)"
+    by (subst divide_eq_eq) simp_all
+  also note Suc
+  finally show ?case by (simp add: add_ac)
+qed
+
+lemma Digamma_one_half: "Digamma (1/2) = - euler_mascheroni - of_real (2 * ln 2)"
+  using Digamma_half_integer[of 0] by simp
+
+lemma Digamma_real_three_halves_pos: "Digamma (3/2 :: real) > 0"
+proof -
+  have "-Digamma (3/2 :: real) = -Digamma (of_nat 1 + 1/2)" by simp
+  also have "\<dots> = 2 * ln 2 + euler_mascheroni - 2" by (subst Digamma_half_integer) simp
+  also from euler_mascheroni_approx have "euler_mascheroni \<le> (0.58::real)" 
+    by (simp add: abs_real_def split: split_if_asm)
+  also from ln_2_bounds have "ln 2 < (0.7 :: real)" by simp
+  finally show ?thesis by simp
+qed
+
+
+lemma has_field_derivative_Polygamma [derivative_intros]:
+  fixes z :: "'a :: {real_normed_field,euclidean_space}"
+  assumes z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z within A)"
+proof (rule has_field_derivative_at_within, cases "n = 0")
+  assume n: "n = 0"
+  let ?f = "\<lambda>k z. inverse (of_nat (Suc k)) - inverse (z + of_nat k)"
+  let ?F = "\<lambda>z. \<Sum>k. ?f k z" and ?f' = "\<lambda>k z. inverse ((z + of_nat k)\<^sup>2)"
+  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
+  from z have summable: "summable (\<lambda>k. inverse (of_nat (Suc k)) - inverse (z + of_nat k))"
+    by (intro summable_Digamma) force
+  from z have conv: "uniformly_convergent_on (ball z d) (\<lambda>k z. \<Sum>i<k. inverse ((z + of_nat i)\<^sup>2))"
+    by (intro Polygamma_converges) auto
+  with d have "summable (\<lambda>k. inverse ((z + of_nat k)\<^sup>2))" unfolding summable_iff_convergent
+    by (auto dest!: uniformly_convergent_imp_convergent simp: summable_iff_convergent )
+
+  have "(?F has_field_derivative (\<Sum>k. ?f' k z)) (at z)"
+  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
+    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
+    from t d(2)[of t] show "((\<lambda>z. ?f k z) has_field_derivative ?f' k t) (at t within ball z d)" 
+      by (auto intro!: derivative_eq_intros simp: power2_eq_square simp del: of_nat_Suc 
+               dest!: plus_of_nat_eq_0_imp elim!: nonpos_Ints_cases)
+  qed (insert d(1) summable conv, (assumption|simp)+)
+  with z show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
+    unfolding Digamma_def [abs_def] Polygamma_def [abs_def] using n
+    by (force simp: power2_eq_square intro!: derivative_eq_intros)
+next
+  assume n: "n \<noteq> 0"
+  from z have z': "z \<noteq> 0" by auto
+  from no_nonpos_Int_in_ball'[OF z] guess d . note d = this
+  def n' \<equiv> "Suc n"
+  from n have n': "n' \<ge> 2" by (simp add: n'_def)
+  have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
+                (\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n'+1)))) (at z)"
+  proof (rule has_field_derivative_series'[of "ball z d" _ _ z])
+    fix k :: nat and t :: 'a assume t: "t \<in> ball z d"
+    with d have t': "t \<notin> \<int>\<^sub>\<le>\<^sub>0" "t \<noteq> 0" by auto
+    show "((\<lambda>a. inverse ((a + of_nat k) ^ n')) has_field_derivative
+                - of_nat n' * inverse ((t + of_nat k) ^ (n'+1))) (at t within ball z d)" using t'
+      by (fastforce intro!: derivative_eq_intros simp: divide_simps power_diff dest: plus_of_nat_eq_0_imp)
+  next
+    have "uniformly_convergent_on (ball z d) 
+              (\<lambda>k z. (- of_nat n' :: 'a) * (\<Sum>i<k. inverse ((z + of_nat i) ^ (n'+1))))" 
+      using z' n by (intro uniformly_convergent_mult Polygamma_converges) (simp_all add: n'_def)
+    thus "uniformly_convergent_on (ball z d) 
+              (\<lambda>k z. \<Sum>i<k. - of_nat n' * inverse ((z + of_nat i :: 'a) ^ (n'+1)))"
+      by (subst (asm) setsum_right_distrib) simp
+  qed (insert Polygamma_converges'[OF z' n'] d, simp_all)
+  also have "(\<Sum>k. - of_nat n' * inverse ((z + of_nat k) ^ (n' + 1))) =
+               (- of_nat n') * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))"
+    using Polygamma_converges'[OF z', of "n'+1"] n' by (subst suminf_mult) simp_all
+  finally have "((\<lambda>z. \<Sum>k. inverse ((z + of_nat k) ^ n')) has_field_derivative
+                    - of_nat n' * (\<Sum>k. inverse ((z + of_nat k) ^ (n' + 1)))) (at z)" .
+  from DERIV_cmult[OF this, of "(-1)^Suc n * fact n :: 'a"]
+    show "(Polygamma n has_field_derivative Polygamma (Suc n) z) (at z)"
+    unfolding n'_def Polygamma_def[abs_def] using n by (simp add: algebra_simps)
+qed
+
+declare has_field_derivative_Polygamma[THEN DERIV_chain2, derivative_intros]
+
+lemma isCont_Polygamma [continuous_intros]: 
+  fixes f :: "_ \<Rightarrow> 'a :: {real_normed_field,euclidean_space}"
+  shows "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Polygamma n (f x)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Polygamma]])
+
+lemma continuous_on_Polygamma: 
+  "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A (Polygamma n :: _ \<Rightarrow> 'a :: {real_normed_field,euclidean_space})"
+  by (intro continuous_at_imp_continuous_on isCont_Polygamma[OF continuous_ident] ballI) blast
+
+lemma isCont_ln_Gamma_complex [continuous_intros]: 
+  "isCont f z \<Longrightarrow> Re (f z) > 0 \<or> Im (f z) \<noteq> 0 \<Longrightarrow> isCont (\<lambda>z. ln_Gamma (f z)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_ln_Gamma_complex]])
+
+lemma continuous_on_ln_Gamma_complex [continuous_intros]:
+   "A \<inter> {z. Re z \<le> 0 \<and> Im z = 0} = {} \<Longrightarrow> continuous_on A ln_Gamma"
+  by (intro continuous_at_imp_continuous_on ballI isCont_ln_Gamma_complex[OF continuous_ident]) 
+     fastforce
+
+
+text \<open>
+  We define a type class that captures all the fundamental properties of the inverse Gamma function 
+  and defines the Gamma function upon that. This allows us to instantiate the type class both for 
+  the reals and for the complex numbers with a minimal amount of proof duplication. 
+\<close>
+
+class Gamma = real_normed_field + complete_space +
+  fixes rGamma :: "'a \<Rightarrow> 'a"
+  assumes rGamma_eq_zero_iff_aux: "rGamma z = 0 \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
+  assumes differentiable_rGamma_aux1: 
+    "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow>
+     let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k)) 
+               \<longlonglongrightarrow> d) - scaleR euler_mascheroni 1
+     in  filterlim (\<lambda>y. (rGamma y - rGamma z + rGamma z * d * (y - z)) /\<^sub>R 
+                        norm (y - z)) (nhds 0) (at z)"
+  assumes differentiable_rGamma_aux2: 
+    "let z = - of_nat n
+     in  filterlim (\<lambda>y. (rGamma y - rGamma z - (-1)^n * (setprod of_nat {1..n}) * (y - z)) /\<^sub>R 
+                        norm (y - z)) (nhds 0) (at z)"
+  assumes rGamma_series_aux: "(\<And>n. z \<noteq> - of_nat n) \<Longrightarrow> 
+             let fact' = (\<lambda>n. setprod of_nat {1..n});
+                 exp = (\<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x^k /\<^sub>R fact k) \<longlonglongrightarrow> e);
+                 pochhammer' = (\<lambda>a n. (\<Prod>n = 0..n. a + of_nat n))
+             in  filterlim (\<lambda>n. pochhammer' z n / (fact' n * exp (z * (ln (of_nat n) *\<^sub>R 1)))) 
+                     (nhds (rGamma z)) sequentially"
+begin
+subclass banach ..
+end
+
+definition "Gamma z = inverse (rGamma z)"
+
+
+subsection \<open>Basic properties\<close>
+
+lemma Gamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z = 0"
+  and rGamma_nonpos_Int: "z \<in> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z = 0"
+  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
+
+lemma Gamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma z \<noteq> 0"
+  and rGamma_nonzero: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> rGamma z \<noteq> 0"
+  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
+
+lemma Gamma_eq_zero_iff: "Gamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0" 
+  and rGamma_eq_zero_iff: "rGamma z = 0 \<longleftrightarrow> z \<in> \<int>\<^sub>\<le>\<^sub>0"
+  using rGamma_eq_zero_iff_aux[of z] unfolding Gamma_def by (auto elim!: nonpos_Ints_cases')
+
+lemma rGamma_inverse_Gamma: "rGamma z = inverse (Gamma z)"
+  unfolding Gamma_def by simp
+
+lemma rGamma_series_LIMSEQ [tendsto_intros]:
+  "rGamma_series z \<longlonglongrightarrow> rGamma z"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case False
+  hence "z \<noteq> - of_nat n" for n by auto
+  from rGamma_series_aux[OF this] show ?thesis
+    by (simp add: rGamma_series_def[abs_def] fact_altdef pochhammer_Suc_setprod
+                  exp_def of_real_def[symmetric] suminf_def sums_def[abs_def])
+qed (insert rGamma_eq_zero_iff[of z], simp_all add: rGamma_series_nonpos_Ints_LIMSEQ)
+
+lemma Gamma_series_LIMSEQ [tendsto_intros]: 
+  "Gamma_series z \<longlonglongrightarrow> Gamma z"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case False
+  hence "(\<lambda>n. inverse (rGamma_series z n)) \<longlonglongrightarrow> inverse (rGamma z)"
+    by (intro tendsto_intros) (simp_all add: rGamma_eq_zero_iff)
+  also have "(\<lambda>n. inverse (rGamma_series z n)) = Gamma_series z" 
+    by (simp add: rGamma_series_def Gamma_series_def[abs_def])
+  finally show ?thesis by (simp add: Gamma_def)
+qed (insert Gamma_eq_zero_iff[of z], simp_all add: Gamma_series_nonpos_Ints_LIMSEQ)
+
+lemma Gamma_altdef: "Gamma z = lim (Gamma_series z)"
+  using Gamma_series_LIMSEQ[of z] by (simp add: limI)
+
+lemma rGamma_1 [simp]: "rGamma 1 = 1"
+proof -
+  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact 
+                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
+  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
+  moreover have "rGamma_series 1 \<longlonglongrightarrow> rGamma 1" by (rule tendsto_intros)
+  ultimately show ?thesis by (intro LIMSEQ_unique)
+qed
+
+lemma rGamma_plus1: "z * rGamma (z + 1) = rGamma z"
+proof -
+  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
+  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
+             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
+      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
+    also from n have "\<dots> = ?f n * rGamma_series z n"
+      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
+    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
+  qed
+  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
+    by (intro tendsto_intros lim_inverse_n)
+  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
+  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
+    by (rule Lim_transform_eventually)
+  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
+    by (intro tendsto_intros)
+  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
+qed
+
+
+lemma pochhammer_rGamma: "rGamma z = pochhammer z n * rGamma (z + of_nat n)"
+proof (induction n arbitrary: z)
+  case (Suc n z)
+  have "rGamma z = pochhammer z n * rGamma (z + of_nat n)" by (rule Suc.IH)
+  also note rGamma_plus1 [symmetric]
+  finally show ?case by (simp add: add_ac pochhammer_rec')
+qed simp_all
+
+lemma Gamma_plus1: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma (z + 1) = z * Gamma z"
+  using rGamma_plus1[of z] by (simp add: rGamma_inverse_Gamma field_simps Gamma_eq_zero_iff)
+
+lemma pochhammer_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> pochhammer z n = Gamma (z + of_nat n) / Gamma z"
+  using pochhammer_rGamma[of z] 
+  by (simp add: rGamma_inverse_Gamma Gamma_eq_zero_iff field_simps)
+
+lemma Gamma_0 [simp]: "Gamma 0 = 0" 
+  and rGamma_0 [simp]: "rGamma 0 = 0"
+  and Gamma_neg_1 [simp]: "Gamma (- 1) = 0"
+  and rGamma_neg_1 [simp]: "rGamma (- 1) = 0"
+  and Gamma_neg_numeral [simp]: "Gamma (- numeral n) = 0"
+  and rGamma_neg_numeral [simp]: "rGamma (- numeral n) = 0"
+  and Gamma_neg_of_nat [simp]: "Gamma (- of_nat m) = 0"
+  and rGamma_neg_of_nat [simp]: "rGamma (- of_nat m) = 0"
+  by (simp_all add: rGamma_eq_zero_iff Gamma_eq_zero_iff)
+
+lemma Gamma_1 [simp]: "Gamma 1 = 1" unfolding Gamma_def by simp
+
+lemma Gamma_fact: "Gamma (of_nat (Suc n)) = fact n"
+  by (simp add: pochhammer_fact pochhammer_Gamma of_nat_in_nonpos_Ints_iff)
+
+lemma Gamma_numeral: "Gamma (numeral n) = fact (pred_numeral n)"
+  by (subst of_nat_numeral[symmetric], subst numeral_eq_Suc, subst Gamma_fact) (rule refl)
+
+lemma Gamma_of_int: "Gamma (of_int n) = (if n > 0 then fact (nat (n - 1)) else 0)"
+proof (cases "n > 0")
+  case True
+  hence "Gamma (of_int n) = Gamma (of_nat (Suc (nat (n - 1))))" by (subst of_nat_Suc) simp_all
+  with True show ?thesis by (subst (asm) Gamma_fact) simp
+qed (simp_all add: Gamma_eq_zero_iff nonpos_Ints_of_int)
+
+lemma rGamma_of_int: "rGamma (of_int n) = (if n > 0 then inverse (fact (nat (n - 1))) else 0)"
+  by (simp add: Gamma_of_int rGamma_inverse_Gamma)
+
+lemma Gamma_seriesI:
+  assumes "(\<lambda>n. g n / Gamma_series z n) \<longlonglongrightarrow> 1"
+  shows   "g \<longlonglongrightarrow> Gamma z"
+proof (rule Lim_transform_eventually)
+  have "1/2 > (0::real)" by simp
+  from tendstoD[OF assms, OF this] 
+    show "eventually (\<lambda>n. g n / Gamma_series z n * Gamma_series z n = g n) sequentially"
+    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
+  from assms have "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> 1 * Gamma z"
+    by (intro tendsto_intros)
+  thus "(\<lambda>n. g n / Gamma_series z n * Gamma_series z n) \<longlonglongrightarrow> Gamma z" by simp
+qed
+
+lemma Gamma_seriesI':
+  assumes "f \<longlonglongrightarrow> rGamma z"
+  assumes "(\<lambda>n. g n * f n) \<longlonglongrightarrow> 1"
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "g \<longlonglongrightarrow> Gamma z"
+proof (rule Lim_transform_eventually)
+  have "1/2 > (0::real)" by simp
+  from tendstoD[OF assms(2), OF this] show "eventually (\<lambda>n. g n * f n / f n = g n) sequentially"
+    by (force elim!: eventually_mono simp: dist_real_def dist_0_norm)
+  from assms have "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> 1 / rGamma z"
+    by (intro tendsto_divide assms) (simp_all add: rGamma_eq_zero_iff)
+  thus "(\<lambda>n. g n * f n / f n) \<longlonglongrightarrow> Gamma z" by (simp add: Gamma_def divide_inverse)
+qed
+
+lemma Gamma_series'_LIMSEQ: "Gamma_series' z \<longlonglongrightarrow> Gamma z"
+  by (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0") (simp_all add: Gamma_nonpos_Int Gamma_seriesI[OF Gamma_series_Gamma_series'] 
+                                      Gamma_series'_nonpos_Ints_LIMSEQ[of z])
+
+
+subsection \<open>Differentiability\<close>
+
+lemma has_field_derivative_rGamma_no_nonpos_int: 
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(rGamma has_field_derivative -rGamma z * Digamma z) (at z within A)"
+proof (rule has_field_derivative_at_within)
+  from assms have "z \<noteq> - of_nat n" for n by auto
+  from differentiable_rGamma_aux1[OF this] 
+    show "(rGamma has_field_derivative -rGamma z * Digamma z) (at z)"
+         unfolding Digamma_def suminf_def sums_def[abs_def]
+                   has_field_derivative_def has_derivative_def netlimit_at
+    by (simp add: Let_def bounded_linear_mult_right mult_ac of_real_def [symmetric])
+qed
+
+lemma has_field_derivative_rGamma_nonpos_int: 
+  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n) within A)" 
+  apply (rule has_field_derivative_at_within)
+  using differentiable_rGamma_aux2[of n]
+  unfolding Let_def has_field_derivative_def has_derivative_def netlimit_at
+  by (simp only: bounded_linear_mult_right mult_ac of_real_def [symmetric] fact_altdef)
+
+lemma has_field_derivative_rGamma [derivative_intros]: 
+  "(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>)
+      else -rGamma z * Digamma z)) (at z within A)"
+using has_field_derivative_rGamma_no_nonpos_int[of z A]
+      has_field_derivative_rGamma_nonpos_int[of "nat \<lfloor>norm z\<rfloor>" A]
+  by (auto elim!: nonpos_Ints_cases')
+
+declare has_field_derivative_rGamma_no_nonpos_int [THEN DERIV_chain2, derivative_intros]
+declare has_field_derivative_rGamma [THEN DERIV_chain2, derivative_intros]
+declare has_field_derivative_rGamma_nonpos_int [derivative_intros]
+declare has_field_derivative_rGamma_no_nonpos_int [derivative_intros]
+declare has_field_derivative_rGamma [derivative_intros]
+
+
+lemma has_field_derivative_Gamma [derivative_intros]: 
+  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> (Gamma has_field_derivative Gamma z * Digamma z) (at z within A)"
+  unfolding Gamma_def [abs_def]
+  by (fastforce intro!: derivative_eq_intros simp: rGamma_eq_zero_iff)
+
+declare has_field_derivative_Gamma[THEN DERIV_chain2, derivative_intros]
+
+(* TODO: Hide ugly facts properly *)
+hide_fact rGamma_eq_zero_iff_aux differentiable_rGamma_aux1 differentiable_rGamma_aux2
+          differentiable_rGamma_aux2 rGamma_series_aux Gamma_class.rGamma_eq_zero_iff_aux
+
+
+
+(* TODO: differentiable etc. *)
+
+
+subsection \<open>Continuity\<close>
+
+lemma continuous_on_rGamma [continuous_intros]: "continuous_on A rGamma"
+  by (rule DERIV_continuous_on has_field_derivative_rGamma)+
+
+lemma continuous_on_Gamma [continuous_intros]: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> continuous_on A Gamma"
+  by (rule DERIV_continuous_on has_field_derivative_Gamma)+ blast
+
+lemma isCont_rGamma [continuous_intros]: 
+  "isCont f z \<Longrightarrow> isCont (\<lambda>x. rGamma (f x)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_rGamma]])
+
+lemma isCont_Gamma [continuous_intros]: 
+  "isCont f z \<Longrightarrow> f z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> isCont (\<lambda>x. Gamma (f x)) z"
+  by (rule isCont_o2[OF _  DERIV_isCont[OF has_field_derivative_Gamma]])
+
+
+
+text \<open>The complex Gamma function\<close>
+
+instantiation complex :: Gamma
+begin
+
+definition rGamma_complex :: "complex \<Rightarrow> complex" where
+  "rGamma_complex z = lim (rGamma_series z)"
+
+lemma rGamma_series_complex_converges: 
+        "convergent (rGamma_series (z :: complex))" (is "?thesis1")
+  and rGamma_complex_altdef:
+        "rGamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (-ln_Gamma z))" (is "?thesis2")
+proof -
+  have "?thesis1 \<and> ?thesis2"
+  proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+    case False
+    have "rGamma_series z \<longlonglongrightarrow> exp (- ln_Gamma z)"
+    proof (rule Lim_transform_eventually)
+      from ln_Gamma_series_complex_converges'[OF False] guess d by (elim exE conjE)
+      from this(1) uniformly_convergent_imp_convergent[OF this(2), of z] 
+        have "ln_Gamma_series z \<longlonglongrightarrow> lim (ln_Gamma_series z)" by (simp add: convergent_LIMSEQ_iff)
+      thus "(\<lambda>n. exp (-ln_Gamma_series z n)) \<longlonglongrightarrow> exp (- ln_Gamma z)"
+        unfolding convergent_def ln_Gamma_def by (intro tendsto_exp tendsto_minus)
+      from eventually_gt_at_top[of "0::nat"] exp_ln_Gamma_series_complex False
+        show "eventually (\<lambda>n. exp (-ln_Gamma_series z n) = rGamma_series z n) sequentially"
+        by (force elim!: eventually_mono simp: exp_minus Gamma_series_def rGamma_series_def)
+    qed
+    with False show ?thesis
+      by (auto simp: convergent_def rGamma_complex_def intro!: limI)
+  next
+    case True
+    then obtain k where "z = - of_nat k" by (erule nonpos_Ints_cases')
+    also have "rGamma_series \<dots> \<longlonglongrightarrow> 0"
+      by (subst tendsto_cong[OF rGamma_series_minus_of_nat]) (simp_all add: convergent_const)
+    finally show ?thesis using True
+      by (auto simp: rGamma_complex_def convergent_def intro!: limI)
+  qed
+  thus "?thesis1" "?thesis2" by blast+
+qed
+
+context
+begin
+
+(* TODO: duplication *)
+private lemma rGamma_complex_plus1: "z * rGamma (z + 1) = rGamma (z :: complex)"
+proof -
+  let ?f = "\<lambda>n. (z + 1) * inverse (of_nat n) + 1"
+  have "eventually (\<lambda>n. ?f n * rGamma_series z n = z * rGamma_series (z + 1) n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    hence "z * rGamma_series (z + 1) n = inverse (of_nat n) *
+             pochhammer z (Suc (Suc n)) / (fact n * exp (z * of_real (ln (of_nat n))))"
+      by (subst pochhammer_rec) (simp add: rGamma_series_def field_simps exp_add exp_of_real)
+    also from n have "\<dots> = ?f n * rGamma_series z n"
+      by (subst pochhammer_rec') (simp_all add: divide_simps rGamma_series_def add_ac)
+    finally show "?f n * rGamma_series z n = z * rGamma_series (z + 1) n" ..
+  qed
+  moreover have "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> ((z+1) * 0 + 1) * rGamma z"
+    using rGamma_series_complex_converges
+    by (intro tendsto_intros lim_inverse_n) 
+       (simp_all add: convergent_LIMSEQ_iff rGamma_complex_def)
+  hence "(\<lambda>n. ?f n * rGamma_series z n) \<longlonglongrightarrow> rGamma z" by simp
+  ultimately have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> rGamma z"
+    by (rule Lim_transform_eventually)
+  moreover have "(\<lambda>n. z * rGamma_series (z + 1) n) \<longlonglongrightarrow> z * rGamma (z + 1)"
+    using rGamma_series_complex_converges 
+    by (auto intro!: tendsto_mult simp: rGamma_complex_def convergent_LIMSEQ_iff)
+  ultimately show "z * rGamma (z + 1) = rGamma z" using LIMSEQ_unique by blast
+qed
+
+private lemma has_field_derivative_rGamma_complex_no_nonpos_Int:
+  assumes "(z :: complex) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
+proof -
+  have diff: "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)" if "Re z > 0" for z 
+  proof (subst DERIV_cong_ev[OF refl _ refl])
+    from that have "eventually (\<lambda>t. t \<in> ball z (Re z/2)) (nhds z)" 
+      by (intro eventually_nhds_in_nhd) simp_all
+    thus "eventually (\<lambda>t. rGamma t = exp (- ln_Gamma t)) (nhds z)"
+      using no_nonpos_Int_in_ball_complex[OF that]
+      by (auto elim!: eventually_mono simp: rGamma_complex_altdef)
+  next
+    show "((\<lambda>t. exp (- ln_Gamma t)) has_field_derivative (-rGamma z * Digamma z)) (at z)"
+      using that by (force elim!: nonpos_Ints_cases intro!: derivative_eq_intros 
+                           simp: rGamma_complex_altdef)
+  qed
+  
+  from assms show "(rGamma has_field_derivative - rGamma z * Digamma z) (at z)"
+  proof (induction "nat \<lfloor>1 - Re z\<rfloor>" arbitrary: z)
+    case (Suc n z)
+    from Suc.prems have z: "z \<noteq> 0" by auto
+    from Suc.hyps have "n = nat \<lfloor>- Re z\<rfloor>" by linarith
+    hence A: "n = nat \<lfloor>1 - Re (z + 1)\<rfloor>" by simp
+    from Suc.prems have B: "z + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by (force dest: plus_one_in_nonpos_Ints_imp)
+    
+    have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1)) z) has_field_derivative 
+      -rGamma (z + 1) * (Digamma (z + 1) * z - 1)) (at z)"
+      by (rule derivative_eq_intros DERIV_chain Suc refl A B)+ (simp add: algebra_simps)
+    also have "(\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) = rGamma" 
+      by (simp add: rGamma_complex_plus1)
+    also from z have "Digamma (z + 1) * z - 1 = z * Digamma z"
+      by (subst Digamma_plus1) (simp_all add: field_simps)
+    also have "-rGamma (z + 1) * (z * Digamma z) = -rGamma z * Digamma z"
+      by (simp add: rGamma_complex_plus1[of z, symmetric])
+    finally show ?case .
+  qed (intro diff, simp)
+qed
+
+private lemma rGamma_complex_1: "rGamma (1 :: complex) = 1"
+proof -
+  have A: "eventually (\<lambda>n. rGamma_series 1 n = of_nat (Suc n) / of_nat n) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+    by (force elim!: eventually_mono simp: rGamma_series_def exp_of_real pochhammer_fact 
+                    divide_simps pochhammer_rec' dest!: pochhammer_eq_0_imp_nonpos_Int)
+  have "rGamma_series 1 \<longlonglongrightarrow> 1" by (subst tendsto_cong[OF A]) (rule LIMSEQ_Suc_n_over_n)
+  thus "rGamma 1 = (1 :: complex)" unfolding rGamma_complex_def by (rule limI)
+qed
+
+private lemma has_field_derivative_rGamma_complex_nonpos_Int:
+  "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: complex))"
+proof (induction n)
+  case 0
+  have A: "(0::complex) + 1 \<notin> \<int>\<^sub>\<le>\<^sub>0" by simp
+  have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative 1) (at 0)"
+    by (rule derivative_eq_intros DERIV_chain refl
+             has_field_derivative_rGamma_complex_no_nonpos_Int A)+ (simp add: rGamma_complex_1)
+    thus ?case by (simp add: rGamma_complex_plus1)
+next
+  case (Suc n)
+  hence A: "(rGamma has_field_derivative (-1)^n * fact n)  
+                (at (- of_nat (Suc n) + 1 :: complex))" by simp
+   have "((\<lambda>z. z * (rGamma \<circ> (\<lambda>z. z + 1 :: complex)) z) has_field_derivative 
+             (- 1) ^ Suc n * fact (Suc n)) (at (- of_nat (Suc n)))"
+     by (rule derivative_eq_intros refl A DERIV_chain)+ 
+        (simp add: algebra_simps rGamma_complex_altdef)
+  thus ?case by (simp add: rGamma_complex_plus1)
+qed
+
+instance proof
+  fix z :: complex show "(rGamma z = 0) \<longleftrightarrow> (\<exists>n. z = - of_nat n)"
+    by (auto simp: rGamma_complex_altdef elim!: nonpos_Ints_cases')
+next
+  fix z :: complex assume "\<And>n. z \<noteq> - of_nat n"
+  hence "z \<notin> \<int>\<^sub>\<le>\<^sub>0" by (auto elim!: nonpos_Ints_cases')
+  from has_field_derivative_rGamma_complex_no_nonpos_Int[OF this]
+    show "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (z + of_nat k))
+                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma z + 
+              rGamma z * d * (y - z)) /\<^sub>R  cmod (y - z)) \<midarrow>z\<rightarrow> 0"
+      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
+                    netlimit_at of_real_def[symmetric] suminf_def)
+next
+  fix n :: nat
+  from has_field_derivative_rGamma_complex_nonpos_Int[of n]
+  show "let z = - of_nat n in (\<lambda>y. (rGamma y - rGamma z - (- 1) ^ n * setprod of_nat {1..n} *
+                  (y - z)) /\<^sub>R cmod (y - z)) \<midarrow>z\<rightarrow> 0"
+    by (simp add: has_field_derivative_def has_derivative_def fact_altdef netlimit_at Let_def)
+next
+  fix z :: complex
+  from rGamma_series_complex_converges[of z] have "rGamma_series z \<longlonglongrightarrow> rGamma z"
+    by (simp add: convergent_LIMSEQ_iff rGamma_complex_def)
+  thus "let fact' = \<lambda>n. setprod of_nat {1..n}; 
+            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
+            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
+        in  (\<lambda>n. pochhammer' z n / (fact' n * exp (z * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma z"
+    by (simp add: fact_altdef pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def
+                  of_real_def [symmetric] suminf_def sums_def [abs_def])
+qed
+
+end
+end
+
+
+lemma Gamma_complex_altdef: 
+  "Gamma z = (if z \<in> \<int>\<^sub>\<le>\<^sub>0 then 0 else exp (ln_Gamma (z :: complex)))"
+  unfolding Gamma_def rGamma_complex_altdef by (simp add: exp_minus)
+
+lemma cnj_rGamma: "cnj (rGamma z) = rGamma (cnj z)"
+proof -
+  have "rGamma_series (cnj z) = (\<lambda>n. cnj (rGamma_series z n))"
+    by (intro ext) (simp_all add: rGamma_series_def exp_cnj)
+  also have "... \<longlonglongrightarrow> cnj (rGamma z)" by (intro tendsto_cnj tendsto_intros)
+  finally show ?thesis unfolding rGamma_complex_def by (intro sym[OF limI])
+qed
+
+lemma cnj_Gamma: "cnj (Gamma z) = Gamma (cnj z)"
+  unfolding Gamma_def by (simp add: cnj_rGamma)
+  
+lemma Gamma_complex_real: 
+  "z \<in> \<real> \<Longrightarrow> Gamma z \<in> (\<real> :: complex set)" and rGamma_complex_real: "z \<in> \<real> \<Longrightarrow> rGamma z \<in> \<real>"
+  by (simp_all add: Reals_cnj_iff cnj_Gamma cnj_rGamma)
+
+lemma complex_differentiable_rGamma: "rGamma complex_differentiable (at z within A)"
+  using has_field_derivative_rGamma[of z] unfolding complex_differentiable_def by blast
+
+lemma holomorphic_on_rGamma: "rGamma holomorphic_on A"
+  unfolding holomorphic_on_def by (auto intro!: complex_differentiable_rGamma)
+
+lemma analytic_on_rGamma: "rGamma analytic_on A"
+  unfolding analytic_on_def by (auto intro!: exI[of _ 1] holomorphic_on_rGamma)
+
+
+lemma complex_differentiable_Gamma: "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Gamma complex_differentiable (at z within A)"
+  using has_field_derivative_Gamma[of z] unfolding complex_differentiable_def by auto
+
+lemma holomorphic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma holomorphic_on A"
+  unfolding holomorphic_on_def by (auto intro!: complex_differentiable_Gamma)
+
+lemma analytic_on_Gamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Gamma analytic_on A"
+  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
+     (auto intro!: holomorphic_on_Gamma)
+
+lemma has_field_derivative_rGamma_complex' [derivative_intros]:
+  "(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 
+        -rGamma z * Digamma z)) (at z within A)"
+  using has_field_derivative_rGamma[of z] by (auto elim!: nonpos_Ints_cases')
+
+declare has_field_derivative_rGamma_complex'[THEN DERIV_chain2, derivative_intros]
+
+
+lemma complex_differentiable_Polygamma:
+  "z \<notin> \<int>\<^sub>\<le>\<^sub>0 \<Longrightarrow> Polygamma n complex_differentiable (at z within A)"
+  using has_field_derivative_Polygamma[of z n] unfolding complex_differentiable_def by auto
+
+lemma holomorphic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n holomorphic_on A"
+  unfolding holomorphic_on_def by (auto intro!: complex_differentiable_Polygamma)
+
+lemma analytic_on_Polygamma: "A \<inter> \<int>\<^sub>\<le>\<^sub>0 = {} \<Longrightarrow> Polygamma n analytic_on A"
+  by (rule analytic_on_subset[of _ "UNIV - \<int>\<^sub>\<le>\<^sub>0"], subst analytic_on_open)
+     (auto intro!: holomorphic_on_Polygamma)
+
+
+
+text \<open>The real Gamma function\<close>
+
+lemma rGamma_series_real:
+  "eventually (\<lambda>n. rGamma_series x n = Re (rGamma_series (of_real x) n)) sequentially"
+  using eventually_gt_at_top[of "0 :: nat"]
+proof eventually_elim
+  fix n :: nat assume n: "n > 0"
+  have "Re (rGamma_series (of_real x) n) =
+          Re (of_real (pochhammer x (Suc n)) / (fact n * exp (of_real (x * ln (real_of_nat n)))))"
+    using n by (simp add: rGamma_series_def powr_def Ln_of_nat pochhammer_of_real)
+  also from n have "\<dots> = Re (of_real ((pochhammer x (Suc n)) /
+                              (fact n * (exp (x * ln (real_of_nat n))))))"
+    by (subst exp_of_real) simp
+  also from n have "\<dots> = rGamma_series x n"
+    by (subst Re_complex_of_real) (simp add: rGamma_series_def powr_def)
+  finally show "rGamma_series x n = Re (rGamma_series (of_real x) n)" ..
+qed
+
+instantiation real :: Gamma
+begin
+
+definition "rGamma_real x = Re (rGamma (of_real x :: complex))"
+
+instance proof
+  fix x :: real
+  have "rGamma x = Re (rGamma (of_real x))" by (simp add: rGamma_real_def)
+  also have "of_real \<dots> = rGamma (of_real x :: complex)"
+    by (intro of_real_Re rGamma_complex_real) simp_all
+  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)
+  also have "\<dots> \<longleftrightarrow> (\<exists>n. x = - of_nat n)" by (auto elim!: nonpos_Ints_cases')
+  finally show "(rGamma x) = 0 \<longleftrightarrow> (\<exists>n. x = - real_of_nat n)" by simp
+next
+  fix x :: real assume "\<And>n. x \<noteq> - of_nat n"
+  hence "complex_of_real x \<notin> \<int>\<^sub>\<le>\<^sub>0" 
+    by (subst of_real_in_nonpos_Ints_iff) (auto elim!: nonpos_Ints_cases')
+  moreover from this have "x \<noteq> 0" by auto
+  ultimately have "(rGamma has_field_derivative - rGamma x * Digamma x) (at x)"
+    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex 
+                  simp: Polygamma_of_real rGamma_real_def [abs_def])
+  thus "let d = (THE d. (\<lambda>n. \<Sum>k<n. inverse (of_nat (Suc k)) - inverse (x + of_nat k))
+                       \<longlonglongrightarrow> d) - euler_mascheroni *\<^sub>R 1 in  (\<lambda>y. (rGamma y - rGamma x + 
+              rGamma x * d * (y - x)) /\<^sub>R  norm (y - x)) \<midarrow>x\<rightarrow> 0"
+      by (simp add: has_field_derivative_def has_derivative_def Digamma_def sums_def [abs_def]
+                    netlimit_at of_real_def[symmetric] suminf_def)
+next
+  fix n :: nat
+  have "(rGamma has_field_derivative (-1)^n * fact n) (at (- of_nat n :: real))"
+    by (fastforce intro!: derivative_eq_intros has_vector_derivative_real_complex 
+                  simp: Polygamma_of_real rGamma_real_def [abs_def])
+  thus "let x = - of_nat n in (\<lambda>y. (rGamma y - rGamma x - (- 1) ^ n * setprod of_nat {1..n} *
+                  (y - x)) /\<^sub>R norm (y - x)) \<midarrow>x::real\<rightarrow> 0"
+    by (simp add: has_field_derivative_def has_derivative_def fact_altdef netlimit_at Let_def)
+next
+  fix x :: real
+  have "rGamma_series x \<longlonglongrightarrow> rGamma x"
+  proof (rule Lim_transform_eventually)
+    show "(\<lambda>n. Re (rGamma_series (of_real x) n)) \<longlonglongrightarrow> rGamma x" unfolding rGamma_real_def
+      by (intro tendsto_intros)
+  qed (insert rGamma_series_real, simp add: eq_commute)
+  thus "let fact' = \<lambda>n. setprod of_nat {1..n}; 
+            exp = \<lambda>x. THE e. (\<lambda>n. \<Sum>k<n. x ^ k /\<^sub>R fact k) \<longlonglongrightarrow> e;
+            pochhammer' = \<lambda>a n. \<Prod>n = 0..n. a + of_nat n
+        in  (\<lambda>n. pochhammer' x n / (fact' n * exp (x * ln (real_of_nat n) *\<^sub>R 1))) \<longlonglongrightarrow> rGamma x"
+    by (simp add: fact_altdef pochhammer_Suc_setprod rGamma_series_def [abs_def] exp_def
+                  of_real_def [symmetric] suminf_def sums_def [abs_def])
+qed
+
+end
+
+
+lemma rGamma_complex_of_real: "rGamma (complex_of_real x) = complex_of_real (rGamma x)"
+  unfolding rGamma_real_def using rGamma_complex_real by simp
+
+lemma Gamma_complex_of_real: "Gamma (complex_of_real x) = complex_of_real (Gamma x)"
+  unfolding Gamma_def by (simp add: rGamma_complex_of_real)
+
+lemma rGamma_real_altdef: "rGamma x = lim (rGamma_series (x :: real))"
+  by (rule sym, rule limI, rule tendsto_intros)
+
+lemma Gamma_real_altdef1: "Gamma x = lim (Gamma_series (x :: real))"
+  by (rule sym, rule limI, rule tendsto_intros)
+
+lemma Gamma_real_altdef2: "Gamma x = Re (Gamma (of_real x))"
+  using rGamma_complex_real[OF Reals_of_real[of x]]
+  by (elim Reals_cases) 
+     (simp only: Gamma_def rGamma_real_def of_real_inverse[symmetric] Re_complex_of_real)
+
+lemma ln_Gamma_series_complex_of_real: 
+  "x > 0 \<Longrightarrow> n > 0 \<Longrightarrow> ln_Gamma_series (complex_of_real x) n = of_real (ln_Gamma_series x n)"
+proof -
+  assume xn: "x > 0" "n > 0"
+  have "Ln (complex_of_real x / of_nat k + 1) = of_real (ln (x / of_nat k + 1))" if "k \<ge> 1" for k
+    using that xn by (subst Ln_of_real [symmetric]) (auto intro!: add_nonneg_pos simp: field_simps)
+  with xn show ?thesis by (simp add: ln_Gamma_series_def Ln_of_nat Ln_of_real)
+qed
+
+lemma ln_Gamma_real_converges: 
+  assumes "(x::real) > 0" 
+  shows   "convergent (ln_Gamma_series x)"
+proof -
+  have "(\<lambda>n. ln_Gamma_series (complex_of_real x) n) \<longlonglongrightarrow> ln_Gamma (of_real x)" using assms
+    by (intro ln_Gamma_complex_LIMSEQ) (auto simp: of_real_in_nonpos_Ints_iff)
+  moreover from eventually_gt_at_top[of "0::nat"]
+    have "eventually (\<lambda>n. complex_of_real (ln_Gamma_series x n) = 
+            ln_Gamma_series (complex_of_real x) n) sequentially" 
+    by eventually_elim (simp add: ln_Gamma_series_complex_of_real assms)
+  ultimately have "(\<lambda>n. complex_of_real (ln_Gamma_series x n)) \<longlonglongrightarrow> ln_Gamma (of_real x)"
+    by (subst tendsto_cong) assumption+
+  from tendsto_Re[OF this] show ?thesis by (auto simp: convergent_def)
+qed
+
+lemma ln_Gamma_real_LIMSEQ: "(x::real) > 0 \<Longrightarrow> ln_Gamma_series x \<longlonglongrightarrow> ln_Gamma x"
+  using ln_Gamma_real_converges[of x] unfolding ln_Gamma_def by (simp add: convergent_LIMSEQ_iff)
+
+lemma ln_Gamma_complex_of_real: "x > 0 \<Longrightarrow> ln_Gamma (complex_of_real x) = of_real (ln_Gamma x)"
+proof (unfold ln_Gamma_def, rule limI, rule Lim_transform_eventually)
+  assume x: "x > 0"
+  show "eventually (\<lambda>n. of_real (ln_Gamma_series x n) = 
+            ln_Gamma_series (complex_of_real x) n) sequentially" 
+    using eventually_gt_at_top[of "0::nat"] 
+    by eventually_elim (simp add: ln_Gamma_series_complex_of_real x)
+qed (intro tendsto_of_real, insert ln_Gamma_real_LIMSEQ[of x], simp add: ln_Gamma_def)
+     
+lemma Gamma_real_pos_exp: "x > (0 :: real) \<Longrightarrow> Gamma x = exp (ln_Gamma x)"
+  by (auto simp: Gamma_real_altdef2 Gamma_complex_altdef of_real_in_nonpos_Ints_iff 
+                 ln_Gamma_complex_of_real exp_of_real)
+
+lemma ln_Gamma_real_pos: "x > 0 \<Longrightarrow> ln_Gamma x = ln (Gamma x :: real)"
+  unfolding Gamma_real_pos_exp by simp
+
+lemma Gamma_real_pos: "x > (0::real) \<Longrightarrow> Gamma x > 0"
+  by (simp add: Gamma_real_pos_exp)
+  
+lemma has_field_derivative_ln_Gamma_real [derivative_intros]:
+  assumes x: "x > (0::real)"
+  shows "(ln_Gamma has_field_derivative Digamma x) (at x)"
+proof (subst DERIV_cong_ev[OF refl _ refl])
+  from assms show "((Re \<circ> ln_Gamma \<circ> complex_of_real) has_field_derivative Digamma x) (at x)"
+    by (auto intro!: derivative_eq_intros has_vector_derivative_real_complex
+             simp: Polygamma_of_real o_def)
+  from eventually_nhds_in_nhd[of x "{0<..}"] assms
+    show "eventually (\<lambda>y. ln_Gamma y = (Re \<circ> ln_Gamma \<circ> of_real) y) (nhds x)"
+    by (auto elim!: eventually_mono simp: ln_Gamma_complex_of_real interior_open)
+qed
+
+declare has_field_derivative_ln_Gamma_real[THEN DERIV_chain2, derivative_intros]
+
+
+lemma has_field_derivative_rGamma_real' [derivative_intros]:
+  "(rGamma has_field_derivative (if x \<in> \<int>\<^sub>\<le>\<^sub>0 then (-1)^(nat \<lfloor>-x\<rfloor>) * fact (nat \<lfloor>-x\<rfloor>) else 
+        -rGamma x * Digamma x)) (at x within A)"
+  using has_field_derivative_rGamma[of x] by (force elim!: nonpos_Ints_cases')
+
+declare has_field_derivative_rGamma_real'[THEN DERIV_chain2, derivative_intros]
+
+lemma Polygamma_real_odd_pos:
+  assumes "(x::real) \<notin> \<int>\<^sub>\<le>\<^sub>0" "odd n"
+  shows   "Polygamma n x > 0"
+proof -
+  from assms have "x \<noteq> 0" by auto
+  with assms show ?thesis
+    unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
+    by (auto simp: zero_less_power_eq simp del: power_Suc 
+             dest: plus_of_nat_eq_0_imp intro!: mult_pos_pos suminf_pos)
+qed
+
+lemma Polygamma_real_even_neg:
+  assumes "(x::real) > 0" "n > 0" "even n"
+  shows   "Polygamma n x < 0"
+  using assms unfolding Polygamma_def using Polygamma_converges'[of x "Suc n"]
+  by (auto intro!: mult_pos_pos suminf_pos)
+  
+lemma Polygamma_real_strict_mono:
+  assumes "x > 0" "x < (y::real)" "even n"
+  shows   "Polygamma n x < Polygamma n y"
+proof -
+  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
+    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
+  then guess \<xi> by (elim exE conjE) note \<xi> = this
+  note \<xi>(3)
+  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> > 0" 
+    by (intro mult_pos_pos Polygamma_real_odd_pos) (auto elim!: nonpos_Ints_cases)
+  finally show ?thesis by simp
+qed
+
+lemma Polygamma_real_strict_antimono:
+  assumes "x > 0" "x < (y::real)" "odd n"
+  shows   "Polygamma n x > Polygamma n y"
+proof -
+  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> Polygamma n y - Polygamma n x = (y - x) * Polygamma (Suc n) \<xi>"
+    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
+  then guess \<xi> by (elim exE conjE) note \<xi> = this
+  note \<xi>(3)
+  also from \<xi>(1,2) assms have "(y - x) * Polygamma (Suc n) \<xi> < 0" 
+    by (intro mult_pos_neg Polygamma_real_even_neg) simp_all
+  finally show ?thesis by simp
+qed
+
+lemma Polygamma_real_mono:
+  assumes "x > 0" "x \<le> (y::real)" "even n"
+  shows   "Polygamma n x \<le> Polygamma n y"
+  using Polygamma_real_strict_mono[OF assms(1) _ assms(3), of y] assms(2) 
+  by (cases "x = y") simp_all
+
+lemma Digamma_real_ge_three_halves_pos:
+  assumes "x \<ge> 3/2"
+  shows   "Digamma (x :: real) > 0"
+proof -
+  have "0 < Digamma (3/2 :: real)" by (fact Digamma_real_three_halves_pos)
+  also from assms have "\<dots> \<le> Digamma x" by (intro Polygamma_real_mono) simp_all
+  finally show ?thesis .
+qed
+
+lemma ln_Gamma_real_strict_mono:
+  assumes "x \<ge> 3/2" "x < y"
+  shows   "ln_Gamma (x :: real) < ln_Gamma y"
+proof -
+  have "\<exists>\<xi>. x < \<xi> \<and> \<xi> < y \<and> ln_Gamma y - ln_Gamma x = (y - x) * Digamma \<xi>"
+    using assms by (intro MVT2 derivative_intros impI allI) (auto elim!: nonpos_Ints_cases)
+  then guess \<xi> by (elim exE conjE) note \<xi> = this
+  note \<xi>(3)
+  also from \<xi>(1,2) assms have "(y - x) * Digamma \<xi> > 0" 
+    by (intro mult_pos_pos Digamma_real_ge_three_halves_pos) simp_all
+  finally show ?thesis by simp
+qed  
+  
+lemma Gamma_real_strict_mono:
+  assumes "x \<ge> 3/2" "x < y"
+  shows   "Gamma (x :: real) < Gamma y"
+proof -
+  from Gamma_real_pos_exp[of x] assms have "Gamma x = exp (ln_Gamma x)" by simp
+  also have "\<dots> < exp (ln_Gamma y)" by (intro exp_less_mono ln_Gamma_real_strict_mono assms)
+  also from Gamma_real_pos_exp[of y] assms have "\<dots> = Gamma y" by simp
+  finally show ?thesis .
+qed
+
+lemma log_convex_Gamma_real: "convex_on {0<..} (ln \<circ> Gamma :: real \<Rightarrow> real)"
+  by (rule convex_on_realI[of _ _ Digamma])
+     (auto intro!: derivative_eq_intros Polygamma_real_mono Gamma_real_pos 
+           simp: o_def Gamma_eq_zero_iff elim!: nonpos_Ints_cases')
+
+
+subsection \<open>Beta function\<close>
+
+definition Beta where "Beta a b = Gamma a * Gamma b / Gamma (a + b)"
+
+lemma Beta_altdef: "Beta a b = Gamma a * Gamma b * rGamma (a + b)"
+  by (simp add: inverse_eq_divide Beta_def Gamma_def)
+
+lemma Beta_commute: "Beta a b = Beta b a"
+  unfolding Beta_def by (simp add: ac_simps)
+
+lemma has_field_derivative_Beta1 [derivative_intros]:
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "((\<lambda>x. Beta x y) has_field_derivative (Beta x y * (Digamma x - Digamma (x + y)))) 
+               (at x within A)" unfolding Beta_altdef
+  by (rule DERIV_cong, (rule derivative_intros assms)+) (simp add: algebra_simps)
+
+lemma has_field_derivative_Beta2 [derivative_intros]:
+  assumes "y \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "((\<lambda>y. Beta x y) has_field_derivative (Beta x y * (Digamma y - Digamma (x + y)))) 
+               (at y within A)"
+  using has_field_derivative_Beta1[of y x A] assms by (simp add: Beta_commute add_ac)
+
+lemma Beta_plus1_plus1: 
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0" 
+  shows   "Beta (x + 1) y + Beta x (y + 1) = Beta x y"
+proof -
+  have "Beta (x + 1) y + Beta x (y + 1) = 
+            (Gamma (x + 1) * Gamma y + Gamma x * Gamma (y + 1)) * rGamma ((x + y) + 1)"
+    by (simp add: Beta_altdef add_divide_distrib algebra_simps)
+  also have "\<dots> = (Gamma x * Gamma y) * ((x + y) * rGamma ((x + y) + 1))"
+    by (subst assms[THEN Gamma_plus1])+ (simp add: algebra_simps)
+  also from assms have "\<dots> = Beta x y" unfolding Beta_altdef by (subst rGamma_plus1) simp
+  finally show ?thesis .
+qed
+
+lemma Beta_plus1_left: 
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(x + y) * Beta (x + 1) y = x * Beta x y"
+proof -
+  have "(x + y) * Beta (x + 1) y = Gamma (x + 1) * Gamma y * ((x + y) * rGamma ((x + y) + 1))"
+    unfolding Beta_altdef by (simp only: ac_simps)
+  also have "\<dots> = x * Beta x y" unfolding Beta_altdef
+     by (subst assms[THEN Gamma_plus1] rGamma_plus1)+ (simp only: ac_simps)
+  finally show ?thesis .
+qed
+
+lemma Beta_plus1_right: 
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "(x + y) * Beta x (y + 1) = y * Beta x y"
+  using Beta_plus1_left[of y x] assms by (simp add: Beta_commute add.commute)
+
+lemma Gamma_Gamma_Beta:
+  assumes "x \<notin> \<int>\<^sub>\<le>\<^sub>0" "y \<notin> \<int>\<^sub>\<le>\<^sub>0" "x + y \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows   "Gamma x * Gamma y = Beta x y * Gamma (x + y)"
+  unfolding Beta_altdef using assms Gamma_eq_zero_iff[of "x+y"]
+  by (simp add: rGamma_inverse_Gamma)
+
+
+
+subsection \<open>Legendre duplication theorem\<close>
+
+context
+begin
+
+private lemma Gamma_legendre_duplication_aux:
+  fixes z :: "'a :: Gamma"
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "Gamma z * Gamma (z + 1/2) = exp ((1 - 2*z) * of_real (ln 2)) * Gamma (1/2) * Gamma (2*z)"
+proof -
+  let ?powr = "\<lambda>b a. exp (a * of_real (ln (of_nat b)))"
+  let ?h = "\<lambda>n. (fact (n-1))\<^sup>2 / fact (2*n-1) * of_nat (2^(2*n)) * 
+                exp (1/2 * of_real (ln (real_of_nat n)))"
+  {
+    fix z :: 'a assume z: "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+    let ?g = "\<lambda>n. ?powr 2 (2*z) * Gamma_series' z n * Gamma_series' (z + 1/2) n / 
+                      Gamma_series' (2*z) (2*n)"
+    have "eventually (\<lambda>n. ?g n = ?h n) sequentially" using eventually_gt_at_top
+    proof eventually_elim
+      fix n :: nat assume n: "n > 0"
+      let ?f = "fact (n - 1) :: 'a" and ?f' = "fact (2*n - 1) :: 'a"
+      have A: "exp t * exp t = exp (2*t :: 'a)" for t by (subst exp_add [symmetric]) simp
+      have A: "Gamma_series' z n * Gamma_series' (z + 1/2) n = ?f^2 * ?powr n (2*z + 1/2) /
+                (pochhammer z n * pochhammer (z + 1/2) n)"
+        by (simp add: Gamma_series'_def exp_add ring_distribs power2_eq_square A mult_ac)
+      have B: "Gamma_series' (2*z) (2*n) = 
+                       ?f' * ?powr 2 (2*z) * ?powr n (2*z) / 
+                       (of_nat (2^(2*n)) * pochhammer z n * pochhammer (z+1/2) n)" using n
+        by (simp add: Gamma_series'_def ln_mult exp_add ring_distribs pochhammer_double)
+      from z have "pochhammer z n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
+      moreover from z have "pochhammer (z + 1/2) n \<noteq> 0" by (auto dest: pochhammer_eq_0_imp_nonpos_Int)
+      ultimately have "?powr 2 (2*z) * (Gamma_series' z n * Gamma_series' (z + 1/2) n) / Gamma_series' (2*z) (2*n) = 
+         ?f^2 / ?f' * of_nat (2^(2*n)) * (?powr n ((4*z + 1)/2) * ?powr n (-2*z))"
+        using n unfolding A B by (simp add: divide_simps exp_minus)
+      also have "?powr n ((4*z + 1)/2) * ?powr n (-2*z) = ?powr n (1/2)"
+        by (simp add: algebra_simps exp_add[symmetric] add_divide_distrib)
+      finally show "?g n = ?h n" by (simp only: mult_ac)
+    qed
+    
+    moreover from z double_in_nonpos_Ints_imp[of z] have "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto     
+    hence "?g \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
+      using lim_subseq[of "op * 2", OF _ Gamma_series'_LIMSEQ, of "2*z"]
+      by (intro tendsto_intros Gamma_series'_LIMSEQ)
+         (simp_all add: o_def subseq_def Gamma_eq_zero_iff)
+    ultimately have "?h \<longlonglongrightarrow> ?powr 2 (2*z) * Gamma z * Gamma (z+1/2) / Gamma (2*z)"
+      by (rule Lim_transform_eventually)
+  } note lim = this
+  
+  from assms double_in_nonpos_Ints_imp[of z] have z': "2 * z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto  
+  from fraction_not_in_ints[of 2 1] have "(1/2 :: 'a) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+    by (intro not_in_Ints_imp_not_in_nonpos_Ints) simp_all
+  with lim[of "1/2 :: 'a"] have "?h \<longlonglongrightarrow> 2 * Gamma (1 / 2 :: 'a)" by (simp add: exp_of_real)
+  from LIMSEQ_unique[OF this lim[OF assms]] z' show ?thesis
+    by (simp add: divide_simps Gamma_eq_zero_iff ring_distribs exp_diff exp_of_real ac_simps)
+qed
+
+(* TODO: perhaps this is unnecessary once we have the fact that a holomorphic function is 
+   infinitely differentiable *)
+private lemma Gamma_reflection_aux:
+  defines "h \<equiv> \<lambda>z::complex. if z \<in> \<int> then 0 else 
+                 (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
+  defines "a \<equiv> complex_of_real pi"
+  obtains h' where "continuous_on UNIV h'" "\<And>z. (h has_field_derivative (h' z)) (at z)"
+proof -
+  def f \<equiv> "\<lambda>n. a * of_real (cos_coeff (n+1) - sin_coeff (n+2))"
+  def F \<equiv> "\<lambda>z. if z = 0 then 0 else (cos (a*z) - sin (a*z)/(a*z)) / z"
+  def g \<equiv> "\<lambda>n. complex_of_real (sin_coeff (n+1))"
+  def G \<equiv> "\<lambda>z. if z = 0 then 1 else sin (a*z)/(a*z)"
+  have a_nz: "a \<noteq> 0" unfolding a_def by simp
+
+  have "(\<lambda>n. f n * (a*z)^n) sums (F z) \<and> (\<lambda>n. g n * (a*z)^n) sums (G z)" 
+    if "abs (Re z) < 1" for z
+  proof (cases "z = 0"; rule conjI)
+    assume "z \<noteq> 0"
+    note z = this that
+
+    from z have sin_nz: "sin (a*z) \<noteq> 0" unfolding a_def by (auto simp: sin_eq_0)
+    have "(\<lambda>n. of_real (sin_coeff n) * (a*z)^n) sums (sin (a*z))" using sin_converges[of "a*z"]
+      by (simp add: scaleR_conv_of_real)
+    from sums_split_initial_segment[OF this, of 1] 
+      have "(\<lambda>n. (a*z) * of_real (sin_coeff (n+1)) * (a*z)^n) sums (sin (a*z))" by (simp add: mult_ac)
+    from sums_mult[OF this, of "inverse (a*z)"] z a_nz
+      have A: "(\<lambda>n. g n * (a*z)^n) sums (sin (a*z)/(a*z))"
+      by (simp add: field_simps g_def)
+    with z show "(\<lambda>n. g n * (a*z)^n) sums (G z)" by (simp add: G_def)
+    from A z a_nz sin_nz have g_nz: "(\<Sum>n. g n * (a*z)^n) \<noteq> 0" by (simp add: sums_iff g_def)
+  
+    have [simp]: "sin_coeff (Suc 0) = 1" by (simp add: sin_coeff_def)
+    from sums_split_initial_segment[OF sums_diff[OF cos_converges[of "a*z"] A], of 1]
+    have "(\<lambda>n. z * f n * (a*z)^n) sums (cos (a*z) - sin (a*z) / (a*z))"
+      by (simp add: mult_ac scaleR_conv_of_real ring_distribs f_def g_def)
+    from sums_mult[OF this, of "inverse z"] z assms
+      show "(\<lambda>n. f n * (a*z)^n) sums (F z)" by (simp add: divide_simps mult_ac f_def F_def)
+  next
+    assume z: "z = 0"
+    have "(\<lambda>n. f n * (a * z) ^ n) sums f 0" using powser_sums_zero[of f] z by simp
+    with z show "(\<lambda>n. f n * (a * z) ^ n) sums (F z)" 
+      by (simp add: f_def F_def sin_coeff_def cos_coeff_def)
+    have "(\<lambda>n. g n * (a * z) ^ n) sums g 0" using powser_sums_zero[of g] z by simp
+    with z show "(\<lambda>n. g n * (a * z) ^ n) sums (G z)" 
+      by (simp add: g_def G_def sin_coeff_def cos_coeff_def)
+  qed
+  note sums = conjunct1[OF this] conjunct2[OF this]
+  
+  def h2 \<equiv> "\<lambda>z. (\<Sum>n. f n * (a*z)^n) / (\<Sum>n. g n * (a*z)^n) + 
+            Digamma (1 + z) - Digamma (1 - z)"
+  def POWSER \<equiv> "\<lambda>f z. (\<Sum>n. f n * (z^n :: complex))"
+  def POWSER' \<equiv> "\<lambda>f z. (\<Sum>n. diffs f n * (z^n :: complex))"
+
+  def h2' \<equiv> "\<lambda>z. a * (POWSER g (a*z) * POWSER' f (a*z) - POWSER f (a*z) * POWSER' g (a*z)) / 
+                     (POWSER g (a*z))^2 + Polygamma 1 (1 + z) + Polygamma 1 (1 - z)" 
+  
+  have h_eq: "h t = h2 t" if "abs (Re t) < 1" for t
+  proof -
+    from that have t: "t \<in> \<int> \<longleftrightarrow> t = 0" by (auto elim!: Ints_cases simp: dist_0_norm)
+    hence "h t = a*cot (a*t) - 1/t + Digamma (1 + t) - Digamma (1 - t)" 
+      unfolding h_def using Digamma_plus1[of t] by (force simp: field_simps a_def)
+    also have "a*cot (a*t) - 1/t = (F t) / (G t)"
+      using t by (auto simp add: divide_simps sin_eq_0 cot_def a_def F_def G_def)
+    also have "\<dots> = (\<Sum>n. f n * (a*t)^n) / (\<Sum>n. g n * (a*t)^n)"
+      using sums[of t] that by (simp add: sums_iff dist_0_norm)
+    finally show "h t = h2 t" by (simp only: h2_def)
+  qed
+  
+  let ?A = "{z. abs (Re z) < 1}"  
+  have "open ({z. Re z < 1} \<inter> {z. Re z > -1})"
+    using open_halfspace_Re_gt open_halfspace_Re_lt by auto
+  also have "({z. Re z < 1} \<inter> {z. Re z > -1}) = {z. abs (Re z) < 1}" by auto
+  finally have open_A: "open ?A" .
+  hence [simp]: "interior ?A = ?A" by (simp add: interior_open)
+  
+  have summable_f: "summable (\<lambda>n. f n * z^n)" for z
+    by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"])
+       (simp_all add: norm_mult a_def del: of_real_add)
+  have summable_g: "summable (\<lambda>n. g n * z^n)" for z
+    by (rule powser_inside, rule sums_summable, rule sums[of "\<i> * of_real (norm z + 1) / a"])
+       (simp_all add: norm_mult a_def del: of_real_add)
+  have summable_fg': "summable (\<lambda>n. diffs f n * z^n)" "summable (\<lambda>n. diffs g n * z^n)" for z
+    by (intro termdiff_converges_all summable_f summable_g)+
+  have "(POWSER f has_field_derivative (POWSER' f z)) (at z)"
+               "(POWSER g has_field_derivative (POWSER' g z)) (at z)" for z
+    unfolding POWSER_def POWSER'_def 
+    by (intro termdiffs_strong_converges_everywhere summable_f summable_g)+
+  note derivs = this[THEN DERIV_chain2[OF _ DERIV_cmult[OF DERIV_ident]], unfolded POWSER_def]
+  have "isCont (POWSER f) z" "isCont (POWSER g) z" "isCont (POWSER' f) z" "isCont (POWSER' g) z" 
+    for z unfolding POWSER_def POWSER'_def
+    by (intro isCont_powser_converges_everywhere summable_f summable_g summable_fg')+
+  note cont = this[THEN isCont_o2[rotated], unfolded POWSER_def POWSER'_def]
+
+  {
+    fix z :: complex assume z: "abs (Re z) < 1"
+    def d \<equiv> "\<i> * of_real (norm z + 1)"
+    have d: "abs (Re d) < 1" "norm z < norm d" by (simp_all add: d_def norm_mult del: of_real_add)
+    have "eventually (\<lambda>z. h z = h2 z) (nhds z)"
+      using eventually_nhds_in_nhd[of z ?A] using h_eq z 
+      by (auto elim!: eventually_mono simp: dist_0_norm)
+    
+    moreover from sums(2)[OF z] z have nz: "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0"
+      unfolding G_def by (auto simp: sums_iff sin_eq_0 a_def)
+    have A: "z \<in> \<int> \<longleftrightarrow> z = 0" using z by (auto elim!: Ints_cases)
+    have no_int: "1 + z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of "1+z" 1] A
+      by (auto elim!: nonpos_Ints_cases)
+    have no_int': "1 - z \<in> \<int> \<longleftrightarrow> z = 0" using z Ints_diff[of 1 "1-z"] A
+      by (auto elim!: nonpos_Ints_cases)
+    from no_int no_int' have no_int: "1 - z \<notin> \<int>\<^sub>\<le>\<^sub>0" "1 + z \<notin> \<int>\<^sub>\<le>\<^sub>0" by auto
+    have "(h2 has_field_derivative h2' z) (at z)" unfolding h2_def
+      by (rule DERIV_cong, (rule derivative_intros refl derivs[unfolded POWSER_def] nz no_int)+)
+         (auto simp: h2'_def POWSER_def field_simps power2_eq_square)
+    ultimately have deriv: "(h has_field_derivative h2' z) (at z)" 
+      by (subst DERIV_cong_ev[OF refl _ refl])
+    
+    from sums(2)[OF z] z have "(\<Sum>n. g n * (a * z) ^ n) \<noteq> 0"
+      unfolding G_def by (auto simp: sums_iff a_def sin_eq_0)
+    hence "isCont h2' z" using no_int unfolding h2'_def[abs_def] POWSER_def POWSER'_def
+      by (intro continuous_intros cont 
+            continuous_on_compose2[OF _ continuous_on_Polygamma[of "{z. Re z > 0}"]]) auto
+    note deriv and this
+  } note A = this
+
+  interpret h: periodic_fun_simple' h
+  proof
+    fix z :: complex
+    show "h (z + 1) = h z"
+    proof (cases "z \<in> \<int>")
+      assume z: "z \<notin> \<int>"
+      hence A: "z + 1 \<notin> \<int>" "z \<noteq> 0" using Ints_diff[of "z+1" 1] by auto
+      hence "Digamma (z + 1) - Digamma (-z) = Digamma z - Digamma (-z + 1)"
+        by (subst (1 2) Digamma_plus1) simp_all
+      with A z show "h (z + 1) = h z" 
+        by (simp add: h_def sin_plus_pi cos_plus_pi ring_distribs cot_def)
+    qed (simp add: h_def)
+  qed
+  
+  have h2'_eq: "h2' (z - 1) = h2' z" if z: "Re z > 0" "Re z < 1" for z
+  proof -
+    have "((\<lambda>z. h (z - 1)) has_field_derivative h2' (z - 1)) (at z)"
+      by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)])
+         (insert z, auto intro!: derivative_eq_intros)
+    hence "(h has_field_derivative h2' (z - 1)) (at z)" by (subst (asm) h.minus_1)
+    moreover from z have "(h has_field_derivative h2' z) (at z)" by (intro A) simp_all
+    ultimately show "h2' (z - 1) = h2' z" by (rule DERIV_unique)
+  qed
+
+  def h2'' \<equiv> "\<lambda>z. h2' (z - of_int \<lfloor>Re z\<rfloor>)"
+  have deriv: "(h has_field_derivative h2'' z) (at z)" for z
+  proof -
+    fix z :: complex
+    have B: "\<bar>Re z - real_of_int \<lfloor>Re z\<rfloor>\<bar> < 1" by linarith
+    have "((\<lambda>t. h (t - of_int \<lfloor>Re z\<rfloor>)) has_field_derivative h2'' z) (at z)"
+      unfolding h2''_def by (rule DERIV_cong, rule DERIV_chain'[OF _ A(1)])
+                            (insert B, auto intro!: derivative_intros)
+    thus "(h has_field_derivative h2'' z) (at z)" by (simp add: h.minus_of_int)
+  qed
+  
+  have cont: "continuous_on UNIV h2''"
+  proof (intro continuous_at_imp_continuous_on ballI)
+    fix z :: complex
+    def r \<equiv> "\<lfloor>Re z\<rfloor>" 
+    def A \<equiv> "{t. of_int r - 1 < Re t \<and> Re t < of_int r + 1}"
+    have "continuous_on A (\<lambda>t. h2' (t - of_int r))" unfolding A_def
+      by (intro continuous_at_imp_continuous_on isCont_o2[OF _ A(2)] ballI continuous_intros)
+         (simp_all add: abs_real_def)
+    moreover have "h2'' t = h2' (t - of_int r)" if t: "t \<in> A" for t
+    proof (cases "Re t \<ge> of_int r")
+      case True
+      from t have "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def) 
+      with True have "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor>" unfolding r_def by linarith
+      thus ?thesis by (auto simp: r_def h2''_def)
+    next
+      case False
+      from t have t: "of_int r - 1 < Re t" "Re t < of_int r + 1" by (simp_all add: A_def)
+      with False have t': "\<lfloor>Re t\<rfloor> = \<lfloor>Re z\<rfloor> - 1" unfolding r_def by linarith
+      moreover from t False have "h2' (t - of_int r + 1 - 1) = h2' (t - of_int r + 1)"
+        by (intro h2'_eq) simp_all
+      ultimately show ?thesis by (auto simp: r_def h2''_def algebra_simps t')
+    qed
+    ultimately have "continuous_on A h2''" by (subst continuous_on_cong[OF refl])
+    moreover {
+      have "open ({t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t})"
+        by (intro open_Int open_halfspace_Re_gt open_halfspace_Re_lt)
+      also have "{t. of_int r - 1 < Re t} \<inter> {t. of_int r + 1 > Re t} = A"
+        unfolding A_def by blast
+      finally have "open A" .
+    }
+    ultimately have C: "isCont h2'' t" if "t \<in> A" for t using that
+      by (subst (asm) continuous_on_eq_continuous_at) auto
+    have "of_int r - 1 < Re z" "Re z  < of_int r + 1" unfolding r_def by linarith+
+    thus "isCont h2'' z" by (intro C) (simp_all add: A_def)
+  qed
+  
+  from that[OF cont deriv] show ?thesis .
+qed
+
+lemma Gamma_reflection_complex: 
+  fixes z :: complex
+  shows "Gamma z * Gamma (1 - z) = of_real pi / sin (of_real pi * z)"
+proof -
+  let ?g = "\<lambda>z::complex. Gamma z * Gamma (1 - z) * sin (of_real pi * z)"
+  def g \<equiv> "\<lambda>z::complex. if z \<in> \<int> then of_real pi else ?g z"
+  let ?h = "\<lambda>z::complex. (of_real pi * cot (of_real pi*z) + Digamma z - Digamma (1 - z))"
+  def h \<equiv> "\<lambda>z::complex. if z \<in> \<int> then 0 else ?h z"
+
+  -- \<open>@{term g} is periodic with period 1.\<close>
+  interpret g: periodic_fun_simple' g
+  proof
+    fix z :: complex
+    show "g (z + 1) = g z"
+    proof (cases "z \<in> \<int>")
+      case False
+      hence "z * g z = z * Beta z (- z + 1) * sin (of_real pi * z)" by (simp add: g_def Beta_def)
+      also have "z * Beta z (- z + 1) = (z + 1 + -z) * Beta (z + 1) (- z + 1)" 
+        using False Ints_diff[of 1 "1 - z"] nonpos_Ints_subset_Ints
+        by (subst Beta_plus1_left [symmetric]) auto
+      also have "\<dots> * sin (of_real pi * z) = z * (Beta (z + 1) (-z) * sin (of_real pi * (z + 1)))"
+        using False Ints_diff[of "z+1" 1] Ints_minus[of "-z"] nonpos_Ints_subset_Ints
+        by (subst Beta_plus1_right) (auto simp: ring_distribs sin_plus_pi)
+      also from False have "Beta (z + 1) (-z) * sin (of_real pi * (z + 1)) = g (z + 1)" 
+        using Ints_diff[of "z+1" 1] by (auto simp: g_def Beta_def)
+      finally show "g (z + 1) = g z" using False by (subst (asm) mult_left_cancel) auto
+    qed (simp add: g_def)
+  qed
+
+  -- \<open>@{term g} is entire.\<close>
+  have g_g': "(g has_field_derivative (h z * g z)) (at z)" for z :: complex
+  proof (cases "z \<in> \<int>")
+    let ?h' = "\<lambda>z. Beta z (1 - z) * ((Digamma z - Digamma (1 - z)) * sin (z * of_real pi) +
+                     of_real pi * cos (z * of_real pi))"
+    case False
+    from False have "eventually (\<lambda>t. t \<in> UNIV - \<int>) (nhds z)"
+      by (intro eventually_nhds_in_open) (auto simp: open_Diff)
+    hence "eventually (\<lambda>t. g t = ?g t) (nhds z)" by eventually_elim (simp add: g_def)
+    moreover {
+      from False Ints_diff[of 1 "1-z"] have "1 - z \<notin> \<int>" by auto 
+      hence "(?g has_field_derivative ?h' z) (at z)" using nonpos_Ints_subset_Ints
+        by (auto intro!: derivative_eq_intros simp: algebra_simps Beta_def)
+      also from False have "sin (of_real pi * z) \<noteq> 0" by (subst sin_eq_0) auto
+      hence "?h' z = h z * g z"
+        using False unfolding g_def h_def cot_def by (simp add: field_simps Beta_def)
+      finally have "(?g has_field_derivative (h z * g z)) (at z)" .
+    }
+    ultimately show ?thesis by (subst DERIV_cong_ev[OF refl _ refl])
+  next
+    case True
+    then obtain n where z: "z = of_int n" by (auto elim!: Ints_cases)
+    let ?t = "(\<lambda>z::complex. if z = 0 then 1 else sin z / z) \<circ> (\<lambda>z. of_real pi * z)"
+    have deriv_0: "(g has_field_derivative 0) (at 0)"
+    proof (subst DERIV_cong_ev[OF refl _ refl])
+      show "eventually (\<lambda>z. g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) (nhds 0)"
+        using eventually_nhds_ball[OF zero_less_one, of "0::complex"]
+      proof eventually_elim
+        fix z :: complex assume z: "z \<in> ball 0 1"
+        show "g z = of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z"
+        proof (cases "z = 0")
+          assume z': "z \<noteq> 0"
+          with z have z'': "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z \<notin> \<int>" by (auto elim!: Ints_cases simp: dist_0_norm)
+          from Gamma_plus1[OF this(1)] have "Gamma z = Gamma (z + 1) / z" by simp
+          with z'' z' show ?thesis by (simp add: g_def ac_simps)
+        qed (simp add: g_def)
+      qed
+      have "(?t has_field_derivative (0 * of_real pi)) (at 0)"
+        using has_field_derivative_sin_z_over_z[of "UNIV :: complex set"] 
+        by (intro DERIV_chain) simp_all
+      thus "((\<lambda>z. of_real pi * Gamma (1 + z) * Gamma (1 - z) * ?t z) has_field_derivative 0) (at 0)"
+        by (auto intro!: derivative_eq_intros simp: o_def)
+    qed
+    
+    have "((g \<circ> (\<lambda>x. x - of_int n)) has_field_derivative 0 * 1) (at (of_int n))"
+      using deriv_0 by (intro DERIV_chain) (auto intro!: derivative_eq_intros)
+    also have "g \<circ> (\<lambda>x. x - of_int n) = g" by (intro ext) (simp add: g.minus_of_int)
+    finally show "(g has_field_derivative (h z * g z)) (at z)" by (simp add: z h_def)
+  qed
+
+  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
+  proof (cases "z \<in> \<int>")
+    case True
+    with that have "z = 0 \<or> z = 1" by (force elim!: Ints_cases)
+    moreover have "g 0 * g (1/2) = Gamma (1/2)^2 * g 0"
+      using fraction_not_in_ints[where 'a = complex, of 2 1] by (simp add: g_def power2_eq_square)
+    moreover have "g (1/2) * g 1 = Gamma (1/2)^2 * g 1"
+        using fraction_not_in_ints[where 'a = complex, of 2 1]
+        by (simp add: g_def power2_eq_square Beta_def algebra_simps)
+    ultimately show ?thesis by force
+  next
+    case False
+    hence z: "z/2 \<notin> \<int>" "(z+1)/2 \<notin> \<int>" using Ints_diff[of "z+1" 1] by (auto elim!: Ints_cases)
+    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)
+    from z have "1-z/2 \<notin> \<int>" "1-((z+1)/2) \<notin> \<int>" 
+      using Ints_diff[of 1 "1-z/2"] Ints_diff[of 1 "1-((z+1)/2)"] by auto
+    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)
+    from z have "g (z/2) * g ((z+1)/2) = 
+      (Gamma (z/2) * Gamma ((z+1)/2)) * (Gamma (1-z/2) * Gamma (1-((z+1)/2))) *
+      (sin (of_real pi * z/2) * sin (of_real pi * (z+1)/2))"
+      by (simp add: g_def)
+    also from z' Gamma_legendre_duplication_aux[of "z/2"] 
+      have "Gamma (z/2) * Gamma ((z+1)/2) = exp ((1-z) * of_real (ln 2)) * Gamma (1/2) * Gamma z"
+      by (simp add: add_divide_distrib)
+    also from z'' Gamma_legendre_duplication_aux[of "1-(z+1)/2"]
+      have "Gamma (1-z/2) * Gamma (1-(z+1)/2) = 
+              Gamma (1-z) * Gamma (1/2) * exp (z * of_real (ln 2))"
+      by (simp add: add_divide_distrib ac_simps)
+    finally have "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * (Gamma z * Gamma (1-z) *
+                    (2 * (sin (of_real pi*z/2) * sin (of_real pi*(z+1)/2))))"
+      by (simp add: add_ac power2_eq_square exp_add ring_distribs exp_diff exp_of_real)
+    also have "sin (of_real pi*(z+1)/2) = cos (of_real pi*z/2)"
+      using cos_sin_eq[of "- of_real pi * z/2", symmetric]
+      by (simp add: ring_distribs add_divide_distrib ac_simps)
+    also have "2 * (sin (of_real pi*z/2) * cos (of_real pi*z/2)) = sin (of_real pi * z)"
+      by (subst sin_times_cos) (simp add: field_simps)
+    also have "Gamma z * Gamma (1 - z) * sin (complex_of_real pi * z) = g z"
+      using \<open>z \<notin> \<int>\<close> by (simp add: g_def)
+    finally show ?thesis .
+  qed
+  have g_eq: "g (z/2) * g ((z+1)/2) = Gamma (1/2)^2 * g z" for z
+  proof -
+    def r \<equiv> "\<lfloor>Re z / 2\<rfloor>"
+    have "Gamma (1/2)^2 * g z = Gamma (1/2)^2 * g (z - of_int (2*r))" by (simp only: g.minus_of_int)
+    also have "of_int (2*r) = 2 * of_int r" by simp
+    also have "Re z - 2 * of_int r > -1" "Re z - 2 * of_int r < 2" unfolding r_def by linarith+
+    hence "Gamma (1/2)^2 * g (z - 2 * of_int r) = 
+                   g ((z - 2 * of_int r)/2) * g ((z - 2 * of_int r + 1)/2)" 
+      unfolding r_def by (intro g_eq[symmetric]) simp_all
+    also have "(z - 2 * of_int r) / 2 = z/2 - of_int r" by simp
+    also have "g \<dots> = g (z/2)" by (rule g.minus_of_int)
+    also have "(z - 2 * of_int r + 1) / 2 = (z + 1)/2 - of_int r" by simp
+    also have "g \<dots> = g ((z+1)/2)" by (rule g.minus_of_int)
+    finally show ?thesis ..
+  qed
+    
+  have g_nz [simp]: "g z \<noteq> 0" for z :: complex
+  unfolding g_def using Ints_diff[of 1 "1 - z"]
+    by (auto simp: Gamma_eq_zero_iff sin_eq_0 dest!: nonpos_Ints_Int)
+  
+  have h_eq: "h z = (h (z/2) + h ((z+1)/2)) / 2" for z
+  proof -
+    have "((\<lambda>t. g (t/2) * g ((t+1)/2)) has_field_derivative
+                       (g (z/2) * g ((z+1)/2)) * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
+      by (auto intro!: derivative_eq_intros g_g'[THEN DERIV_chain2] simp: field_simps)
+    hence "((\<lambda>t. Gamma (1/2)^2 * g t) has_field_derivative
+              Gamma (1/2)^2 * g z * ((h (z/2) + h ((z+1)/2)) / 2)) (at z)"
+      by (subst (1 2) g_eq[symmetric]) simp
+    from DERIV_cmult[OF this, of "inverse ((Gamma (1/2))^2)"]
+      have "(g has_field_derivative (g z * ((h (z/2) + h ((z+1)/2))/2))) (at z)"
+      using fraction_not_in_ints[where 'a = complex, of 2 1]
+      by (simp add: divide_simps Gamma_eq_zero_iff not_in_Ints_imp_not_in_nonpos_Ints)
+    moreover have "(g has_field_derivative (g z * h z)) (at z)"
+      using g_g'[of z] by (simp add: ac_simps)
+    ultimately have "g z * h z = g z * ((h (z/2) + h ((z+1)/2))/2)"
+      by (intro DERIV_unique)
+    thus "h z = (h (z/2) + h ((z+1)/2)) / 2" by simp
+  qed
+  
+  obtain h' where h'_cont: "continuous_on UNIV h'" and
+                  h_h': "\<And>z. (h has_field_derivative h' z) (at z)"
+     unfolding h_def by (erule Gamma_reflection_aux) 
+  
+  have h'_eq: "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" for z
+  proof -
+    have "((\<lambda>t. (h (t/2) + h ((t+1)/2)) / 2) has_field_derivative
+                       ((h' (z/2) + h' ((z+1)/2)) / 4)) (at z)" 
+      by (fastforce intro!: derivative_eq_intros h_h'[THEN DERIV_chain2])
+    hence "(h has_field_derivative ((h' (z/2) + h' ((z+1)/2))/4)) (at z)"
+      by (subst (asm) h_eq[symmetric])
+    from h_h' and this show "h' z = (h' (z/2) + h' ((z+1)/2)) / 4" by (rule DERIV_unique)
+  qed
+  
+  have h'_zero: "h' z = 0" for z
+  proof -
+    def m \<equiv> "max 1 \<bar>Re z\<bar>"
+    def B \<equiv> "{t. abs (Re t) \<le> m \<and> abs (Im t) \<le> abs (Im z)}"
+    have "closed ({t. Re t \<ge> -m} \<inter> {t. Re t \<le> m} \<inter> 
+                  {t. Im t \<ge> -\<bar>Im z\<bar>} \<inter> {t. Im t \<le> \<bar>Im z\<bar>})"
+      (is "closed ?B") by (intro closed_Int closed_halfspace_Re_ge closed_halfspace_Re_le 
+                                 closed_halfspace_Im_ge closed_halfspace_Im_le)
+    also have "?B = B" unfolding B_def by fastforce
+    finally have "closed B" .
+    moreover have "bounded B" unfolding bounded_iff
+    proof (intro ballI exI)
+      fix t assume t: "t \<in> B"
+      have "norm t \<le> \<bar>Re t\<bar> + \<bar>Im t\<bar>" by (rule cmod_le)
+      also from t have "\<bar>Re t\<bar> \<le> m" unfolding B_def by blast
+      also from t have "\<bar>Im t\<bar> \<le> \<bar>Im z\<bar>" unfolding B_def by blast
+      finally show "norm t \<le> m + \<bar>Im z\<bar>" by - simp
+    qed
+    ultimately have compact: "compact B" by (subst compact_eq_bounded_closed) blast
+    
+    def M \<equiv> "SUP z:B. norm (h' z)"
+    have "compact (h' ` B)"
+      by (intro compact_continuous_image continuous_on_subset[OF h'_cont] compact) blast+
+    hence bdd: "bdd_above ((\<lambda>z. norm (h' z)) ` B)"
+      using bdd_above_norm[of "h' ` B"] by (simp add: image_comp o_def compact_imp_bounded)
+    have "norm (h' z) \<le> M" unfolding M_def by (intro cSUP_upper bdd) (simp_all add: B_def m_def)
+    also have "M \<le> M/2"
+    proof (subst M_def, subst cSUP_le_iff)
+      have "z \<in> B" unfolding B_def m_def by simp
+      thus "B \<noteq> {}" by auto
+    next
+      show "\<forall>z\<in>B. norm (h' z) \<le> M/2"
+      proof
+        fix t :: complex assume t: "t \<in> B"
+        from h'_eq[of t] t have "h' t = (h' (t/2) + h' ((t+1)/2)) / 4" by (simp add: dist_0_norm)
+        also have "norm \<dots> = norm (h' (t/2) + h' ((t+1)/2)) / 4" by simp
+        also have "norm (h' (t/2) + h' ((t+1)/2)) \<le> norm (h' (t/2)) + norm (h' ((t+1)/2))"
+          by (rule norm_triangle_ineq)
+        also from t have "abs (Re ((t + 1)/2)) \<le> m" unfolding m_def B_def by auto
+        with t have "t/2 \<in> B" "(t+1)/2 \<in> B" unfolding B_def by auto        
+        hence "norm (h' (t/2)) + norm (h' ((t+1)/2)) \<le> M + M" unfolding M_def 
+          by (intro add_mono cSUP_upper bdd) (auto simp: B_def)
+        also have "(M + M) / 4 = M / 2" by simp
+        finally show "norm (h' t) \<le> M/2" by - simp_all
+      qed
+    qed (insert bdd, auto simp: cball_eq_empty)
+    hence "M \<le> 0" by simp
+    finally show "h' z = 0" by simp
+  qed
+  have h_h'_2: "(h has_field_derivative 0) (at z)" for z
+    using h_h'[of z] h'_zero[of z] by simp
+  
+  have g_real: "g z \<in> \<real>" if "z \<in> \<real>" for z
+    unfolding g_def using that by (auto intro!: Reals_mult Gamma_complex_real)
+  have h_real: "h z \<in> \<real>" if "z \<in> \<real>" for z
+    unfolding h_def using that by (auto intro!: Reals_mult Reals_add Reals_diff Polygamma_Real)
+  have g_nz: "g z \<noteq> 0" for z unfolding g_def using Ints_diff[of 1 "1-z"]
+    by (auto simp: Gamma_eq_zero_iff sin_eq_0)
+  
+  from h'_zero h_h'_2 have "\<exists>c. \<forall>z\<in>UNIV. h z = c"
+    by (intro has_field_derivative_zero_constant) (simp_all add: dist_0_norm)
+  then obtain c where c: "\<And>z. h z = c" by auto
+  have "\<exists>u. u \<in> closed_segment 0 1 \<and> Re (g 1) - Re (g 0) = Re (h u * g u * (1 - 0))"
+    by (intro complex_mvt_line g_g')
+    find_theorems name:deriv Reals
+  then guess u by (elim exE conjE) note u = this
+  from u(1) have u': "u \<in> \<real>" unfolding closed_segment_def 
+    by (auto simp: scaleR_conv_of_real)
+  from u' g_real[of u] g_nz[of u] have "Re (g u) \<noteq> 0" by (auto elim!: Reals_cases)
+  with u(2) c[of u] g_real[of u] g_nz[of u] u'
+    have "Re c = 0" by (simp add: complex_is_Real_iff g.of_1)
+  with h_real[of 0] c[of 0] have "c = 0" by (auto elim!: Reals_cases)
+  with c have A: "h z * g z = 0" for z by simp
+  hence "(g has_field_derivative 0) (at z)" for z using g_g'[of z] by simp
+  hence "\<exists>c'. \<forall>z\<in>UNIV. g z = c'" by (intro has_field_derivative_zero_constant) simp_all
+  then obtain c' where c: "\<And>z. g z = c'" by (force simp: dist_0_norm)
+  moreover from this[of 0] have "c' = pi" unfolding g_def by simp
+  ultimately have "g z = pi" by simp
+  
+  show ?thesis
+  proof (cases "z \<in> \<int>")
+    case False
+    with `g z = pi` show ?thesis by (auto simp: g_def divide_simps)
+  next
+    case True
+    then obtain n where n: "z = of_int n" by (elim Ints_cases)
+    with sin_eq_0[of "of_real pi * z"] have "sin (of_real pi * z) = 0" by force
+    moreover have "of_int (1 - n) \<in> \<int>\<^sub>\<le>\<^sub>0" if "n > 0" using that by (intro nonpos_Ints_of_int) simp
+    ultimately show ?thesis using n
+      by (cases "n \<le> 0") (auto simp: Gamma_eq_zero_iff nonpos_Ints_of_int)
+  qed
+qed
+
+lemma rGamma_reflection_complex: 
+  "rGamma z * rGamma (1 - z :: complex) = sin (of_real pi * z) / of_real pi"
+  using Gamma_reflection_complex[of z]
+    by (simp add: Gamma_def divide_simps split: split_if_asm)
+
+lemma rGamma_reflection_complex': 
+  "rGamma z * rGamma (- z :: complex) = -z * sin (of_real pi * z) / of_real pi"
+proof -
+  have "rGamma z * rGamma (-z) = -z * (rGamma z * rGamma (1 - z))"
+    using rGamma_plus1[of "-z", symmetric] by simp
+  also have "rGamma z * rGamma (1 - z) = sin (of_real pi * z) / of_real pi"
+    by (rule rGamma_reflection_complex)
+  finally show ?thesis by simp
+qed
+
+lemma Gamma_reflection_complex': 
+  "Gamma z * Gamma (- z :: complex) = - of_real pi / (z * sin (of_real pi * z))"
+  using rGamma_reflection_complex'[of z] by (force simp add: Gamma_def divide_simps mult_ac)
+
+
+
+lemma Gamma_one_half_real: "Gamma (1/2 :: real) = sqrt pi"
+proof -
+  from Gamma_reflection_complex[of "1/2"] fraction_not_in_ints[where 'a = complex, of 2 1]
+    have "Gamma (1/2 :: complex)^2 = of_real pi" by (simp add: power2_eq_square)
+  hence "of_real pi = Gamma (complex_of_real (1/2))^2" by simp
+  also have "\<dots> = of_real ((Gamma (1/2))^2)" by (subst Gamma_complex_of_real) simp_all
+  finally have "Gamma (1/2)^2 = pi" by (subst (asm) of_real_eq_iff) simp_all
+  moreover have "Gamma (1/2 :: real) \<ge> 0" using Gamma_real_pos[of "1/2"] by simp
+  ultimately show ?thesis by (rule real_sqrt_unique [symmetric])
+qed
+
+lemma Gamma_one_half_complex: "Gamma (1/2 :: complex) = of_real (sqrt pi)"
+proof -
+  have "Gamma (1/2 :: complex) = Gamma (of_real (1/2))" by simp
+  also have "\<dots> = of_real (sqrt pi)" by (simp only: Gamma_complex_of_real Gamma_one_half_real)
+  finally show ?thesis .
+qed
+
+lemma Gamma_legendre_duplication:
+  fixes z :: complex
+  assumes "z \<notin> \<int>\<^sub>\<le>\<^sub>0" "z + 1/2 \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "Gamma z * Gamma (z + 1/2) = 
+             exp ((1 - 2*z) * of_real (ln 2)) * of_real (sqrt pi) * Gamma (2*z)"
+  using Gamma_legendre_duplication_aux[OF assms] by (simp add: Gamma_one_half_complex)
+
+end
+
+
+subsection \<open>Limits and residues\<close>
+
+text \<open>
+  The inverse Gamma function has simple zeros:
+\<close>
+
+lemma rGamma_zeros: 
+  "(\<lambda>z. rGamma z / (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n * fact n :: 'a :: Gamma)"
+proof (subst tendsto_cong)
+  let ?f = "\<lambda>z. pochhammer z n * rGamma (z + of_nat (Suc n)) :: 'a"
+  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
+    show "eventually (\<lambda>z. rGamma z / (z + of_nat n) = ?f z) (at (- of_nat n))" 
+    by (subst pochhammer_rGamma[of _ "Suc n"])
+       (auto elim!: eventually_mono simp: divide_simps pochhammer_rec' eq_neg_iff_add_eq_0)
+  have "isCont ?f (- of_nat n)" by (intro continuous_intros)
+  thus "?f \<midarrow> (- of_nat n) \<rightarrow> (- 1) ^ n * fact n" unfolding isCont_def
+    by (simp add: pochhammer_same)
+qed
+
+
+text \<open>
+  The simple zeros of the inverse Gamma function correspond to simple poles of the Gamma function, 
+  and their residues can easily be computed from the limit we have just proven:
+\<close>
+
+lemma Gamma_poles: "filterlim Gamma at_infinity (at (- of_nat n :: 'a :: Gamma))"
+proof -
+  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
+    have "eventually (\<lambda>z. rGamma z \<noteq> (0 :: 'a)) (at (- of_nat n))"
+    by (auto elim!: eventually_mono nonpos_Ints_cases'
+             simp: rGamma_eq_zero_iff dist_of_nat dist_minus)
+  with isCont_rGamma[of "- of_nat n :: 'a", OF continuous_ident] 
+    have "filterlim (\<lambda>z. inverse (rGamma z) :: 'a) at_infinity (at (- of_nat n))"
+    unfolding isCont_def by (intro filterlim_compose[OF filterlim_inverse_at_infinity])
+                            (simp_all add: filterlim_at)
+  moreover have "(\<lambda>z. inverse (rGamma z) :: 'a) = Gamma" 
+    by (intro ext) (simp add: rGamma_inverse_Gamma)
+  ultimately show ?thesis by (simp only: )
+qed
+
+lemma Gamma_residues: 
+  "(\<lambda>z. Gamma z * (z + of_nat n)) \<midarrow> (- of_nat n) \<rightarrow> ((-1)^n / fact n :: 'a :: Gamma)"
+proof (subst tendsto_cong)
+  let ?c = "(- 1) ^ n / fact n :: 'a"
+  from eventually_at_ball'[OF zero_less_one, of "- of_nat n :: 'a" UNIV]
+    show "eventually (\<lambda>z. Gamma z * (z + of_nat n) = inverse (rGamma z / (z + of_nat n))) 
+            (at (- of_nat n))" 
+    by (auto elim!: eventually_mono simp: divide_simps rGamma_inverse_Gamma)
+  have "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> 
+          inverse ((- 1) ^ n * fact n :: 'a)"
+    by (intro tendsto_intros rGamma_zeros) simp_all
+  also have "inverse ((- 1) ^ n * fact n) = ?c" 
+    by (simp_all add: field_simps power_mult_distrib [symmetric] del: power_mult_distrib)
+  finally show "(\<lambda>z. inverse (rGamma z / (z + of_nat n))) \<midarrow> (- of_nat n) \<rightarrow> ?c" .
+qed
+
+
+
+subsection \<open>Alternative definitions\<close>
+
+
+subsubsection \<open>Variant of the Euler form\<close>
+
+
+definition Gamma_series_euler' where
+  "Gamma_series_euler' z n = 
+     inverse z * (\<Prod>k=1..n. exp (z * of_real (ln (1 + inverse (of_nat k)))) / (1 + z / of_nat k))"
+
+context
+begin
+private lemma Gamma_euler'_aux1:
+  fixes z :: "'a :: {real_normed_field,banach}"
+  assumes n: "n > 0"
+  shows "exp (z * of_real (ln (of_nat n + 1))) = (\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k))))"
+proof -
+  have "(\<Prod>k=1..n. exp (z * of_real (ln (1 + 1 / of_nat k)))) = 
+          exp (z * of_real (\<Sum>k = 1..n. ln (1 + 1 / real_of_nat k)))"
+    by (subst exp_setsum [symmetric]) (simp_all add: setsum_right_distrib)
+  also have "(\<Sum>k=1..n. ln (1 + 1 / of_nat k) :: real) = ln (\<Prod>k=1..n. 1 + 1 / real_of_nat k)"
+    by (subst ln_setprod [symmetric]) (auto intro!: add_pos_nonneg)
+  also have "(\<Prod>k=1..n. 1 + 1 / of_nat k :: real) = (\<Prod>k=1..n. (of_nat k + 1) / of_nat k)"
+    by (intro setprod.cong) (simp_all add: divide_simps)
+  also have "(\<Prod>k=1..n. (of_nat k + 1) / of_nat k :: real) = of_nat n + 1"
+    by (induction n) (simp_all add: setprod_nat_ivl_Suc' divide_simps)
+  finally show ?thesis ..
+qed
+
+lemma Gamma_series_euler':
+  assumes z: "(z :: 'a :: Gamma) \<notin> \<int>\<^sub>\<le>\<^sub>0"
+  shows "(\<lambda>n. Gamma_series_euler' z n) \<longlonglongrightarrow> Gamma z"
+proof (rule Gamma_seriesI, rule Lim_transform_eventually)
+  let ?f = "\<lambda>n. fact n * exp (z * of_real (ln (of_nat n + 1))) / pochhammer z (n + 1)"
+  let ?r = "\<lambda>n. ?f n / Gamma_series z n"
+  let ?r' = "\<lambda>n. exp (z * of_real (ln (of_nat (Suc n) / of_nat n)))"
+  from z have z': "z \<noteq> 0" by auto
+
+  have "eventually (\<lambda>n. ?r' n = ?r n) sequentially" using eventually_gt_at_top[of "0::nat"]
+    using z by (auto simp: divide_simps Gamma_series_def ring_distribs exp_diff ln_div add_ac
+                     elim!: eventually_mono dest: pochhammer_eq_0_imp_nonpos_Int)
+  moreover have "?r' \<longlonglongrightarrow> exp (z * of_real (ln 1))"
+    by (intro tendsto_intros LIMSEQ_Suc_n_over_n) simp_all
+  ultimately show "?r \<longlonglongrightarrow> 1" by (force dest!: Lim_transform_eventually)
+  
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. ?r n = Gamma_series_euler' z n / Gamma_series z n) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    from n z' have "Gamma_series_euler' z n = 
+      exp (z * of_real (ln (of_nat n + 1))) / (z * (\<Prod>k=1..n. (1 + z / of_nat k)))"
+      by (subst Gamma_euler'_aux1) 
+         (simp_all add: Gamma_series_euler'_def setprod.distrib 
+                        setprod_inversef[symmetric] divide_inverse)
+    also have "(\<Prod>k=1..n. (1 + z / of_nat k)) = pochhammer (z + 1) n / fact n"
+      by (cases n) (simp_all add: pochhammer_def fact_altdef setprod_shift_bounds_cl_Suc_ivl 
+                                  setprod_dividef[symmetric] divide_simps add_ac)
+    also have "z * \<dots> = pochhammer z (Suc n) / fact n" by (simp add: pochhammer_rec)
+    finally show "?r n = Gamma_series_euler' z n / Gamma_series z n" by simp
+  qed
+qed
+
+end
+
+
+
+subsubsection \<open>Weierstrass form\<close>
+
+definition Gamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
+  "Gamma_series_weierstrass z n = 
+     exp (-euler_mascheroni * z) / z * (\<Prod>k=1..n. exp (z / of_nat k) / (1 + z / of_nat k))"
+
+definition rGamma_series_weierstrass :: "'a :: {banach,real_normed_field} \<Rightarrow> nat \<Rightarrow> 'a" where
+  "rGamma_series_weierstrass z n = 
+     exp (euler_mascheroni * z) * z * (\<Prod>k=1..n. (1 + z / of_nat k) * exp (-z / of_nat k))"
+
+lemma Gamma_series_weierstrass_nonpos_Ints:
+  "eventually (\<lambda>k. Gamma_series_weierstrass (- of_nat n) k = 0) sequentially"
+  using eventually_ge_at_top[of n] by eventually_elim (auto simp: Gamma_series_weierstrass_def)
+
+lemma rGamma_series_weierstrass_nonpos_Ints:
+  "eventually (\<lambda>k. rGamma_series_weierstrass (- of_nat n) k = 0) sequentially"
+  using eventually_ge_at_top[of n] by eventually_elim (auto simp: rGamma_series_weierstrass_def)
+  
+lemma Gamma_weierstrass_complex: "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma (z :: complex)"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case True
+  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
+  also from True have "Gamma_series_weierstrass \<dots> \<longlonglongrightarrow> Gamma z"
+    by (simp add: tendsto_cong[OF Gamma_series_weierstrass_nonpos_Ints] Gamma_nonpos_Int)
+  finally show ?thesis .
+next
+  case False
+  hence z: "z \<noteq> 0" by auto
+  let ?f = "(\<lambda>x. \<Prod>x = Suc 0..x. exp (z / of_nat x) / (1 + z / of_nat x))"
+  have A: "exp (ln (1 + z / of_nat n)) = (1 + z / of_nat n)" if "n \<ge> 1" for n :: nat
+    using False that by (subst exp_Ln) (auto simp: field_simps dest!: plus_of_nat_eq_0_imp)
+  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"
+    using ln_Gamma_series'_aux[OF False]
+    by (simp only: atLeastLessThanSuc_atLeastAtMost [symmetric] One_nat_def 
+                   setsum_shift_bounds_Suc_ivl sums_def atLeast0LessThan)
+  from tendsto_exp[OF this] False z have "?f \<longlonglongrightarrow> z * exp (euler_mascheroni * z) * Gamma z"
+    by (simp add: exp_add exp_setsum exp_diff mult_ac Gamma_complex_altdef A)
+  from tendsto_mult[OF tendsto_const[of "exp (-euler_mascheroni * z) / z"] this] z
+    show "Gamma_series_weierstrass z \<longlonglongrightarrow> Gamma z" 
+    by (simp add: exp_minus divide_simps Gamma_series_weierstrass_def [abs_def])
+qed
+
+lemma Gamma_weierstrass_real: "Gamma_series_weierstrass x \<longlonglongrightarrow> Gamma (x :: real)"
+  using Gamma_weierstrass_complex[of "of_real x"] unfolding Gamma_series_weierstrass_def[abs_def]
+  by (subst tendsto_complex_of_real_iff [symmetric])
+     (simp_all add: exp_of_real[symmetric] Gamma_complex_of_real)
+
+lemma rGamma_weierstrass_complex: "rGamma_series_weierstrass z \<longlonglongrightarrow> rGamma (z :: complex)"
+proof (cases "z \<in> \<int>\<^sub>\<le>\<^sub>0")
+  case True
+  then obtain n where "z = - of_nat n" by (elim nonpos_Ints_cases')
+  also from True have "rGamma_series_weierstrass \<dots> \<longlonglongrightarrow> rGamma z"
+    by (simp add: tendsto_cong[OF rGamma_series_weierstrass_nonpos_Ints] rGamma_nonpos_Int)
+  finally show ?thesis .
+next
+  case False
+  have "rGamma_series_weierstrass z = (\<lambda>n. inverse (Gamma_series_weierstrass z n))"
+    by (simp add: rGamma_series_weierstrass_def[abs_def] Gamma_series_weierstrass_def 
+                  exp_minus divide_inverse setprod_inversef[symmetric] mult_ac)
+  also from False have "\<dots> \<longlonglongrightarrow> inverse (Gamma z)"
+    by (intro tendsto_intros Gamma_weierstrass_complex) (simp add: Gamma_eq_zero_iff)
+  finally show ?thesis by (simp add: Gamma_def)
+qed
+
+lemma sin_product_formula_complex:
+  fixes z :: complex
+  shows "(\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k^2)) \<longlonglongrightarrow> sin (of_real pi * z)"
+proof -
+  let ?f = "rGamma_series_weierstrass"
+  have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (- z) n))
+            \<longlonglongrightarrow> (- of_real pi * inverse z) * (rGamma z * rGamma (- z))"
+    by (intro tendsto_intros rGamma_weierstrass_complex)
+  also have "(\<lambda>n. (- of_real pi * inverse z) * (?f z n * ?f (-z) n)) =
+                    (\<lambda>n. of_real pi * z * (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2))"
+  proof
+    fix n :: nat
+    have "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = 
+              of_real pi * z * (\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2)"
+      by (simp add: rGamma_series_weierstrass_def mult_ac exp_minus
+                    divide_simps setprod.distrib[symmetric] power2_eq_square)
+    also have "(\<Prod>k=1..n. (of_nat k - z) * (of_nat k + z) / of_nat k ^ 2) =
+                 (\<Prod>k=1..n. 1 - z^2 / of_nat k ^ 2)"
+      by (intro setprod.cong) (simp_all add: power2_eq_square field_simps)
+    finally show "(- of_real pi * inverse z) * (?f z n * ?f (-z) n) = of_real pi * z * \<dots>" 
+      by (simp add: divide_simps)
+  qed
+  also have "(- of_real pi * inverse z) * (rGamma z * rGamma (- z)) = sin (of_real pi * z)"
+    by (subst rGamma_reflection_complex') (simp add: divide_simps)
+  finally show ?thesis .
+qed
+
+lemma sin_product_formula_real:
+  "(\<lambda>n. pi * (x::real) * (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x)"
+proof -
+  from sin_product_formula_complex[of "of_real x"]
+    have "(\<lambda>n. of_real pi * of_real x * (\<Prod>k=1..n. 1 - (of_real x)^2 / (of_nat k)^2))
+              \<longlonglongrightarrow> sin (of_real pi * of_real x :: complex)" (is "?f \<longlonglongrightarrow> ?y") .
+  also have "?f = (\<lambda>n. of_real (pi * x * (\<Prod>k=1..n. 1 - x^2 / (of_nat k^2))))" by simp
+  also have "?y = of_real (sin (pi * x))" by (simp only: sin_of_real [symmetric] of_real_mult)
+  finally show ?thesis by (subst (asm) tendsto_of_real_iff)
+qed
+
+lemma sin_product_formula_real':
+  assumes "x \<noteq> (0::real)"
+  shows   "(\<lambda>n. (\<Prod>k=1..n. 1 - x^2 / of_nat k^2)) \<longlonglongrightarrow> sin (pi * x) / (pi * x)" 
+  using tendsto_divide[OF sin_product_formula_real[of x] tendsto_const[of "pi * x"]] assms
+  by simp
+
+theorem inverse_squares_sums: "(\<lambda>n. 1 / (n + 1)\<^sup>2) sums (pi\<^sup>2 / 6)"
+proof -
+  def P \<equiv> "\<lambda>x n. (\<Prod>k=1..n. 1 - x^2 / of_nat k^2 :: real)"
+  def K \<equiv> "\<Sum>n. inverse (real_of_nat (Suc n))^2"
+  def f \<equiv> "\<lambda>x. \<Sum>n. P x n / of_nat (Suc n)^2"
+  def g \<equiv> "\<lambda>x. (1 - sin (pi * x) / (pi * x))"
+  
+  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
+  proof (cases "x = 0")
+    assume x: "x = 0"
+    have "summable (\<lambda>n. inverse ((real_of_nat (Suc n))\<^sup>2))"
+      using inverse_power_summable[of 2] by (subst summable_Suc_iff) simp
+    thus ?thesis by (simp add: x g_def P_def K_def inverse_eq_divide power_divide summable_sums)
+  next
+    assume x: "x \<noteq> 0"
+    have "(\<lambda>n. P x n - P x (Suc n)) sums (P x 0 - sin (pi * x) / (pi * x))"
+      unfolding P_def using x by (intro telescope_sums' sin_product_formula_real')
+    also have "(\<lambda>n. P x n - P x (Suc n)) = (\<lambda>n. (x^2 / of_nat (Suc n)^2) * P x n)"
+      unfolding P_def by (simp add: setprod_nat_ivl_Suc' algebra_simps)
+    also have "P x 0 = 1" by (simp add: P_def)
+    finally have "(\<lambda>n. x\<^sup>2 / (of_nat (Suc n))\<^sup>2 * P x n) sums (1 - sin (pi * x) / (pi * x))" .
+    from sums_divide[OF this, of "x^2"] x show ?thesis unfolding g_def by simp
+  qed
+  
+  have "continuous_on (ball 0 1) f"
+  proof (rule uniform_limit_theorem; (intro always_eventually allI)?)
+    show "uniform_limit (ball 0 1) (\<lambda>n x. \<Sum>k<n. P x k / of_nat (Suc k)^2) f sequentially"
+    proof (unfold f_def, rule weierstrass_m_test)
+      fix n :: nat and x :: real assume x: "x \<in> ball 0 1"
+      {
+        fix k :: nat assume k: "k \<ge> 1"
+        from x have "x^2 < 1" by (auto simp: dist_0_norm abs_square_less_1)
+        also from k have "\<dots> \<le> of_nat k^2" by simp
+        finally have "(1 - x^2 / of_nat k^2) \<in> {0..1}" using k 
+          by (simp_all add: field_simps del: of_nat_Suc)
+      }
+      hence "(\<Prod>k=1..n. abs (1 - x^2 / of_nat k^2)) \<le> (\<Prod>k=1..n. 1)" by (intro setprod_mono) simp
+      thus "norm (P x n / (of_nat (Suc n)^2)) \<le> 1 / of_nat (Suc n)^2"
+        unfolding P_def by (simp add: field_simps abs_setprod del: of_nat_Suc)
+    qed (subst summable_Suc_iff, insert inverse_power_summable[of 2], simp add: inverse_eq_divide)
+  qed (auto simp: P_def intro!: continuous_intros)
+  hence "isCont f 0" by (subst (asm) continuous_on_eq_continuous_at) simp_all
+  hence "(f \<midarrow> 0 \<rightarrow> f 0)" by (simp add: isCont_def)
+  also have "f 0 = K" unfolding f_def P_def K_def by (simp add: inverse_eq_divide power_divide)
+  finally have "f \<midarrow> 0 \<rightarrow> K" . 
+  
+  moreover have "f \<midarrow> 0 \<rightarrow> pi^2 / 6"
+  proof (rule Lim_transform_eventually)
+    def f' \<equiv> "\<lambda>x. \<Sum>n. - sin_coeff (n+3) * pi ^ (n+2) * x^n"
+    have "eventually (\<lambda>x. x \<noteq> (0::real)) (at 0)" 
+      by (auto simp add: eventually_at intro!: exI[of _ 1])
+    thus "eventually (\<lambda>x. f' x = f x) (at 0)"
+    proof eventually_elim
+      fix x :: real assume x: "x \<noteq> 0"
+      have "sin_coeff 1 = (1 :: real)" "sin_coeff 2 = (0::real)" by (simp_all add: sin_coeff_def)
+      with sums_split_initial_segment[OF sums_minus[OF sin_converges], of 3 "pi*x"]
+      have "(\<lambda>n. - (sin_coeff (n+3) * (pi*x)^(n+3))) sums (pi * x - sin (pi*x))"
+        by (simp add: eval_nat_numeral)
+      from sums_divide[OF this, of "x^3 * pi"] x
+        have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums ((1 - sin (pi*x) / (pi*x)) / x^2)"
+        by (simp add: divide_simps eval_nat_numeral power_mult_distrib mult_ac)
+      with x have "(\<lambda>n. - (sin_coeff (n+3) * pi^(n+2) * x^n)) sums (g x / x^2)"
+        by (simp add: g_def)
+      hence "f' x = g x / x^2" by (simp add: sums_iff f'_def)
+      also have "\<dots> = f x" using sums[of x] x by (simp add: sums_iff g_def f_def)
+      finally show "f' x = f x" .
+    qed
+    
+    have "isCont f' 0" unfolding f'_def
+    proof (intro isCont_powser_converges_everywhere) 
+      fix x :: real show "summable (\<lambda>n. -sin_coeff (n+3) * pi^(n+2) * x^n)"
+      proof (cases "x = 0")
+        assume x: "x \<noteq> 0"
+        from summable_divide[OF sums_summable[OF sums_split_initial_segment[OF 
+               sin_converges[of "pi*x"]], of 3], of "-pi*x^3"] x
+          show ?thesis by (simp add: mult_ac power_mult_distrib divide_simps eval_nat_numeral)
+      qed (simp only: summable_0_powser)
+    qed
+    hence "f' \<midarrow> 0 \<rightarrow> f' 0" by (simp add: isCont_def)
+    also have "f' 0 = pi * pi / fact 3" unfolding f'_def 
+      by (subst powser_zero) (simp add: sin_coeff_def)
+    finally show "f' \<midarrow> 0 \<rightarrow> pi^2 / 6" by (simp add: eval_nat_numeral)
+  qed
+  
+  ultimately have "K = pi^2 / 6" by (rule LIM_unique)
+  moreover from inverse_power_summable[of 2]
+    have "summable (\<lambda>n. (inverse (real_of_nat (Suc n)))\<^sup>2)"
+    by (subst summable_Suc_iff) (simp add: power_inverse)
+  ultimately show ?thesis unfolding K_def 
+    by (auto simp add: sums_iff power_divide inverse_eq_divide)
+qed
+
+
+
+subsection \<open>Binomial coefficient form\<close>
+
+lemma Gamma_binomial:
+  "(\<lambda>n. ((z + of_nat n) gchoose n) * exp (-z * of_real (ln (of_nat n)))) \<longlonglongrightarrow> rGamma (z+1)"
+proof (cases "z = 0")
+  case False
+  show ?thesis
+  proof (rule Lim_transform_eventually)
+    let ?powr = "\<lambda>a b. exp (b * of_real (ln (of_nat a)))"
+    show "eventually (\<lambda>n. rGamma_series z n / z = 
+            ((z + of_nat n) gchoose n) * ?powr n (-z)) sequentially"
+    proof (intro always_eventually allI)
+      fix n :: nat
+      from False have "((z + of_nat n) gchoose n) = pochhammer z (Suc n) / z / fact n"
+        by (simp add: gbinomial_pochhammer' pochhammer_rec)
+      also have "pochhammer z (Suc n) / z / fact n * ?powr n (-z) = rGamma_series z n / z"
+        by (simp add: rGamma_series_def divide_simps exp_minus)
+      finally show "rGamma_series z n / z = ((z + of_nat n) gchoose n) * ?powr n (-z)" ..
+    qed
+  
+    from False have "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma z / z" by (intro tendsto_intros)
+    also from False have "rGamma z / z = rGamma (z + 1)" using rGamma_plus1[of z] 
+      by (simp add: field_simps)
+    finally show "(\<lambda>n. rGamma_series z n / z) \<longlonglongrightarrow> rGamma (z+1)" .
+  qed
+qed (simp_all add: binomial_gbinomial [symmetric])
+
+lemma fact_binomial_limit: 
+  "(\<lambda>n. of_nat ((k + n) choose n) / of_nat (n ^ k) :: 'a :: Gamma) \<longlonglongrightarrow> 1 / fact k"
+proof (rule Lim_transform_eventually)
+  have "(\<lambda>n. of_nat ((k + n) choose n) / of_real (exp (of_nat k * ln (real_of_nat n))))
+            \<longlonglongrightarrow> 1 / Gamma (of_nat (Suc k) :: 'a)" (is "?f \<longlonglongrightarrow> _")
+    using Gamma_binomial[of "of_nat k :: 'a"] 
+    by (simp add: binomial_gbinomial add_ac Gamma_def divide_simps exp_of_real [symmetric] exp_minus)
+  also have "Gamma (of_nat (Suc k)) = fact k" by (rule Gamma_fact)
+  finally show "?f \<longlonglongrightarrow> 1 / fact k" .
+
+  show "eventually (\<lambda>n. ?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)) sequentially"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    from n have "exp (real_of_nat k * ln (real_of_nat n)) = real_of_nat (n^k)"
+      by (simp add: exp_of_nat_mult)
+    thus "?f n = of_nat ((k + n) choose n) / of_nat (n ^ k)" by simp
+  qed
+qed
+
+lemma binomial_asymptotic: 
+  "(\<lambda>n. of_nat ((k + n) choose n) / (of_nat (n ^ k) / fact k) :: 'a :: Gamma) \<longlonglongrightarrow> 1"
+  using tendsto_mult[OF fact_binomial_limit[of k] tendsto_const[of "fact k :: 'a"]] by simp
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -0,0 +1,255 @@
+(*
+  Title:    HOL/Multivariate_Analysis/Generalised_Binomial_Theorem.thy
+  Author:   Manuel Eberl, TU München
+  
+  The proof of the Generalised Binomial Theorem and related results.
+*)
+theory Generalised_Binomial_Theorem
+imports 
+  Complex_Main 
+  Complex_Transcendental
+  Summation
+begin
+
+subsection \<open>The generalised binomial theorem\<close>
+
+text \<open>
+  We prove the generalised binomial theorem for complex numbers, following the proof at:
+  https://proofwiki.org/wiki/Binomial_Theorem/General_Binomial_Theorem
+\<close>
+
+lemma gbinomial_ratio_limit:
+  fixes a :: "'a :: real_normed_field"
+  assumes "a \<notin> \<nat>"
+  shows "(\<lambda>n. (a gchoose n) / (a gchoose Suc n)) \<longlonglongrightarrow> -1"
+proof (rule Lim_transform_eventually)
+  let ?f = "\<lambda>n. inverse (a / of_nat (Suc n) - of_nat n / of_nat (Suc n))"
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. ?f n = (a gchoose n) /(a gchoose Suc n)) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    let ?P = "\<Prod>i = 0..n - 1. a - of_nat i"
+    from n have "(a gchoose n) / (a gchoose Suc n) = (of_nat (Suc n) :: 'a) *
+                   (?P / (\<Prod>i = 0..n. a - of_nat i))" by (simp add: gbinomial_def)
+    also from n have "(\<Prod>i = 0..n. a - of_nat i) = ?P * (a - of_nat n)"
+      by (cases n) (simp_all add: setprod_nat_ivl_Suc)
+    also have "?P / \<dots> = (?P / ?P) / (a - of_nat n)" by (rule divide_divide_eq_left[symmetric])
+    also from assms have "?P / ?P = 1" by auto
+    also have "of_nat (Suc n) * (1 / (a - of_nat n)) = 
+                   inverse (inverse (of_nat (Suc n)) * (a - of_nat n))" by (simp add: field_simps)
+    also have "inverse (of_nat (Suc n)) * (a - of_nat n) = a / of_nat (Suc n) - of_nat n / of_nat (Suc n)"
+      by (simp add: field_simps del: of_nat_Suc)
+    finally show "?f n = (a gchoose n) / (a gchoose Suc n)" by simp
+  qed
+
+  have "(\<lambda>n. norm a / (of_nat (Suc n))) \<longlonglongrightarrow> 0" 
+    unfolding divide_inverse
+    by (intro tendsto_mult_right_zero LIMSEQ_inverse_real_of_nat)
+  hence "(\<lambda>n. a / of_nat (Suc n)) \<longlonglongrightarrow> 0"
+    by (subst tendsto_norm_zero_iff[symmetric]) (simp add: norm_divide del: of_nat_Suc)
+  hence "?f \<longlonglongrightarrow> inverse (0 - 1)"
+    by (intro tendsto_inverse tendsto_diff LIMSEQ_n_over_Suc_n) simp_all
+  thus "?f \<longlonglongrightarrow> -1" by simp
+qed
+
+lemma conv_radius_gchoose:
+  fixes a :: "'a :: {real_normed_field,banach}"
+  shows "conv_radius (\<lambda>n. a gchoose n) = (if a \<in> \<nat> then \<infinity> else 1)"
+proof (cases "a \<in> \<nat>")
+  assume a: "a \<in> \<nat>"
+  have "eventually (\<lambda>n. (a gchoose n) = 0) sequentially"
+    using eventually_gt_at_top[of "nat \<lfloor>norm a\<rfloor>"]
+    by eventually_elim (insert a, auto elim!: Nats_cases simp: binomial_gbinomial[symmetric])
+  from conv_radius_cong[OF this] a show ?thesis by simp
+next
+  assume a: "a \<notin> \<nat>"
+  from tendsto_norm[OF gbinomial_ratio_limit[OF this]]
+    have "conv_radius (\<lambda>n. a gchoose n) = 1"
+    by (intro conv_radius_ratio_limit_nonzero[of _ 1]) (simp_all add: norm_divide)
+  with a show ?thesis by simp
+qed
+
+lemma gen_binomial_complex:
+  fixes z :: complex
+  assumes "norm z < 1"
+  shows   "(\<lambda>n. (a gchoose n) * z^n) sums (1 + z) powr a"
+proof -
+  def K \<equiv> "1 - (1 - norm z) / 2"
+  from assms have K: "K > 0" "K < 1" "norm z < K"
+     unfolding K_def by (auto simp: field_simps intro!: add_pos_nonneg)
+  let ?f = "\<lambda>n. a gchoose n" and ?f' = "diffs (\<lambda>n. a gchoose n)"
+  have summable_strong: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < 1" for z using that
+    by (intro summable_in_conv_radius) (simp_all add: conv_radius_gchoose)
+  with K have summable: "summable (\<lambda>n. ?f n * z ^ n)" if "norm z < K" for z using that by auto
+  hence summable': "summable (\<lambda>n. ?f' n * z ^ n)" if "norm z < K" for z using that
+    by (intro termdiff_converges[of _ K]) simp_all
+  
+  def f \<equiv> "\<lambda>z. \<Sum>n. ?f n * z ^ n" and f' \<equiv> "\<lambda>z. \<Sum>n. ?f' n * z ^ n"
+  {
+    fix z :: complex assume z: "norm z < K"
+    from summable_mult2[OF summable'[OF z], of z]
+      have summable1: "summable (\<lambda>n. ?f' n * z ^ Suc n)" by (simp add: mult_ac)
+    hence summable2: "summable (\<lambda>n. of_nat n * ?f n * z^n)" 
+      unfolding diffs_def by (subst (asm) summable_Suc_iff)
+
+    have "(1 + z) * f' z = (\<Sum>n. ?f' n * z^n) + (\<Sum>n. ?f' n * z^Suc n)"
+      unfolding f'_def using summable' z by (simp add: algebra_simps suminf_mult)
+    also have "(\<Sum>n. ?f' n * z^n) = (\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n)"
+      by (intro suminf_cong) (simp add: diffs_def)
+    also have "(\<Sum>n. ?f' n * z^Suc n) = (\<Sum>n. of_nat n * ?f n * z ^ n)" 
+      using summable1 suminf_split_initial_segment[OF summable1] unfolding diffs_def
+      by (subst suminf_split_head, subst (asm) summable_Suc_iff) simp_all
+    also have "(\<Sum>n. of_nat (Suc n) * ?f (Suc n) * z^n) + (\<Sum>n. of_nat n * ?f n * z^n) =
+                 (\<Sum>n. a * ?f n * z^n)"
+      by (subst gbinomial_mult_1, subst suminf_add)
+         (insert summable'[OF z] summable2, 
+          simp_all add: summable_powser_split_head algebra_simps diffs_def)
+    also have "\<dots> = a * f z" unfolding f_def
+      by (subst suminf_mult[symmetric]) (simp_all add: summable[OF z] mult_ac)
+    finally have "a * f z = (1 + z) * f' z" by simp
+  } note deriv = this
+
+  have [derivative_intros]: "(f has_field_derivative f' z) (at z)" if "norm z < of_real K" for z
+    unfolding f_def f'_def using K that
+    by (intro termdiffs_strong[of "?f" K z] summable_strong) simp_all
+  have "f 0 = (\<Sum>n. if n = 0 then 1 else 0)" unfolding f_def by (intro suminf_cong) simp
+  also have "\<dots> = 1" using sums_single[of 0 "\<lambda>_. 1::complex"] unfolding sums_iff by simp
+  finally have [simp]: "f 0 = 1" .
+
+  have "\<exists>c. \<forall>z\<in>ball 0 K. f z * (1 + z) powr (-a) = c"
+  proof (rule has_field_derivative_zero_constant)
+    fix z :: complex assume z': "z \<in> ball 0 K"
+    hence z: "norm z < K" by (simp add: dist_0_norm)
+    with K have nz: "1 + z \<noteq> 0" by (auto dest!: minus_unique)
+    from z K have "norm z < 1" by simp
+    hence "Im (1 + z) \<noteq> 0 \<or> Re (1 + z) > 0" by (cases z) auto
+    hence "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 
+              f' z * (1 + z) powr (-a) - a * f z * (1 + z) powr (-a-1)) (at z)" using z
+      by (auto intro!: derivative_eq_intros)
+    also from z have "a * f z = (1 + z) * f' z" by (rule deriv)
+    finally show "((\<lambda>z. f z * (1 + z) powr (-a)) has_field_derivative 0) (at z within ball 0 K)" 
+      using nz by (simp add: field_simps powr_diff_complex at_within_open[OF z'])
+  qed simp_all
+  then obtain c where c: "\<And>z. z \<in> ball 0 K \<Longrightarrow> f z * (1 + z) powr (-a) = c" by blast
+  from c[of 0] and K have "c = 1" by simp
+  with c[of z] have "f z = (1 + z) powr a" using K 
+    by (simp add: powr_minus_complex field_simps dist_complex_def)
+  with summable K show ?thesis unfolding f_def by (simp add: sums_iff)
+qed
+
+lemma gen_binomial_complex':
+  fixes x y :: real and a :: complex
+  assumes "\<bar>x\<bar> < \<bar>y\<bar>"
+  shows   "(\<lambda>n. (a gchoose n) * of_real x^n * of_real y powr (a - of_nat n)) sums 
+               of_real (x + y) powr a" (is "?P x y")
+proof -
+  {
+    fix x y :: real assume xy: "\<bar>x\<bar> < \<bar>y\<bar>" "y \<ge> 0"
+    hence "y > 0" by simp
+    note xy = xy this
+    from xy have "(\<lambda>n. (a gchoose n) * of_real (x / y) ^ n) sums (1 + of_real (x / y)) powr a"
+        by (intro gen_binomial_complex) (simp add: norm_divide)
+    hence "(\<lambda>n. (a gchoose n) * of_real (x / y) ^ n * y powr a) sums 
+               ((1 + of_real (x / y)) powr a * y powr a)"
+      by (rule sums_mult2)
+    also have "(1 + complex_of_real (x / y)) = complex_of_real (1 + x/y)" by simp
+    also from xy have "\<dots> powr a * of_real y powr a = (\<dots> * y) powr a"
+      by (subst powr_times_real[symmetric]) (simp_all add: field_simps)
+    also from xy have "complex_of_real (1 + x / y) * complex_of_real y = of_real (x + y)"
+      by (simp add: field_simps)
+    finally have "?P x y" using xy by (simp add: field_simps powr_diff_complex powr_nat)
+  } note A = this
+
+  show ?thesis
+  proof (cases "y < 0")
+    assume y: "y < 0"
+    with assms have xy: "x + y < 0" by simp
+    with assms have "\<bar>-x\<bar> < \<bar>-y\<bar>" "-y \<ge> 0" by simp_all
+    note A[OF this]
+    also have "complex_of_real (-x + -y) = - complex_of_real (x + y)" by simp
+    also from xy assms have "... powr a = (-1) powr -a * of_real (x + y) powr a"
+      by (subst powr_neg_real_complex) (simp add: abs_real_def split: split_if_asm)
+    also {
+      fix n :: nat
+      from y have "(a gchoose n) * of_real (-x) ^ n * of_real (-y) powr (a - of_nat n) = 
+                       (a gchoose n) * (-of_real x / -of_real y) ^ n * (- of_real y) powr a"
+        by (subst power_divide) (simp add: powr_diff_complex powr_nat)
+      also from y have "(- of_real y) powr a = (-1) powr -a * of_real y powr a"
+        by (subst powr_neg_real_complex) simp
+      also have "-complex_of_real x / -complex_of_real y = complex_of_real x / complex_of_real y"
+        by simp
+      also have "... ^ n = of_real x ^ n / of_real y ^ n" by (simp add: power_divide)
+      also have "(a gchoose n) * ... * ((-1) powr -a * of_real y powr a) = 
+                   (-1) powr -a * ((a gchoose n) * of_real x ^ n * of_real y powr (a - n))"
+        by (simp add: algebra_simps powr_diff_complex powr_nat)
+      finally have "(a gchoose n) * of_real (- x) ^ n * of_real (- y) powr (a - of_nat n) =
+                      (-1) powr -a * ((a gchoose n) * of_real x ^ n * of_real y powr (a - of_nat n))" .
+    }
+    note sums_cong[OF this]
+    finally show ?thesis by (simp add: sums_mult_iff)
+  qed (insert A[of x y] assms, simp_all add: not_less)
+qed
+
+lemma gen_binomial_complex'':
+  fixes x y :: real and a :: complex
+  assumes "\<bar>y\<bar> < \<bar>x\<bar>"
+  shows   "(\<lambda>n. (a gchoose n) * of_real x powr (a - of_nat n) * of_real y ^ n) sums 
+               of_real (x + y) powr a"
+  using gen_binomial_complex'[OF assms] by (simp add: mult_ac add.commute)
+
+lemma gen_binomial_real:
+  fixes z :: real
+  assumes "\<bar>z\<bar> < 1"
+  shows   "(\<lambda>n. (a gchoose n) * z^n) sums (1 + z) powr a"
+proof -
+  from assms have "norm (of_real z :: complex) < 1" by simp
+  from gen_binomial_complex[OF this]
+    have "(\<lambda>n. (of_real a gchoose n :: complex) * of_real z ^ n) sums
+              (of_real (1 + z)) powr (of_real a)" by simp
+  also have "(of_real (1 + z) :: complex) powr (of_real a) = of_real ((1 + z) powr a)"
+    using assms by (subst powr_of_real) simp_all
+  also have "(of_real a gchoose n :: complex) = of_real (a gchoose n)" for n 
+    by (simp add: gbinomial_def)
+  hence "(\<lambda>n. (of_real a gchoose n :: complex) * of_real z ^ n) =
+           (\<lambda>n. of_real ((a gchoose n) * z ^ n))" by (intro ext) simp
+  finally show ?thesis by (simp only: sums_of_real_iff)
+qed 
+
+lemma gen_binomial_real':
+  fixes x y a :: real
+  assumes "\<bar>x\<bar> < y"
+  shows   "(\<lambda>n. (a gchoose n) * x^n * y powr (a - of_nat n)) sums (x + y) powr a"
+proof -
+  from assms have "y > 0" by simp
+  note xy = this assms
+  from assms have "\<bar>x / y\<bar> < 1" by simp
+  hence "(\<lambda>n. (a gchoose n) * (x / y) ^ n) sums (1 + x / y) powr a"
+    by (rule gen_binomial_real)
+  hence "(\<lambda>n. (a gchoose n) * (x / y) ^ n * y powr a) sums ((1 + x / y) powr a * y powr a)"
+    by (rule sums_mult2)
+  with xy show ?thesis 
+    by (simp add: field_simps powr_divide powr_divide2[symmetric] powr_realpow)
+qed
+
+lemma one_plus_neg_powr_powser:
+  fixes z s :: complex
+  assumes "norm (z :: complex) < 1"
+  shows "(\<lambda>n. (-1)^n * ((s + n - 1) gchoose n) * z^n) sums (1 + z) powr (-s)"
+    using gen_binomial_complex[OF assms, of "-s"] by (simp add: gbinomial_minus)
+
+lemma gen_binomial_real'':
+  fixes x y a :: real
+  assumes "\<bar>y\<bar> < x"
+  shows   "(\<lambda>n. (a gchoose n) * x powr (a - of_nat n) * y^n) sums (x + y) powr a"
+  using gen_binomial_real'[OF assms] by (simp add: mult_ac add.commute)
+
+lemma sqrt_series':
+  "\<bar>z\<bar> < a \<Longrightarrow> (\<lambda>n. ((1/2) gchoose n) * a powr (1/2 - real_of_nat n) * z ^ n) sums 
+                  sqrt (a + z :: real)"
+  using gen_binomial_real''[of z a "1/2"] by (simp add: powr_half_sqrt)
+
+lemma sqrt_series:
+  "\<bar>z\<bar> < 1 \<Longrightarrow> (\<lambda>n. ((1/2) gchoose n) * z ^ n) sums sqrt (1 + z)"
+  using gen_binomial_real[of z "1/2"] by (simp add: powr_half_sqrt)
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Harmonic_Numbers.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -0,0 +1,573 @@
+(*
+  Title:    HOL/Multivariate_Analysis/Harmonic_Numbers.thy
+  Author:   Manuel Eberl, TU München
+  
+  The definition of the Harmonic Numbers and the Euler–Mascheroni constant.
+  Also provides a reasonably accurate approximation of @{term "ln 2 :: real"} 
+  and the Euler–Mascheroni constant.
+*)
+theory Harmonic_Numbers
+imports 
+  Complex_Transcendental
+  Summation
+  Integral_Test
+begin
+
+lemma ln_2_less_1: "ln 2 < (1::real)"
+proof -
+  have "2 < 5/(2::real)" by simp
+  also have "5/2 \<le> exp (1::real)" using exp_lower_taylor_quadratic[of 1, simplified] by simp
+  finally have "exp (ln 2) < exp (1::real)" by simp
+  thus "ln 2 < (1::real)" by (subst (asm) exp_less_cancel_iff) simp
+qed
+
+lemma setsum_Suc_diff':
+  fixes f :: "nat \<Rightarrow> 'a::ab_group_add"
+  assumes "m \<le> n"
+  shows "(\<Sum>i = m..<n. f(Suc i) - f i) = f n - f m"
+using assms by (induct n) (auto simp: le_Suc_eq)
+
+lemma eval_fact:
+  "fact 0 = 1"
+  "fact (Suc 0) = 1"
+  "fact (numeral n) = numeral n * fact (pred_numeral n)"
+  by (simp, simp, simp_all only: numeral_eq_Suc fact_Suc,
+      simp only: numeral_eq_Suc [symmetric] of_nat_numeral)
+
+lemma setsum_poly_horner_expand:
+  "(\<Sum>k<(numeral n::nat). f k * x^k) = f 0 + (\<Sum>k<pred_numeral n. f (k+1) * x^k) * x"
+  "(\<Sum>k<Suc 0. f k * x^k) = (f 0 :: 'a :: semiring_1)"
+  "(\<Sum>k<(0::nat). f k * x^k) = 0"
+proof -
+  {
+    fix m :: nat
+    have "(\<Sum>k<Suc m. f k * x^k) = f 0 + (\<Sum>k=Suc 0..<Suc m. f k * x^k)"
+      by (subst atLeast0LessThan [symmetric], subst setsum_head_upt_Suc) simp_all
+    also have "(\<Sum>k=Suc 0..<Suc m. f k * x^k) = (\<Sum>k<m. f (k+1) * x^k) * x"
+      by (subst setsum_shift_bounds_Suc_ivl)
+         (simp add: setsum_left_distrib algebra_simps atLeast0LessThan power_commutes)
+    finally have "(\<Sum>k<Suc m. f k * x ^ k) = f 0 + (\<Sum>k<m. f (k + 1) * x ^ k) * x" .
+  }
+  from this[of "pred_numeral n"] 
+    show "(\<Sum>k<numeral n. f k * x^k) = f 0 + (\<Sum>k<pred_numeral n. f (k+1) * x^k) * x" 
+    by (simp add: numeral_eq_Suc)
+qed simp_all
+
+
+subsection \<open>The Harmonic numbers\<close>
+
+definition harm :: "nat \<Rightarrow> 'a :: real_normed_field" where
+  "harm n = (\<Sum>k=1..n. inverse (of_nat k))"
+
+lemma harm_altdef: "harm n = (\<Sum>k<n. inverse (of_nat (Suc k)))"
+  unfolding harm_def by (induction n) simp_all
+
+lemma harm_Suc: "harm (Suc n) = harm n + inverse (of_nat (Suc n))"
+  by (simp add: harm_def)
+
+lemma harm_nonneg: "harm n \<ge> (0 :: 'a :: {real_normed_field,linordered_field})"
+  unfolding harm_def by (intro setsum_nonneg) simp_all
+
+lemma of_real_harm: "of_real (harm n) = harm n"
+  unfolding harm_def by simp
+  
+lemma norm_harm: "norm (harm n) = harm n"
+  by (subst of_real_harm [symmetric]) (simp add: harm_nonneg)
+
+lemma harm_expand: 
+  "harm (Suc 0) = 1"
+  "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)"
+proof -
+  have "numeral n = Suc (pred_numeral n)" by simp
+  also have "harm \<dots> = harm (pred_numeral n) + inverse (numeral n)"
+    by (subst harm_Suc, subst numeral_eq_Suc[symmetric]) simp
+  finally show "harm (numeral n) = harm (pred_numeral n) + inverse (numeral n)" .
+qed (simp add: harm_def)
+
+lemma not_convergent_harm: "\<not>convergent (harm :: nat \<Rightarrow> 'a :: real_normed_field)"
+proof -
+  have "convergent (\<lambda>n. norm (harm n :: 'a)) \<longleftrightarrow>
+            convergent (harm :: nat \<Rightarrow> real)" by (simp add: norm_harm)
+  also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k=Suc 0..Suc n. inverse (of_nat k) :: real)"
+    unfolding harm_def[abs_def] by (subst convergent_Suc_iff) simp_all
+  also have "... \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k\<le>n. inverse (of_nat (Suc k)) :: real)"
+    by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: atLeast0AtMost)
+  also have "... \<longleftrightarrow> summable (\<lambda>n. inverse (of_nat n) :: real)"
+    by (subst summable_Suc_iff [symmetric]) (simp add: summable_iff_convergent')
+  also have "\<not>..." by (rule not_summable_harmonic)
+  finally show ?thesis by (blast dest: convergent_norm)
+qed
+
+
+subsection \<open>The Euler–Mascheroni constant\<close>
+
+text \<open>
+  The limit of the difference between the partial harmonic sum and the natural logarithm
+  (approximately 0.577216). This value occurs e.g. in the definition of the Gamma function.
+ \<close>
+definition euler_mascheroni :: "'a :: real_normed_algebra_1" where
+  "euler_mascheroni = of_real (lim (\<lambda>n. harm n - ln (of_nat n)))"
+
+lemma of_real_euler_mascheroni [simp]: "of_real euler_mascheroni = euler_mascheroni"
+  by (simp add: euler_mascheroni_def)
+
+interpretation euler_mascheroni: antimono_fun_sum_integral_diff "\<lambda>x. inverse (x + 1)"
+  by unfold_locales (auto intro!: continuous_intros)
+
+lemma euler_mascheroni_sum_integral_diff_series:
+  "euler_mascheroni.sum_integral_diff_series n = harm (Suc n) - ln (of_nat (Suc n))"
+proof -
+  have "harm (Suc n) = (\<Sum>k=0..n. inverse (of_nat k + 1) :: real)" unfolding harm_def
+    unfolding One_nat_def by (subst setsum_shift_bounds_cl_Suc_ivl) (simp add: add_ac)
+  moreover have "((\<lambda>x. inverse (x + 1) :: real) has_integral ln (of_nat n + 1) - ln (0 + 1))
+                   {0..of_nat n}"
+    by (intro fundamental_theorem_of_calculus)
+       (auto intro!: derivative_eq_intros simp: divide_inverse
+           has_field_derivative_iff_has_vector_derivative[symmetric])
+  hence "integral {0..of_nat n} (\<lambda>x. inverse (x + 1) :: real) = ln (of_nat (Suc n))"
+    by (auto dest!: integral_unique)
+  ultimately show ?thesis 
+    by (simp add: euler_mascheroni.sum_integral_diff_series_def atLeast0AtMost)
+qed
+
+lemma euler_mascheroni_sequence_decreasing:
+  "m > 0 \<Longrightarrow> m \<le> n \<Longrightarrow> harm n - ln (of_nat n) \<le> harm m - ln (of_nat m :: real)"
+  by (cases m, simp, cases n, simp, hypsubst,
+      subst (1 2) euler_mascheroni_sum_integral_diff_series [symmetric],
+      rule euler_mascheroni.sum_integral_diff_series_antimono, simp)
+
+lemma euler_mascheroni_sequence_nonneg:
+  "n > 0 \<Longrightarrow> harm n - ln (of_nat n) \<ge> (0::real)"
+  by (cases n, simp, hypsubst, subst euler_mascheroni_sum_integral_diff_series [symmetric],
+      rule euler_mascheroni.sum_integral_diff_series_nonneg)
+
+lemma euler_mascheroni_convergent: "convergent (\<lambda>n. harm n - ln (of_nat n) :: real)"
+proof -
+  have A: "(\<lambda>n. harm (Suc n) - ln (of_nat (Suc n))) = 
+             euler_mascheroni.sum_integral_diff_series"
+    by (subst euler_mascheroni_sum_integral_diff_series [symmetric]) (rule refl)
+  have "convergent (\<lambda>n. harm (Suc n) - ln (of_nat (Suc n) :: real))"
+    by (subst A) (fact euler_mascheroni.sum_integral_diff_series_convergent)
+  thus ?thesis by (subst (asm) convergent_Suc_iff)
+qed
+
+lemma euler_mascheroni_LIMSEQ: 
+  "(\<lambda>n. harm n - ln (of_nat n) :: real) \<longlonglongrightarrow> euler_mascheroni"
+  unfolding euler_mascheroni_def
+  by (simp add: convergent_LIMSEQ_iff [symmetric] euler_mascheroni_convergent)
+
+lemma euler_mascheroni_LIMSEQ_of_real: 
+  "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow> 
+      (euler_mascheroni :: 'a :: {real_normed_algebra_1, topological_space})"
+proof -
+  have "(\<lambda>n. of_real (harm n - ln (of_nat n))) \<longlonglongrightarrow> (of_real (euler_mascheroni) :: 'a)"
+    by (intro tendsto_of_real euler_mascheroni_LIMSEQ)
+  thus ?thesis by simp
+qed
+
+lemma euler_mascheroni_sum:
+  "(\<lambda>n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real)
+       sums euler_mascheroni" 
+ using sums_add[OF telescope_sums[OF LIMSEQ_Suc[OF euler_mascheroni_LIMSEQ]]
+                   telescope_sums'[OF LIMSEQ_inverse_real_of_nat]]
+  by (simp_all add: harm_def algebra_simps)
+
+lemma alternating_harmonic_series_sums: "(\<lambda>k. (-1)^k / real_of_nat (Suc k)) sums ln 2"
+proof -
+  let ?f = "\<lambda>n. harm n - ln (real_of_nat n)"
+  let ?g = "\<lambda>n. if even n then 0 else (2::real)"
+  let ?em = "\<lambda>n. harm n - ln (real_of_nat n)"
+  have "eventually (\<lambda>n. ?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))) at_top"
+    using eventually_gt_at_top[of "0::nat"]
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0"
+    have "(\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k)) =
+              (\<Sum>k<2*n. ((-1)^k + ?g k) / of_nat (Suc k)) - (\<Sum>k<2*n. ?g k / of_nat (Suc k))"
+      by (simp add: setsum.distrib algebra_simps divide_inverse)
+    also have "(\<Sum>k<2*n. ((-1)^k + ?g k) / real_of_nat (Suc k)) = harm (2*n)"
+      unfolding harm_altdef by (intro setsum.cong) (auto simp: field_simps)
+    also have "(\<Sum>k<2*n. ?g k / real_of_nat (Suc k)) = (\<Sum>k|k<2*n \<and> odd k. ?g k / of_nat (Suc k))"
+      by (intro setsum.mono_neutral_right) auto
+    also have "\<dots> = (\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k)))"
+      by (intro setsum.cong) auto
+    also have "(\<Sum>k|k<2*n \<and> odd k. 2 / (real_of_nat (Suc k))) = harm n" 
+      unfolding harm_altdef
+      by (intro setsum.reindex_cong[of "\<lambda>n. 2*n+1"]) (auto simp: inj_on_def field_simps elim!: oddE)
+    also have "harm (2*n) - harm n = ?em (2*n) - ?em n + ln 2" using n
+      by (simp_all add: algebra_simps ln_mult)
+    finally show "?em (2*n) - ?em n + ln 2 = (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))" ..
+  qed
+  moreover have "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real)) 
+                     \<longlonglongrightarrow> euler_mascheroni - euler_mascheroni + ln 2"
+    by (intro tendsto_intros euler_mascheroni_LIMSEQ filterlim_compose[OF euler_mascheroni_LIMSEQ]
+              filterlim_subseq) (auto simp: subseq_def)
+  hence "(\<lambda>n. ?em (2*n) - ?em n + ln (2::real)) \<longlonglongrightarrow> ln 2" by simp
+  ultimately have "(\<lambda>n. (\<Sum>k<2*n. (-1)^k / real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2"
+    by (rule Lim_transform_eventually)
+  
+  moreover have "summable (\<lambda>k. (-1)^k * inverse (real_of_nat (Suc k)))"
+    using LIMSEQ_inverse_real_of_nat
+    by (intro summable_Leibniz(1) decseq_imp_monoseq decseq_SucI) simp_all
+  hence A: "(\<lambda>n. \<Sum>k<n. (-1)^k / real_of_nat (Suc k)) \<longlonglongrightarrow> (\<Sum>k. (-1)^k / real_of_nat (Suc k))"
+    by (simp add: summable_sums_iff divide_inverse sums_def)
+  from filterlim_compose[OF this filterlim_subseq[of "op * (2::nat)"]]
+    have "(\<lambda>n. \<Sum>k<2*n. (-1)^k / real_of_nat (Suc k)) \<longlonglongrightarrow> (\<Sum>k. (-1)^k / real_of_nat (Suc k))"
+    by (simp add: subseq_def)
+  ultimately have "(\<Sum>k. (- 1) ^ k / real_of_nat (Suc k)) = ln 2" by (intro LIMSEQ_unique)
+  with A show ?thesis by (simp add: sums_def)
+qed
+
+lemma alternating_harmonic_series_sums': 
+  "(\<lambda>k. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2))) sums ln 2"
+unfolding sums_def
+proof (rule Lim_transform_eventually)
+  show "(\<lambda>n. \<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) \<longlonglongrightarrow> ln 2"
+    using alternating_harmonic_series_sums unfolding sums_def 
+    by (rule filterlim_compose) (rule mult_nat_left_at_top, simp)
+  show "eventually (\<lambda>n. (\<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) =
+            (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))) sequentially"
+  proof (intro always_eventually allI)
+    fix n :: nat
+    show "(\<Sum>k<2*n. (-1)^k / (real_of_nat (Suc k))) =
+              (\<Sum>k<n. inverse (real_of_nat (2*k+1)) - inverse (real_of_nat (2*k+2)))"
+      by (induction n) (simp_all add: inverse_eq_divide)
+  qed
+qed               
+
+
+subsection \<open>Approximation of the Euler--Mascheroni constant\<close>
+
+(* FIXME: ugly *)
+(* TODO: Move ? *)
+lemma ln_inverse_approx_le:
+  assumes "(x::real) > 0" "a > 0"
+  shows   "ln (x + a) - ln x \<le> a * (inverse x + inverse (x + a))/2" (is "_ \<le> ?A")
+proof -
+  def f' \<equiv> "(inverse (x + a) - inverse x)/a"
+  have f'_nonpos: "f' \<le> 0" using assms by (simp add: f'_def divide_simps)
+  let ?f = "\<lambda>t. (t - x) * f' + inverse x"
+  let ?F = "\<lambda>t. (t - x)^2 * f' / 2 + t * inverse x"
+  have diff: "\<forall>t\<in>{x..x+a}. (?F has_vector_derivative ?f t) 
+                               (at t within {x..x+a})" using assms
+    by (auto intro!: derivative_eq_intros 
+             simp: has_field_derivative_iff_has_vector_derivative[symmetric])
+  from assms have "(?f has_integral (?F (x+a) - ?F x)) {x..x+a}"
+    by (intro fundamental_theorem_of_calculus[OF _ diff])
+       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] field_simps
+             intro!: derivative_eq_intros)
+  also have "?F (x+a) - ?F x = (a*2 + f'*a\<^sup>2*x) / (2*x)" using assms by (simp add: field_simps)
+  also have "f'*a^2 = - (a^2) / (x*(x + a))" using assms 
+    by (simp add: divide_simps f'_def power2_eq_square)
+  also have "(a*2 + - a\<^sup>2/(x*(x+a))*x) / (2*x) = ?A" using assms
+    by (simp add: divide_simps power2_eq_square) (simp add: algebra_simps)
+  finally have int1: "((\<lambda>t. (t - x) * f' + inverse x) has_integral ?A) {x..x + a}" .
+
+  from assms have int2: "(inverse has_integral (ln (x + a) - ln x)) {x..x+a}"
+    by (intro fundamental_theorem_of_calculus)
+       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps
+             intro!: derivative_eq_intros)
+  hence "ln (x + a) - ln x = integral {x..x+a} inverse" by (simp add: integral_unique)
+  also have ineq: "\<forall>xa\<in>{x..x + a}. inverse xa \<le> (xa - x) * f' + inverse x"
+  proof
+    fix t assume t': "t \<in> {x..x+a}"
+    with assms have t: "0 \<le> (t - x) / a" "(t - x) / a \<le> 1" by simp_all
+    have "inverse t = inverse ((1 - (t - x) / a) *\<^sub>R x + ((t - x) / a) *\<^sub>R (x + a))" (is "_ = ?A")
+      using assms t' by (simp add: field_simps)
+    also from assms have "convex_on {x..x+a} inverse" by (intro convex_on_inverse) auto
+    from convex_onD_Icc[OF this _ t] assms 
+      have "?A \<le> (1 - (t - x) / a) * inverse x + (t - x) / a * inverse (x + a)" by simp
+    also have "\<dots> = (t - x) * f' + inverse x" using assms
+      by (simp add: f'_def divide_simps) (simp add: f'_def field_simps)
+    finally show "inverse t \<le> (t - x) * f' + inverse x" .
+  qed
+  hence "integral {x..x+a} inverse \<le> integral {x..x+a} ?f" using f'_nonpos assms
+    by (intro integral_le has_integral_integrable[OF int1] has_integral_integrable[OF int2] ineq)
+  also have "\<dots> = ?A" using int1 by (rule integral_unique)
+  finally show ?thesis .
+qed
+
+lemma ln_inverse_approx_ge:
+  assumes "(x::real) > 0" "x < y"
+  shows   "ln y - ln x \<ge> 2 * (y - x) / (x + y)" (is "_ \<ge> ?A")
+proof -
+  def m \<equiv> "(x+y)/2"
+  def f' \<equiv> "-inverse (m^2)"
+  from assms have m: "m > 0" by (simp add: m_def)
+  let ?F = "\<lambda>t. (t - m)^2 * f' / 2 + t / m"
+  from assms have "((\<lambda>t. (t - m) * f' + inverse m) has_integral (?F y - ?F x)) {x..y}"
+    by (intro fundamental_theorem_of_calculus)
+       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps
+             intro!: derivative_eq_intros)
+  also from m have "?F y - ?F x = ((y - m)^2 - (x - m)^2) * f' / 2 + (y - x) / m" 
+    by (simp add: field_simps)
+  also have "((y - m)^2 - (x - m)^2) = 0" by (simp add: m_def power2_eq_square field_simps)
+  also have "0 * f' / 2 + (y - x) / m = ?A" by (simp add: m_def)
+  finally have int1: "((\<lambda>t. (t - m) * f' + inverse m) has_integral ?A) {x..y}" .
+
+  from assms have int2: "(inverse has_integral (ln y - ln x)) {x..y}"
+    by (intro fundamental_theorem_of_calculus)
+       (auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] divide_simps
+             intro!: derivative_eq_intros)
+  hence "ln y - ln x = integral {x..y} inverse" by (simp add: integral_unique)
+  also have ineq: "\<forall>xa\<in>{x..y}. inverse xa \<ge> (xa - m) * f' + inverse m"
+  proof
+    fix t assume t: "t \<in> {x..y}"
+    from t assms have "inverse t - inverse m \<ge> f' * (t - m)"
+      by (intro convex_on_imp_above_tangent[of "{0<..}"] convex_on_inverse)
+         (auto simp: m_def interior_open f'_def power2_eq_square intro!: derivative_eq_intros)
+    thus "(t - m) * f' + inverse m \<le> inverse t" by (simp add: algebra_simps)
+  qed
+  hence "integral {x..y} inverse \<ge> integral {x..y} (\<lambda>t. (t - m) * f' + inverse m)"
+    using int1 int2 by (intro integral_le has_integral_integrable)
+  also have "integral {x..y} (\<lambda>t. (t - m) * f' + inverse m) = ?A"
+    using integral_unique[OF int1] by simp
+  finally show ?thesis .
+qed
+
+
+lemma euler_mascheroni_lower: 
+        "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))"
+  and euler_mascheroni_upper:
+        "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))"
+proof -
+  def D \<equiv> "\<lambda>n. inverse (of_nat (n+1)) + ln (of_nat (n+1)) - ln (of_nat (n+2)) :: real"
+  let ?g = "\<lambda>n. ln (of_nat (n+2)) - ln (of_nat (n+1)) - inverse (of_nat (n+1)) :: real"
+  def inv \<equiv> "\<lambda>n. inverse (real_of_nat n)"
+  fix n :: nat
+  note summable = sums_summable[OF euler_mascheroni_sum, folded D_def]
+  have sums: "(\<lambda>k. (inv (Suc (k + (n+1))) - inv (Suc (Suc k + (n+1))))/2) sums ((inv (Suc (0 + (n+1))) - 0)/2)"
+    unfolding inv_def
+    by (intro sums_divide telescope_sums' LIMSEQ_ignore_initial_segment LIMSEQ_inverse_real_of_nat)
+  have sums': "(\<lambda>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) sums ((inv (Suc (0 + n)) - 0)/2)"
+    unfolding inv_def
+    by (intro sums_divide telescope_sums' LIMSEQ_ignore_initial_segment LIMSEQ_inverse_real_of_nat)
+  from euler_mascheroni_sum have "euler_mascheroni = (\<Sum>k. D k)"
+    by (simp add: sums_iff D_def)
+  also have "\<dots> = (\<Sum>k. D (k + Suc n)) + (\<Sum>k\<le>n. D k)"
+    by (subst suminf_split_initial_segment[OF summable, of "Suc n"], subst lessThan_Suc_atMost) simp
+  finally have sum: "(\<Sum>k\<le>n. D k) - euler_mascheroni = -(\<Sum>k. D (k + Suc n))" by simp
+
+  note sum
+  also have "\<dots> \<le> -(\<Sum>k. (inv (k + Suc n + 1) - inv (k + Suc n + 2)) / 2)"
+  proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable])
+    fix k' :: nat
+    def k \<equiv> "k' + Suc n"
+    hence k: "k > 0" by (simp add: k_def)
+    have "real_of_nat (k+1) > 0" by (simp add: k_def)
+    with ln_inverse_approx_le[OF this zero_less_one]
+      have "ln (of_nat k + 2) - ln (of_nat k + 1) \<le> (inv (k+1) + inv (k+2))/2"
+      by (simp add: inv_def add_ac)
+    hence "(inv (k+1) - inv (k+2))/2 \<le> inv (k+1) + ln (of_nat (k+1)) - ln (of_nat (k+2))"
+      by (simp add: field_simps)
+    also have "\<dots> = D k" unfolding D_def inv_def ..
+    finally show "D (k' + Suc n) \<ge> (inv (k' + Suc n + 1) - inv (k' + Suc n + 2)) / 2"
+      by (simp add: k_def)
+    from sums_summable[OF sums] 
+      show "summable (\<lambda>k. (inv (k + Suc n + 1) - inv (k + Suc n + 2))/2)" by simp
+  qed
+  also from sums have "\<dots> = -inv (n+2) / 2" by (simp add: sums_iff)
+  finally have "euler_mascheroni \<ge> (\<Sum>k\<le>n. D k) + 1 / (of_nat (2 * (n+2)))" 
+    by (simp add: inv_def field_simps of_nat_mult)
+  also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
+    unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add:  setsum.distrib setsum_subtractf)
+  also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))"
+    by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all
+  finally show "euler_mascheroni \<ge> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 2))"
+    by simp
+  
+  note sum
+  also have "-(\<Sum>k. D (k + Suc n)) \<ge> -(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)"
+  proof (intro le_imp_neg_le suminf_le allI summable_ignore_initial_segment[OF summable])
+    fix k' :: nat
+    def k \<equiv> "k' + Suc n"
+    hence k: "k > 0" by (simp add: k_def)
+    have "real_of_nat (k+1) > 0" by (simp add: k_def)
+    from ln_inverse_approx_ge[of "of_nat k + 1" "of_nat k + 2"]
+      have "2 / (2 * real_of_nat k + 3) \<le> ln (of_nat (k+2)) - ln (real_of_nat (k+1))"
+      by (simp add: add_ac)
+    hence "D k \<le> 1 / real_of_nat (k+1) - 2 / (2 * real_of_nat k + 3)" 
+      by (simp add: D_def inverse_eq_divide inv_def)
+    also have "\<dots> = inv ((k+1)*(2*k+3))" unfolding inv_def by (simp add: field_simps)
+    also have "\<dots> \<le> inv (2*k*(k+1))" unfolding inv_def using k
+      by (intro le_imp_inverse_le) 
+         (simp add: algebra_simps, simp del: of_nat_add)
+    also have "\<dots> = (inv k - inv (k+1))/2" unfolding inv_def using k
+      by (simp add: divide_simps del: of_nat_mult) (simp add: algebra_simps)
+    finally show "D k \<le> (inv (Suc (k' + n)) - inv (Suc (Suc k' + n)))/2" unfolding k_def by simp
+  next
+    from sums_summable[OF sums'] 
+      show "summable (\<lambda>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2)" by simp
+  qed
+  also from sums' have "(\<Sum>k. (inv (Suc (k + n)) - inv (Suc (Suc k + n)))/2) = inv (n+1)/2"
+    by (simp add: sums_iff)
+  finally have "euler_mascheroni \<le> (\<Sum>k\<le>n. D k) + 1 / of_nat (2 * (n+1))" 
+    by (simp add: inv_def field_simps)
+  also have "(\<Sum>k\<le>n. D k) = harm (Suc n) - (\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1)))"
+    unfolding harm_altdef D_def by (subst lessThan_Suc_atMost) (simp add:  setsum.distrib setsum_subtractf)
+  also have "(\<Sum>k\<le>n. ln (real_of_nat (Suc k+1)) - ln (of_nat (k+1))) = ln (of_nat (n+2))"
+    by (subst atLeast0AtMost [symmetric], subst setsum_Suc_diff) simp_all
+  finally show "euler_mascheroni \<le> harm (Suc n) - ln (real_of_nat (n + 2)) + 1/real_of_nat (2 * (n + 1))"
+    by simp
+qed
+
+lemma euler_mascheroni_pos: "euler_mascheroni > (0::real)"
+  using euler_mascheroni_lower[of 0] ln_2_less_1 by (simp add: harm_def)
+
+lemma ln_approx_aux:
+  fixes n :: nat and x :: real
+  defines "y \<equiv> (x-1)/(x+1)"
+  assumes x: "x > 0" "x \<noteq> 1"
+  shows "inverse (2*y^(2*n+1)) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in> 
+            {0..(1 / (1 - y^2) / of_nat (2*n+1))}"
+proof -
+  from x have norm_y: "norm y < 1" unfolding y_def by simp
+  from power_strict_mono[OF this, of 2] have norm_y': "norm y^2 < 1" by simp
+
+  let ?f = "\<lambda>k. 2 * y ^ (2*k+1) / of_nat (2*k+1)"
+  note sums = ln_series_quadratic[OF x(1)]
+  def c \<equiv> "inverse (2*y^(2*n+1))"
+  let ?d = "c * (ln x - (\<Sum>k<n. ?f k))"
+  have "\<forall>k. y\<^sup>2^k / of_nat (2*(k+n)+1) \<le> y\<^sup>2 ^ k / of_nat (2*n+1)"
+    by (intro allI divide_left_mono mult_right_mono mult_pos_pos zero_le_power[of "y^2"]) simp_all
+  moreover {
+    have "(\<lambda>k. ?f (k + n)) sums (ln x - (\<Sum>k<n. ?f k))"
+      using sums_split_initial_segment[OF sums] by (simp add: y_def)
+    hence "(\<lambda>k. c * ?f (k + n)) sums ?d" by (rule sums_mult)
+    also have "(\<lambda>k. c * (2*y^(2*(k+n)+1) / of_nat (2*(k+n)+1))) =
+                   (\<lambda>k. (c * (2*y^(2*n+1))) * ((y^2)^k / of_nat (2*(k+n)+1)))"
+      by (simp only: ring_distribs power_add power_mult) (simp add: mult_ac)
+    also from x have "c * (2*y^(2*n+1)) = 1" by (simp add: c_def y_def)
+    finally have "(\<lambda>k. (y^2)^k / of_nat (2*(k+n)+1)) sums ?d" by simp
+  } note sums' = this
+  moreover from norm_y' have "(\<lambda>k. (y^2)^k / of_nat (2*n+1)) sums (1 / (1 - y^2) / of_nat (2*n+1))"
+    by (intro sums_divide geometric_sums) (simp_all add: norm_power)
+  ultimately have "?d \<le> (1 / (1 - y^2) / of_nat (2*n+1))" by (rule sums_le)
+  moreover have "c * (ln x - (\<Sum>k<n. 2 * y ^ (2 * k + 1) / real_of_nat (2 * k + 1))) \<ge> 0"
+    by (intro sums_le[OF _ sums_zero sums']) simp_all
+  ultimately show ?thesis unfolding c_def by simp
+qed
+
+lemma
+  fixes n :: nat and x :: real
+  defines "y \<equiv> (x-1)/(x+1)"
+  defines "approx \<equiv> (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))"
+  defines "d \<equiv> y^(2*n+1) / (1 - y^2) / of_nat (2*n+1)"
+  assumes x: "x > 1"
+  shows   ln_approx_bounds: "ln x \<in> {approx..approx + 2*d}"
+  and     ln_approx_abs:    "abs (ln x - (approx + d)) \<le> d"
+proof -
+  def c \<equiv> "2*y^(2*n+1)"
+  from x have c_pos: "c > 0" unfolding c_def y_def 
+    by (intro mult_pos_pos zero_less_power) simp_all
+  have A: "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) \<in>
+              {0.. (1 / (1 - y^2) / of_nat (2*n+1))}" using assms unfolding y_def c_def
+    by (intro ln_approx_aux) simp_all
+  hence "inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1)/of_nat (2*k+1))) \<le> (1 / (1-y^2) / of_nat (2*n+1))"
+    by simp
+  hence "(ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))) / c \<le> (1 / (1 - y^2) / of_nat (2*n+1))" 
+    by (auto simp add: divide_simps)
+  with c_pos have "ln x \<le> c / (1 - y^2) / of_nat (2*n+1) + approx"
+    by (subst (asm) pos_divide_le_eq) (simp_all add: mult_ac approx_def)
+  moreover {
+    from A c_pos have "0 \<le> c * (inverse c * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1))))"
+      by (intro mult_nonneg_nonneg[of c]) simp_all
+    also have "\<dots> = (c * inverse c) * (ln x - (\<Sum>k<n. 2*y^(2*k+1) / of_nat (2*k+1)))"  
+      by (simp add: mult_ac)
+    also from c_pos have "c * inverse c = 1" by simp
+    finally have "ln x \<ge> approx" by (simp add: approx_def)
+  }
+  ultimately show "ln x \<in> {approx..approx + 2*d}" by (simp add: c_def d_def)
+  thus "abs (ln x - (approx + d)) \<le> d" by auto
+qed
+
+context
+begin
+
+qualified lemma ln_approx_abs': 
+  assumes "x > (1::real)"
+  assumes "(x-1)/(x+1) = y"
+  assumes "y^2 = ysqr"
+  assumes "(\<Sum>k<n. inverse (of_nat (2*k+1)) * ysqr^k) = approx"
+  assumes "y*ysqr^n / (1 - ysqr) / of_nat (2*n+1) = d"
+  assumes "d \<le> e"
+  shows   "abs (ln x - (2*y*approx + d)) \<le> e"
+proof -
+  note ln_approx_abs[OF assms(1), of n]
+  also note assms(2)
+  also have "y^(2*n+1) = y*ysqr^n" by (simp add: assms(3)[symmetric] power_mult)
+  also note assms(3)
+  also note assms(5)
+  also note assms(5)
+  also note assms(6)
+  also have "(\<Sum>k<n. 2*y^(2*k+1) / real_of_nat (2 * k + 1)) = (2*y) * approx"
+    apply (subst assms(4)[symmetric], subst setsum_right_distrib)
+    apply (simp add: assms(3)[symmetric] power_mult)
+    apply (simp add: mult_ac divide_simps)?
+    done
+  finally show ?thesis .
+qed
+
+lemma ln_2_approx: "\<bar>ln 2 - 0.69314718055\<bar> < inverse (2 ^ 36 :: real)" (is ?thesis1)
+  and ln_2_bounds: "ln (2::real) \<in> {0.693147180549..0.693147180561}" (is ?thesis2)
+proof -
+  def approx \<equiv> "0.69314718055 :: real" and approx' \<equiv> "4465284211343447 / 6442043387911560 :: real"
+  def d \<equiv> "inverse (195259926456::real)"
+  have "dist (ln 2) approx \<le> dist (ln 2) approx' + dist approx' approx" by (rule dist_triangle)
+  also have "\<bar>ln (2::real) - (2 * (1/3) * (651187280816108 / 626309773824735) +
+                 inverse 195259926456)\<bar> \<le> inverse 195259926456"
+  proof (rule ln_approx_abs'[where n = 10])
+    show "(1/3::real)^2 = 1/9" by (simp add: power2_eq_square)
+  qed (simp_all add: eval_nat_numeral)
+  hence A: "dist (ln 2) approx' \<le> d" by (simp add: dist_real_def approx'_def d_def)
+  hence "dist (ln 2) approx' + dist approx' approx \<le> \<dots> + dist approx' approx"
+    by (rule add_right_mono)
+  also have "\<dots> < inverse (2 ^ 36)" by (simp add: dist_real_def approx'_def approx_def d_def)
+  finally show ?thesis1 unfolding dist_real_def approx_def .
+  
+  from A have "ln 2 \<in> {approx' - d..approx' + d}" 
+    by (simp add: dist_real_def abs_real_def split: split_if_asm)
+  also have "\<dots> \<subseteq> {0.693147180549..0.693147180561}"
+    by (subst atLeastatMost_subset_iff, rule disjI2) (simp add: approx'_def d_def)
+  finally show ?thesis2 .
+qed
+
+end
+
+
+lemma euler_mascheroni_bounds:
+  fixes n :: nat assumes "n \<ge> 1" defines "t \<equiv> harm n - ln (of_nat (Suc n)) :: real"
+  shows "euler_mascheroni \<in> {t + inverse (of_nat (2*(n+1)))..t + inverse (of_nat (2*n))}"
+  using assms euler_mascheroni_upper[of "n-1"] euler_mascheroni_lower[of "n-1"]
+  unfolding t_def by (cases n) (simp_all add: harm_Suc t_def inverse_eq_divide of_nat_mult)
+
+lemma euler_mascheroni_bounds':
+  fixes n :: nat assumes "n \<ge> 1" "ln (real_of_nat (Suc n)) \<in> {l<..<u}"
+  shows "euler_mascheroni \<in> 
+           {harm n - u + inverse (of_nat (2*(n+1)))<..<harm n - l + inverse (of_nat (2*n))}"
+  using euler_mascheroni_bounds[OF assms(1)] assms(2) by auto
+
+lemma euler_mascheroni_approx: 
+  defines "approx \<equiv> 0.577257 :: real" and "e \<equiv> 0.000063 :: real"
+  shows   "abs (euler_mascheroni - approx :: real) < e"
+  (is "abs (_ - ?approx) < ?e")
+proof -
+  def l \<equiv> "47388813395531028639296492901910937/82101866951584879688289000000000000 :: real"
+  def u \<equiv> "142196984054132045946501548559032969 / 246305600854754639064867000000000000 :: real"
+  have impI: "P \<longrightarrow> Q" if Q for P Q using that by blast
+  have hsum_63: "harm 63 = (310559566510213034489743057 / 65681493561267903750631200 ::real)"
+    by (simp add: harm_expand)
+  from harm_Suc[of 63] have hsum_64: "harm 64 = 
+          623171679694215690971693339 / (131362987122535807501262400::real)" 
+    by (subst (asm) hsum_63) simp
+  have "ln (64::real) = real (6::nat) * ln 2" by (subst ln_realpow[symmetric]) simp_all
+  hence "ln (real_of_nat (Suc 63)) \<in> {4.158883083293<..<4.158883083367}" using ln_2_bounds by simp
+  from euler_mascheroni_bounds'[OF _ this]
+    have "(euler_mascheroni :: real) \<in> {l<..<u}" 
+    by (simp add: hsum_63 del: greaterThanLessThan_iff) (simp only: l_def u_def)
+  also have "\<dots> \<subseteq> {approx - e<..<approx + e}"
+    by (subst greaterThanLessThan_subseteq_greaterThanLessThan, rule impI) 
+       (simp add: approx_def e_def u_def l_def)
+  finally show ?thesis by (simp add: abs_real_def)
+qed
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Integral_Test.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -0,0 +1,111 @@
+(*
+  Title:    HOL/Multivariate_Analysis/Integral_Test.thy
+  Author:   Manuel Eberl, TU München
+  
+  The integral test for summability. We show here that for a decreasing non-negative 
+  function, the infinite sum over that function evaluated at the natural numbers 
+  converges iff the corresponding integral converges.
+  
+  As a useful side result, we also provide some results on the difference between
+  the integral and the partial sum. (This is useful e.g. for the definition of the
+  Euler–Mascheroni constant)
+*)
+theory Integral_Test
+imports Integration
+begin
+
+subsubsection \<open>Integral test\<close>
+
+(* TODO: continuous_in \<rightarrow> integrable_on *)
+locale antimono_fun_sum_integral_diff =
+  fixes f :: "real \<Rightarrow> real"
+  assumes dec: "\<And>x y. x \<ge> 0 \<Longrightarrow> x \<le> y \<Longrightarrow> f x \<ge> f y"
+  assumes nonneg: "\<And>x. x \<ge> 0 \<Longrightarrow> f x \<ge> 0"
+  assumes cont: "continuous_on {0..} f"
+begin
+
+definition "sum_integral_diff_series n = (\<Sum>k\<le>n. f (of_nat k)) - (integral {0..of_nat n} f)"
+
+lemma sum_integral_diff_series_nonneg:
+  "sum_integral_diff_series n \<ge> 0"
+proof -
+  note int = integrable_continuous_real[OF continuous_on_subset[OF cont]]
+  let ?int = "\<lambda>a b. integral {of_nat a..of_nat b} f"
+  have "-sum_integral_diff_series n = ?int 0 n - (\<Sum>k\<le>n. f (of_nat k))" 
+    by (simp add: sum_integral_diff_series_def)
+  also have "?int 0 n = (\<Sum>k<n. ?int k (Suc k))"
+  proof (induction n)
+    case (Suc n)
+    have "?int 0 (Suc n) = ?int 0 n + ?int n (Suc n)"
+      by (intro integral_combine[symmetric] int) simp_all
+    with Suc show ?case by simp
+  qed simp_all
+  also have "... \<le> (\<Sum>k<n. integral {of_nat k..of_nat (Suc k)} (\<lambda>_::real. f (of_nat k)))"
+    by (intro setsum_mono integral_le int) (auto intro: dec)
+  also have "... = (\<Sum>k<n. f (of_nat k))" by simp
+  also have "\<dots> - (\<Sum>k\<le>n. f (of_nat k)) = -(\<Sum>k\<in>{..n} - {..<n}. f (of_nat k))"
+    by (subst setsum_diff) auto
+  also have "\<dots> \<le> 0" by (auto intro!: setsum_nonneg nonneg)
+  finally show "sum_integral_diff_series n \<ge> 0" by simp
+qed
+
+lemma sum_integral_diff_series_antimono:
+  assumes "m \<le> n"
+  shows   "sum_integral_diff_series m \<ge> sum_integral_diff_series n"
+proof -
+  let ?int = "\<lambda>a b. integral {of_nat a..of_nat b} f"
+  note int = integrable_continuous_real[OF continuous_on_subset[OF cont]]
+  have d_mono: "sum_integral_diff_series (Suc n) \<le> sum_integral_diff_series n" for n
+  proof -
+    fix n :: nat
+    have "sum_integral_diff_series (Suc n) - sum_integral_diff_series n = 
+            f (of_nat (Suc n)) + (?int 0 n - ?int 0 (Suc n))"
+      unfolding sum_integral_diff_series_def by (simp add: algebra_simps)
+    also have "?int 0 n - ?int 0 (Suc n) = -?int n (Suc n)"
+      by (subst integral_combine [symmetric, of "of_nat 0" "of_nat n" "of_nat (Suc n)"])
+         (auto intro!: int simp: algebra_simps)
+    also have "?int n (Suc n) \<ge> integral {of_nat n..of_nat (Suc n)} (\<lambda>_::real. f (of_nat (Suc n)))"
+      by (intro integral_le int) (auto intro: dec)
+    hence "f (of_nat (Suc n)) + -?int n (Suc n) \<le> 0" by (simp add: algebra_simps)
+    finally show "sum_integral_diff_series (Suc n) \<le> sum_integral_diff_series n" by simp
+  qed
+  with assms show ?thesis
+    by (induction rule: inc_induct) (auto intro: order.trans[OF _ d_mono])
+qed
+
+lemma sum_integral_diff_series_Bseq: "Bseq sum_integral_diff_series"
+proof -
+  from sum_integral_diff_series_nonneg and sum_integral_diff_series_antimono 
+    have "norm (sum_integral_diff_series n) \<le> sum_integral_diff_series 0" for n by simp
+  thus "Bseq sum_integral_diff_series" by (rule BseqI')
+qed
+
+lemma sum_integral_diff_series_monoseq: "monoseq sum_integral_diff_series"
+  using sum_integral_diff_series_antimono unfolding monoseq_def by blast
+
+lemma sum_integral_diff_series_convergent: "convergent sum_integral_diff_series"
+  using sum_integral_diff_series_Bseq sum_integral_diff_series_monoseq
+  by (blast intro!: Bseq_monoseq_convergent)
+
+lemma integral_test:
+  "summable (\<lambda>n. f (of_nat n)) \<longleftrightarrow> convergent (\<lambda>n. integral {0..of_nat n} f)"
+proof -
+  have "summable (\<lambda>n. f (of_nat n)) \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k\<le>n. f (of_nat k))"
+    by (simp add: summable_iff_convergent')
+  also have "... \<longleftrightarrow> convergent (\<lambda>n. integral {0..of_nat n} f)"
+  proof
+    assume "convergent (\<lambda>n. \<Sum>k\<le>n. f (of_nat k))"
+    from convergent_diff[OF this sum_integral_diff_series_convergent] 
+      show "convergent (\<lambda>n. integral {0..of_nat n} f)" 
+        unfolding sum_integral_diff_series_def by simp
+  next
+    assume "convergent (\<lambda>n. integral {0..of_nat n} f)"
+    from convergent_add[OF this sum_integral_diff_series_convergent] 
+      show "convergent (\<lambda>n. \<Sum>k\<le>n. f (of_nat k))" unfolding sum_integral_diff_series_def by simp
+  qed
+  finally show ?thesis by simp
+qed
+
+end
+
+end
\ No newline at end of file
--- a/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Multivariate_Analysis/Multivariate_Analysis.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -8,6 +8,8 @@
   Bounded_Continuous_Function
   Weierstrass
   Cauchy_Integral_Thm
+  Generalised_Binomial_Theorem
+  Gamma
 begin
 
 end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Multivariate_Analysis/Summation.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -0,0 +1,907 @@
+(*
+  Title:    HOL/Multivariate_Analysis/Summation.thy
+  Author:   Manuel Eberl, TU München
+  
+  The definition of the radius of convergence of a power series, 
+  various summability tests, lemmas to compute the radius of convergence etc.
+*)
+theory Summation
+imports
+  Complex_Main
+  "~~/src/HOL/Library/Extended_Real" 
+  "~~/src/HOL/Library/Liminf_Limsup"
+begin
+
+subsection \<open>Rounded dual logarithm\<close>
+
+(* This is required for the Cauchy condensation criterion *)
+
+definition "natlog2 n = (if n = 0 then 0 else nat \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
+
+lemma natlog2_0 [simp]: "natlog2 0 = 0" by (simp add: natlog2_def)
+lemma natlog2_1 [simp]: "natlog2 1 = 0" by (simp add: natlog2_def)
+lemma natlog2_eq_0_iff: "natlog2 n = 0 \<longleftrightarrow> n < 2" by (simp add: natlog2_def)
+
+lemma natlog2_power_of_two [simp]: "natlog2 (2 ^ n) = n"
+  by (simp add: natlog2_def log_nat_power)
+
+lemma natlog2_mono: "m \<le> n \<Longrightarrow> natlog2 m \<le> natlog2 n"
+  unfolding natlog2_def by (simp_all add: nat_mono floor_mono)
+
+lemma pow_natlog2_le: "n > 0 \<Longrightarrow> 2 ^ natlog2 n \<le> n"
+proof -
+  assume n: "n > 0"
+  from n have "of_nat (2 ^ natlog2 n) = 2 powr real_of_nat (nat \<lfloor>log 2 (real_of_nat n)\<rfloor>)"
+    by (subst powr_realpow) (simp_all add: natlog2_def)
+  also have "\<dots> = 2 powr of_int \<lfloor>log 2 (real_of_nat n)\<rfloor>" using n by simp
+  also have "\<dots> \<le> 2 powr log 2 (real_of_nat n)" by (intro powr_mono) (linarith, simp_all)
+  also have "\<dots> = of_nat n" using n by simp
+  finally show ?thesis by simp
+qed
+
+lemma pow_natlog2_gt: "n > 0 \<Longrightarrow> 2 * 2 ^ natlog2 n > n"
+  and pow_natlog2_ge: "n > 0 \<Longrightarrow> 2 * 2 ^ natlog2 n \<ge> n"
+proof -
+  assume n: "n > 0"
+  from n have "of_nat n = 2 powr log 2 (real_of_nat n)" by simp
+  also have "\<dots> < 2 powr (1 + of_int \<lfloor>log 2 (real_of_nat n)\<rfloor>)" 
+    by (intro powr_less_mono) (linarith, simp_all)
+  also from n have "\<dots> = 2 powr (1 + real_of_nat (nat \<lfloor>log 2 (real_of_nat n)\<rfloor>))" by simp
+  also from n have "\<dots> = of_nat (2 * 2 ^ natlog2 n)"
+    by (simp_all add: natlog2_def powr_real_of_int powr_add)
+  finally show "2 * 2 ^ natlog2 n > n" by (rule of_nat_less_imp_less)
+  thus "2 * 2 ^ natlog2 n \<ge> n" by simp
+qed
+
+lemma natlog2_eqI:
+  assumes "n > 0" "2^k \<le> n" "n < 2 * 2^k"
+  shows   "natlog2 n = k"
+proof -
+  from assms have "of_nat (2 ^ k) \<le> real_of_nat n"  by (subst of_nat_le_iff) simp_all
+  hence "real_of_int (int k) \<le> log (of_nat 2) (real_of_nat n)"
+    by (subst le_log_iff) (simp_all add: powr_realpow assms del: of_nat_le_iff)
+  moreover from assms have "real_of_nat n < of_nat (2 ^ Suc k)" by (subst of_nat_less_iff) simp_all
+  hence "log 2 (real_of_nat n) < of_nat k + 1"
+    by (subst log_less_iff) (simp_all add: assms powr_realpow powr_add)
+  ultimately have "\<lfloor>log 2 (real_of_nat n)\<rfloor> = of_nat k" by (intro floor_unique) simp_all
+  with assms show ?thesis by (simp add: natlog2_def)
+qed
+
+lemma natlog2_rec: 
+  assumes "n \<ge> 2"
+  shows   "natlog2 n = 1 + natlog2 (n div 2)"
+proof (rule natlog2_eqI)
+  from assms have "2 ^ (1 + natlog2 (n div 2)) \<le> 2 * (n div 2)" 
+    by (simp add: pow_natlog2_le)
+  also have "\<dots> \<le> n" by simp
+  finally show "2 ^ (1 + natlog2 (n div 2)) \<le> n" .
+next
+  from assms have "n < 2 * (n div 2 + 1)" by simp 
+  also from assms have "(n div 2) < 2 ^ (1 + natlog2 (n div 2))" 
+    by (simp add: pow_natlog2_gt)
+  hence "2 * (n div 2 + 1) \<le> 2 * (2 ^ (1 + natlog2 (n div 2)))" 
+    by (intro mult_left_mono) simp_all
+  finally show "n < 2 * 2 ^ (1 + natlog2 (n div 2))" .
+qed (insert assms, simp_all)
+
+fun natlog2_aux where
+  "natlog2_aux n acc = (if (n::nat) < 2 then acc else natlog2_aux (n div 2) (acc + 1))"
+
+lemma natlog2_aux_correct:
+  "natlog2_aux n acc = acc + natlog2 n"
+  by (induction n acc rule: natlog2_aux.induct) (auto simp: natlog2_rec natlog2_eq_0_iff)
+  
+lemma natlog2_code [code]: "natlog2 n = natlog2_aux n 0"
+  by (subst natlog2_aux_correct) simp
+
+
+subsection \<open>Convergence tests for infinite sums\<close>
+
+subsubsection \<open>Root test\<close>
+
+lemma limsup_root_powser:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  shows "limsup (\<lambda>n. ereal (root n (norm (f n * z ^ n)))) = 
+             limsup (\<lambda>n. ereal (root n (norm (f n)))) * ereal (norm z)"
+proof -
+  have A: "(\<lambda>n. ereal (root n (norm (f n * z ^ n)))) = 
+              (\<lambda>n. ereal (root n (norm (f n))) * ereal (norm z))" (is "?g = ?h")
+  proof
+    fix n show "?g n = ?h n"
+    by (cases "n = 0") (simp_all add: norm_mult real_root_mult real_root_pos2 norm_power)
+  qed
+  show ?thesis by (subst A, subst limsup_ereal_mult_right) simp_all
+qed
+
+lemma limsup_root_limit:
+  assumes "(\<lambda>n. ereal (root n (norm (f n)))) \<longlonglongrightarrow> l" (is "?g \<longlonglongrightarrow> _")
+  shows   "limsup (\<lambda>n. ereal (root n (norm (f n)))) = l"
+proof -
+  from assms have "convergent ?g" "lim ?g = l"
+    unfolding convergent_def by (blast intro: limI)+
+  with convergent_limsup_cl show ?thesis by force
+qed
+
+lemma limsup_root_limit':
+  assumes "(\<lambda>n. root n (norm (f n))) \<longlonglongrightarrow> l"
+  shows   "limsup (\<lambda>n. ereal (root n (norm (f n)))) = ereal l"
+  by (intro limsup_root_limit tendsto_ereal assms)
+
+lemma root_test_convergence':
+  fixes f :: "nat \<Rightarrow> 'a :: banach"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  assumes l: "l < 1"
+  shows   "summable f"
+proof -
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have "l \<ge> 0" by simp
+  with l obtain l' where l': "l = ereal l'" by (cases l) simp_all
+
+  def c \<equiv> "(1 - l') / 2"
+  from l and `l \<ge> 0` have c: "l + c > l" "l' + c \<ge> 0" "l' + c < 1" unfolding c_def 
+    by (simp_all add: field_simps l')
+  have "\<forall>C>l. eventually (\<lambda>n. ereal (root n (norm (f n))) < C) sequentially"
+    by (subst Limsup_le_iff[symmetric]) (simp add: l_def)
+  with c have "eventually (\<lambda>n. ereal (root n (norm (f n))) < l + ereal c) sequentially" by simp
+  with eventually_gt_at_top[of "0::nat"]
+    have "eventually (\<lambda>n. norm (f n) \<le> (l' + c) ^ n) sequentially"
+  proof eventually_elim
+    fix n :: nat assume n: "n > 0" 
+    assume "ereal (root n (norm (f n))) < l + ereal c"
+    hence "root n (norm (f n)) \<le> l' + c" by (simp add: l')
+    with c n have "root n (norm (f n)) ^ n \<le> (l' + c) ^ n"
+      by (intro power_mono) (simp_all add: real_root_ge_zero)
+    also from n have "root n (norm (f n)) ^ n = norm (f n)" by simp
+    finally show "norm (f n) \<le> (l' + c) ^ n" by simp
+  qed
+  thus ?thesis
+    by (rule summable_comparison_test_ev[OF _ summable_geometric]) (simp add: c)
+qed
+
+lemma root_test_divergence:
+  fixes f :: "nat \<Rightarrow> 'a :: banach"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  assumes l: "l > 1"
+  shows   "\<not>summable f"
+proof
+  assume "summable f"
+  hence bounded: "Bseq f" by (simp add: summable_imp_Bseq)
+
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have l_nonneg: "l \<ge> 0" by simp
+
+  def c \<equiv> "if l = \<infinity> then 2 else 1 + (real_of_ereal l - 1) / 2"
+  from l l_nonneg consider "l = \<infinity>" | "\<exists>l'. l = ereal l'" by (cases l) simp_all
+  hence c: "c > 1 \<and> ereal c < l" by cases (insert l, auto simp: c_def field_simps)
+
+  have unbounded: "\<not>bdd_above {n. root n (norm (f n)) > c}"
+  proof
+    assume "bdd_above {n. root n (norm (f n)) > c}"
+    then obtain N where "\<forall>n. root n (norm (f n)) > c \<longrightarrow> n \<le> N" unfolding bdd_above_def by blast
+    hence "\<exists>N. \<forall>n\<ge>N. root n (norm (f n)) \<le> c"
+      by (intro exI[of _ "N + 1"]) (force simp: not_less_eq_eq[symmetric])
+    hence "eventually (\<lambda>n. root n (norm (f n)) \<le> c) sequentially"
+      by (auto simp: eventually_at_top_linorder)
+    hence "l \<le> c" unfolding l_def by (intro Limsup_bounded) simp_all
+    with c show False by auto
+  qed
+  
+  from bounded obtain K where K: "K > 0" "\<And>n. norm (f n) \<le> K" using BseqE by blast
+  def n \<equiv> "nat \<lceil>log c K\<rceil>"
+  from unbounded have "\<exists>m>n. c < root m (norm (f m))" unfolding bdd_above_def
+    by (auto simp: not_le)
+  then guess m by (elim exE conjE) note m = this
+  from c K have "K = c powr log c K" by (simp add: powr_def log_def)
+  also from c have "c powr log c K \<le> c powr real n" unfolding n_def
+    by (intro powr_mono, linarith, simp)
+  finally have "K \<le> c ^ n" using c by (simp add: powr_realpow)
+  also from c m have "c ^ n < c ^ m" by simp
+  also from c m have "c ^ m < root m (norm (f m)) ^ m" by (intro power_strict_mono) simp_all
+  also from m have "... = norm (f m)" by simp
+  finally show False using K(2)[of m]  by simp
+qed
+
+
+subsection \<open>Cauchy's condensation test\<close>
+
+context
+fixes f :: "nat \<Rightarrow> real"
+begin
+
+private lemma condensation_inequality:
+  assumes mono: "\<And>m n. 0 < m \<Longrightarrow> m \<le> n \<Longrightarrow> f n \<le> f m"
+  shows   "(\<Sum>k=1..<n. f k) \<ge> (\<Sum>k=1..<n. f (2 * 2 ^ natlog2 k))" (is "?thesis1")
+          "(\<Sum>k=1..<n. f k) \<le> (\<Sum>k=1..<n. f (2 ^ natlog2 k))" (is "?thesis2")
+  by (intro setsum_mono mono pow_natlog2_ge pow_natlog2_le, simp, simp)+
+
+private lemma condensation_condense1: "(\<Sum>k=1..<2^n. f (2 ^ natlog2 k)) = (\<Sum>k<n. 2^k * f (2 ^ k))"
+proof (induction n)
+  case (Suc n)
+  have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto  
+  also have "(\<Sum>k\<in>\<dots>. f (2 ^ natlog2 k)) = 
+                 (\<Sum>k<n. 2^k * f (2^k)) + (\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k))" 
+    by (subst setsum.union_disjoint) (insert Suc, auto)
+  also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
+  hence "(\<Sum>k = 2^n..<2^Suc n. f (2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^n))"
+    by (intro setsum.cong) simp_all
+  also have "\<dots> = 2^n * f (2^n)" by (simp add: of_nat_power)
+  finally show ?case by simp
+qed simp
+
+private lemma condensation_condense2: "(\<Sum>k=1..<2^n. f (2 * 2 ^ natlog2 k)) = (\<Sum>k<n. 2^k * f (2 ^ Suc k))"
+proof (induction n)
+  case (Suc n)
+  have "{1..<2^Suc n} = {1..<2^n} \<union> {2^n..<(2^Suc n :: nat)}" by auto  
+  also have "(\<Sum>k\<in>\<dots>. f (2 * 2 ^ natlog2 k)) = 
+                 (\<Sum>k<n. 2^k * f (2^Suc k)) + (\<Sum>k = 2^n..<2^Suc n. f (2 * 2^natlog2 k))" 
+    by (subst setsum.union_disjoint) (insert Suc, auto)
+  also have "natlog2 k = n" if "k \<in> {2^n..<2^Suc n}" for k using that by (intro natlog2_eqI) simp_all
+  hence "(\<Sum>k = 2^n..<2^Suc n. f (2*2^natlog2 k)) = (\<Sum>(_::nat) = 2^n..<2^Suc n. f (2^Suc n))"
+    by (intro setsum.cong) simp_all
+  also have "\<dots> = 2^n * f (2^Suc n)" by (simp add: of_nat_power)
+  finally show ?case by simp
+qed simp
+
+lemma condensation_test:
+  assumes mono: "\<And>m. 0 < m \<Longrightarrow> f (Suc m) \<le> f m"
+  assumes nonneg: "\<And>n. f n \<ge> 0"
+  shows "summable f \<longleftrightarrow> summable (\<lambda>n. 2^n * f (2^n))"
+proof -
+  def f' \<equiv> "\<lambda>n. if n = 0 then 0 else f n"
+  from mono have mono': "decseq (\<lambda>n. f (Suc n))" by (intro decseq_SucI) simp
+  hence mono': "f n \<le> f m" if "m \<le> n" "m > 0" for m n 
+    using that decseqD[OF mono', of "m - 1" "n - 1"] by simp
+  
+  have "(\<lambda>n. f (Suc n)) = (\<lambda>n. f' (Suc n))" by (intro ext) (simp add: f'_def)
+  hence "summable f \<longleftrightarrow> summable f'"
+    by (subst (1 2) summable_Suc_iff [symmetric]) (simp only:)
+  also have "\<dots> \<longleftrightarrow> convergent (\<lambda>n. \<Sum>k<n. f' k)" unfolding summable_iff_convergent ..
+  also have "monoseq (\<lambda>n. \<Sum>k<n. f' k)" unfolding f'_def
+    by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
+  hence "convergent (\<lambda>n. \<Sum>k<n. f' k) \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. f' k)"
+    by (rule monoseq_imp_convergent_iff_Bseq)
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f' k)" unfolding One_nat_def
+    by (subst setsum_shift_lb_Suc0_0_upt) (simp_all add: f'_def atLeast0LessThan)
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<n. f k)" unfolding f'_def by simp
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
+    by (rule nonneg_incseq_Bseq_subseq_iff[symmetric])
+       (auto intro!: setsum_nonneg incseq_SucI nonneg simp: subseq_def)
+  also have "\<dots> \<longleftrightarrow> Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^k))"
+  proof (intro iffI)
+    assume A: "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)"
+    have "eventually (\<lambda>n. norm (\<Sum>k<n. 2^k * f (2^Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)) sequentially"
+    proof (intro always_eventually allI)
+      fix n :: nat
+      have "norm (\<Sum>k<n. 2^k * f (2^Suc k)) = (\<Sum>k<n. 2^k * f (2^Suc k))" unfolding real_norm_def
+        by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
+      also have "\<dots> \<le> (\<Sum>k=1..<2^n. f k)"
+        by (subst condensation_condense2 [symmetric]) (intro condensation_inequality mono')
+      also have "\<dots> = norm \<dots>" unfolding real_norm_def
+        by (intro abs_of_nonneg[symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg)
+      finally show "norm (\<Sum>k<n. 2 ^ k * f (2 ^ Suc k)) \<le> norm (\<Sum>k=1..<2^n. f k)" .
+    qed
+    from this and A have "Bseq (\<lambda>n. \<Sum>k<n. 2^k * f (2^Suc k))" by (rule Bseq_eventually_mono)
+    from Bseq_mult[OF Bfun_const[of 2] this] have "Bseq (\<lambda>n. \<Sum>k<n. 2^Suc k * f (2^Suc k))"
+      by (simp add: setsum_right_distrib setsum_left_distrib mult_ac)
+    hence "Bseq (\<lambda>n. (\<Sum>k=Suc 0..<Suc n. 2^k * f (2^k)) + f 1)"
+      by (intro Bseq_add, subst setsum_shift_bounds_Suc_ivl) (simp add: atLeast0LessThan)
+    hence "Bseq (\<lambda>n. (\<Sum>k=0..<Suc n. 2^k * f (2^k)))"
+      by (subst setsum_head_upt_Suc) (simp_all add: add_ac)
+    thus "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))" 
+      by (subst (asm) Bseq_Suc_iff) (simp add: atLeast0LessThan)
+  next
+    assume A: "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
+    have "eventually (\<lambda>n. norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))) sequentially"
+    proof (intro always_eventually allI)
+      fix n :: nat
+      have "norm (\<Sum>k=1..<2^n. f k) = (\<Sum>k=1..<2^n. f k)" unfolding real_norm_def
+        by (intro abs_of_nonneg setsum_nonneg ballI mult_nonneg_nonneg nonneg)
+      also have "\<dots> \<le> (\<Sum>k<n. 2^k * f (2^k))"
+        by (subst condensation_condense1 [symmetric]) (intro condensation_inequality mono')
+      also have "\<dots> = norm \<dots>" unfolding real_norm_def
+        by (intro abs_of_nonneg [symmetric] setsum_nonneg ballI mult_nonneg_nonneg nonneg) simp_all
+      finally show "norm (\<Sum>k=1..<2^n. f k) \<le> norm (\<Sum>k<n. 2^k * f (2^k))" .
+    qed
+    from this and A show "Bseq (\<lambda>n. \<Sum>k=1..<2^n. f k)" by (rule Bseq_eventually_mono)
+  qed
+  also have "monoseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
+    by (intro mono_SucI1) (auto intro!: mult_nonneg_nonneg nonneg)
+  hence "Bseq (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k))) \<longleftrightarrow> convergent (\<lambda>n. (\<Sum>k<n. 2^k * f (2^k)))"
+    by (rule monoseq_imp_convergent_iff_Bseq [symmetric])
+  also have "\<dots> \<longleftrightarrow> summable (\<lambda>k. 2^k * f (2^k))" by (simp only: summable_iff_convergent)
+  finally show ?thesis .
+qed
+
+end
+
+
+subsection \<open>Summability of powers\<close>
+
+lemma abs_summable_complex_powr_iff: 
+    "summable (\<lambda>n. norm (exp (of_real (ln (of_nat n)) * s))) \<longleftrightarrow> Re s < -1"
+proof (cases "Re s \<le> 0")
+  let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
+  case False
+  with eventually_gt_at_top[of "0::nat"]
+    have "eventually (\<lambda>n. norm (1 :: real) \<le> norm (exp (?l n * s))) sequentially" 
+    by (auto intro!: ge_one_powr_ge_zero elim!: eventually_mono)
+  from summable_comparison_test_ev[OF this] False show ?thesis by (auto simp: summable_const_iff)
+next
+  let ?l = "\<lambda>n. complex_of_real (ln (of_nat n))"
+  case True
+  hence "summable (\<lambda>n. norm (exp (?l n * s))) \<longleftrightarrow> summable (\<lambda>n. 2^n * norm (exp (?l (2^n) * s)))"
+    by (intro condensation_test) (auto intro!: mult_right_mono_neg)
+  also have "(\<lambda>n. 2^n * norm (exp (?l (2^n) * s))) = (\<lambda>n. (2 powr (Re s + 1)) ^ n)"
+  proof
+    fix n :: nat
+    have "2^n * norm (exp (?l (2^n) * s)) = exp (real n * ln 2) * exp (real n * ln 2 * Re s)"
+      using True by (subst exp_of_nat_mult) (simp add: ln_realpow algebra_simps) 
+    also have "\<dots> = exp (real n * (ln 2 * (Re s + 1)))"
+      by (simp add: algebra_simps exp_add)
+    also have "\<dots> = exp (ln 2 * (Re s + 1)) ^ n" by (subst exp_of_nat_mult) simp
+    also have "exp (ln 2 * (Re s + 1)) = 2 powr (Re s + 1)" by (simp add: powr_def)
+    finally show "2^n * norm (exp (?l (2^n) * s)) = (2 powr (Re s + 1)) ^ n" .
+  qed
+  also have "summable \<dots> \<longleftrightarrow> 2 powr (Re s + 1) < 2 powr 0"
+    by (subst summable_geometric_iff) simp
+  also have "\<dots> \<longleftrightarrow> Re s < -1" by (subst powr_less_cancel_iff) (simp, linarith)
+  finally show ?thesis .
+qed
+
+lemma summable_complex_powr_iff: 
+  assumes "Re s < -1"
+  shows   "summable (\<lambda>n. exp (of_real (ln (of_nat n)) * s))"
+  by (rule summable_norm_cancel, subst abs_summable_complex_powr_iff) fact
+
+lemma summable_real_powr_iff: "summable (\<lambda>n. of_nat n powr s :: real) \<longleftrightarrow> s < -1"
+proof -
+  from eventually_gt_at_top[of "0::nat"]
+    have "summable (\<lambda>n. of_nat n powr s) \<longleftrightarrow> summable (\<lambda>n. exp (ln (of_nat n) * s))"
+    by (intro summable_cong) (auto elim!: eventually_mono simp: powr_def)
+  also have "\<dots> \<longleftrightarrow> s < -1" using abs_summable_complex_powr_iff[of "of_real s"] by simp
+  finally show ?thesis .
+qed
+
+lemma inverse_power_summable:
+  assumes s: "s \<ge> 2"
+  shows "summable (\<lambda>n. inverse (of_nat n ^ s :: 'a :: {real_normed_div_algebra,banach}))"
+proof (rule summable_norm_cancel, subst summable_cong)
+  from eventually_gt_at_top[of "0::nat"]
+    show "eventually (\<lambda>n. norm (inverse (of_nat n ^ s:: 'a)) = real_of_nat n powr (-real s)) at_top"
+    by eventually_elim (simp add: norm_inverse norm_power powr_minus powr_realpow)
+qed (insert s summable_real_powr_iff[of "-s"], simp_all)
+
+lemma not_summable_harmonic: "\<not>summable (\<lambda>n. inverse (of_nat n) :: 'a :: real_normed_field)"
+proof
+  assume "summable (\<lambda>n. inverse (of_nat n) :: 'a)"
+  hence "convergent (\<lambda>n. norm (of_real (\<Sum>k<n. inverse (of_nat k)) :: 'a))" 
+    by (simp add: summable_iff_convergent convergent_norm)
+  hence "convergent (\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real)" by (simp only: norm_of_real)
+  also have "(\<lambda>n. abs (\<Sum>k<n. inverse (of_nat k)) :: real) = (\<lambda>n. \<Sum>k<n. inverse (of_nat k))"
+    by (intro ext abs_of_nonneg setsum_nonneg) auto
+  also have "convergent \<dots> \<longleftrightarrow> summable (\<lambda>k. inverse (of_nat k) :: real)"
+    by (simp add: summable_iff_convergent)
+  finally show False using summable_real_powr_iff[of "-1"] by (simp add: powr_minus)
+qed
+
+
+subsection \<open>Kummer's test\<close>
+
+lemma kummers_test_convergence:
+  fixes f p :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
+  assumes nonneg_p: "eventually (\<lambda>n. p n \<ge> 0) sequentially"
+  defines "l \<equiv> liminf (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
+  assumes l: "l > 0"
+  shows   "summable f"
+  unfolding summable_iff_convergent'
+proof -
+  def r \<equiv> "(if l = \<infinity> then 1 else real_of_ereal l / 2)"
+  from l have "r > 0 \<and> of_real r < l" by (cases l) (simp_all add: r_def)
+  hence r: "r > 0" "of_real r < l" by simp_all
+  hence "eventually (\<lambda>n. p n * f n / f (Suc n) - p (Suc n) > r) sequentially"
+    unfolding l_def by (force dest: less_LiminfD)
+  moreover from pos_f have "eventually (\<lambda>n. f (Suc n) > 0) sequentially" 
+    by (subst eventually_sequentially_Suc)
+  ultimately have "eventually (\<lambda>n. p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)) sequentially"
+    by eventually_elim (simp add: field_simps)
+  from eventually_conj[OF pos_f eventually_conj[OF nonneg_p this]]
+    obtain m where m: "\<And>n. n \<ge> m \<Longrightarrow> f n > 0" "\<And>n. n \<ge> m \<Longrightarrow> p n \<ge> 0"
+        "\<And>n. n \<ge> m \<Longrightarrow> p n * f n - p (Suc n) * f (Suc n) > r * f (Suc n)"
+    unfolding eventually_at_top_linorder by blast
+
+  let ?c = "(norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r"
+  have "Bseq (\<lambda>n. (\<Sum>k\<le>n + Suc m. f k))"
+  proof (rule BseqI')
+    fix k :: nat
+    def n \<equiv> "k + Suc m"
+    have n: "n > m" by (simp add: n_def)
+
+    from r have "r * norm (\<Sum>k\<le>n. f k) = norm (\<Sum>k\<le>n. r * f k)"
+      by (simp add: setsum_right_distrib[symmetric] abs_mult)
+    also from n have "{..n} = {..m} \<union> {Suc m..n}" by auto
+    hence "(\<Sum>k\<le>n. r * f k) = (\<Sum>k\<in>{..m} \<union> {Suc m..n}. r * f k)" by (simp only:)
+    also have "\<dots> = (\<Sum>k\<le>m. r * f k) + (\<Sum>k=Suc m..n. r * f k)"
+      by (subst setsum.union_disjoint) auto
+    also have "norm \<dots> \<le> norm (\<Sum>k\<le>m. r * f k) + norm (\<Sum>k=Suc m..n. r * f k)"
+      by (rule norm_triangle_ineq)
+    also from r less_imp_le[OF m(1)] have "(\<Sum>k=Suc m..n. r * f k) \<ge> 0" 
+      by (intro setsum_nonneg) auto
+    hence "norm (\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=Suc m..n. r * f k)" by simp
+    also have "(\<Sum>k=Suc m..n. r * f k) = (\<Sum>k=m..<n. r * f (Suc k))"
+     by (subst setsum_shift_bounds_Suc_ivl [symmetric])
+          (simp only: atLeastLessThanSuc_atLeastAtMost)
+    also from m have "\<dots> \<le> (\<Sum>k=m..<n. p k * f k - p (Suc k) * f (Suc k))"
+      by (intro setsum_mono[OF less_imp_le]) simp_all
+    also have "\<dots> = -(\<Sum>k=m..<n. p (Suc k) * f (Suc k) - p k * f k)"
+      by (simp add: setsum_negf [symmetric] algebra_simps)
+    also from n have "\<dots> = p m * f m - p n * f n"
+      by (cases n, simp, simp only: atLeastLessThanSuc_atLeastAtMost, subst setsum_Suc_diff) simp_all
+    also from less_imp_le[OF m(1)] m(2) n have "\<dots> \<le> p m * f m" by simp
+    finally show "norm (\<Sum>k\<le>n. f k) \<le> (norm (\<Sum>k\<le>m. r * f k) + p m * f m) / r" using r
+      by (subst pos_le_divide_eq[OF r(1)]) (simp only: mult_ac)
+  qed
+  moreover have "(\<Sum>k\<le>n. f k) \<le> (\<Sum>k\<le>n'. f k)" if "Suc m \<le> n" "n \<le> n'" for n n'
+    using less_imp_le[OF m(1)] that by (intro setsum_mono2) auto
+  ultimately show "convergent (\<lambda>n. \<Sum>k\<le>n. f k)" by (rule Bseq_monoseq_convergent'_inc)
+qed
+
+
+lemma kummers_test_divergence:
+  fixes f p :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
+  assumes pos_p: "eventually (\<lambda>n. p n > 0) sequentially"
+  assumes divergent_p: "\<not>summable (\<lambda>n. inverse (p n))"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (p n * f n / f (Suc n) - p (Suc n)))"
+  assumes l: "l < 0"
+  shows   "\<not>summable f"
+proof
+  assume "summable f"
+  from eventually_conj[OF pos_f eventually_conj[OF pos_p Limsup_lessD[OF l[unfolded l_def]]]]
+    obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> p n > 0" "\<And>n. n \<ge> N \<Longrightarrow> f n > 0"
+                      "\<And>n. n \<ge> N \<Longrightarrow> p n * f n / f (Suc n) - p (Suc n) < 0"
+    by (auto simp: eventually_at_top_linorder)
+  hence A: "p n * f n < p (Suc n) * f (Suc n)" if "n \<ge> N" for n using that N[of n] N[of "Suc n"] 
+    by (simp add: field_simps)
+  have "p n * f n \<ge> p N * f N" if "n \<ge> N" for n using that and A
+      by (induction n rule: dec_induct) (auto intro!: less_imp_le elim!: order.trans)
+  from eventually_ge_at_top[of N] N this
+    have "eventually (\<lambda>n. norm (p N * f N * inverse (p n)) \<le> f n) sequentially"
+    by (auto elim!: eventually_mono simp: field_simps abs_of_pos)
+  from this and \<open>summable f\<close> have "summable (\<lambda>n. p N * f N * inverse (p n))"
+    by (rule summable_comparison_test_ev)
+  from summable_mult[OF this, of "inverse (p N * f N)"] N[OF le_refl] 
+    have "summable (\<lambda>n. inverse (p n))" by (simp add: divide_simps)
+  with divergent_p show False by contradiction
+qed
+
+
+subsection \<open>Ratio test\<close>
+
+lemma ratio_test_convergence:
+  fixes f :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
+  defines "l \<equiv> liminf (\<lambda>n. ereal (f n / f (Suc n)))"
+  assumes l: "l > 1"
+  shows   "summable f"
+proof (rule kummers_test_convergence[OF pos_f])
+  note l
+  also have "l = liminf (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1" 
+    by (subst Liminf_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
+  finally show "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) > 0"
+    by (cases "liminf (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
+qed simp
+
+lemma ratio_test_divergence:
+  fixes f :: "nat \<Rightarrow> real"
+  assumes pos_f: "eventually (\<lambda>n. f n > 0) sequentially" 
+  defines "l \<equiv> limsup (\<lambda>n. ereal (f n / f (Suc n)))"
+  assumes l: "l < 1"
+  shows   "\<not>summable f"
+proof (rule kummers_test_divergence[OF pos_f])
+  have "limsup (\<lambda>n. ereal (f n / f (Suc n) - 1)) + 1 = l" 
+    by (subst Limsup_add_ereal_right[symmetric]) (simp_all add: minus_ereal_def l_def one_ereal_def)
+  also note l
+  finally show "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1)) < 0"
+    by (cases "limsup (\<lambda>n. ereal (1 * f n / f (Suc n) - 1))") simp_all
+qed (simp_all add: summable_const_iff)
+
+
+subsection \<open>Raabe's test\<close>
+
+lemma raabes_test_convergence:
+fixes f :: "nat \<Rightarrow> real"
+  assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
+  defines "l \<equiv> liminf (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
+  assumes l: "l > 1"
+  shows   "summable f"
+proof (rule kummers_test_convergence)
+  let ?l' = "liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
+  have "1 < l" by fact
+  also have "l = liminf (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
+    by (simp add: l_def algebra_simps)
+  also have "\<dots> = ?l' + 1" by (subst Liminf_add_ereal_right) simp_all
+  finally show "?l' > 0" by (cases ?l') (simp_all add: algebra_simps)
+qed (simp_all add: pos)
+
+lemma raabes_test_divergence:
+fixes f :: "nat \<Rightarrow> real"
+  assumes pos: "eventually (\<lambda>n. f n > 0) sequentially"
+  defines "l \<equiv> limsup (\<lambda>n. ereal (of_nat n * (f n / f (Suc n) - 1)))"
+  assumes l: "l < 1"
+  shows   "\<not>summable f"
+proof (rule kummers_test_divergence)
+  let ?l' = "limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)))"
+  note l
+  also have "l = limsup (\<lambda>n. ereal (of_nat n * f n / f (Suc n) - of_nat (Suc n)) + 1)"
+    by (simp add: l_def algebra_simps)
+  also have "\<dots> = ?l' + 1" by (subst Limsup_add_ereal_right) simp_all
+  finally show "?l' < 0" by (cases ?l') (simp_all add: algebra_simps)
+qed (insert pos eventually_gt_at_top[of "0::nat"] not_summable_harmonic, simp_all)
+
+
+
+subsection \<open>Radius of convergence\<close>
+
+text \<open>
+  The radius of convergence of a power series. This value always exists, ranges from
+  @{term "0::ereal"} to @{term "\<infinity>::ereal"}, and the power series is guaranteed to converge for 
+  all inputs with a norm that is smaller than that radius and to diverge for all inputs with a
+  norm that is greater. 
+\<close>
+definition conv_radius :: "(nat \<Rightarrow> 'a :: banach) \<Rightarrow> ereal" where
+  "conv_radius f = inverse (limsup (\<lambda>n. ereal (root n (norm (f n)))))"
+
+lemma conv_radius_nonneg: "conv_radius f \<ge> 0"
+proof -
+  have "0 = limsup (\<lambda>n. 0)" by (subst Limsup_const) simp_all
+  also have "\<dots> \<le> limsup (\<lambda>n. ereal (root n (norm (f n))))"
+    by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally show ?thesis
+    unfolding conv_radius_def by (auto simp: ereal_inverse_nonneg_iff)
+qed
+
+lemma conv_radius_zero [simp]: "conv_radius (\<lambda>_. 0) = \<infinity>"
+  by (auto simp: conv_radius_def zero_ereal_def [symmetric] Limsup_const)
+
+lemma conv_radius_cong:
+  assumes "eventually (\<lambda>x. f x = g x) sequentially"
+  shows   "conv_radius f = conv_radius g"
+proof -
+  have "eventually (\<lambda>n. ereal (root n (norm (f n))) = ereal (root n (norm (g n)))) sequentially"
+    using assms by eventually_elim simp
+  from Limsup_eq[OF this] show ?thesis unfolding conv_radius_def by simp
+qed
+
+lemma conv_radius_altdef:
+  "conv_radius f = liminf (\<lambda>n. inverse (ereal (root n (norm (f n)))))"
+  by (subst Liminf_inverse_ereal) (simp_all add: real_root_ge_zero conv_radius_def)
+
+
+lemma abs_summable_in_conv_radius:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "ereal (norm z) < conv_radius f"
+  shows   "summable (\<lambda>n. norm (f n * z ^ n))"
+proof (rule root_test_convergence')
+  def l \<equiv> "limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have l_nonneg: "l \<ge> 0" .
+
+  have "limsup (\<lambda>n. root n (norm (f n * z^n))) = l * ereal (norm z)" unfolding l_def
+    by (rule limsup_root_powser)
+  also from l_nonneg consider "l = 0" | "l = \<infinity>" | "\<exists>l'. l = ereal l' \<and> l' > 0"
+    by (cases "l") (auto simp: less_le)
+  hence "l * ereal (norm z) < 1"
+  proof cases
+    assume "l = \<infinity>"
+    hence "conv_radius f = 0" unfolding conv_radius_def l_def by simp
+    with assms show ?thesis by simp
+  next
+    assume "\<exists>l'. l = ereal l' \<and> l' > 0"
+    then guess l' by (elim exE conjE) note l' = this
+    hence "l \<noteq> \<infinity>" by auto
+    have "l * ereal (norm z) < l * conv_radius f"
+      by (intro ereal_mult_strict_left_mono) (simp_all add: l' assms)
+    also have "conv_radius f = inverse l" by (simp add: conv_radius_def l_def)
+    also from l' have "l * inverse l = 1" by simp
+    finally show ?thesis .
+  qed simp_all
+  finally show "limsup (\<lambda>n. ereal (root n (norm (norm (f n * z ^ n))))) < 1" by simp
+qed
+
+lemma summable_in_conv_radius:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "ereal (norm z) < conv_radius f"
+  shows   "summable (\<lambda>n. f n * z ^ n)"
+  by (rule summable_norm_cancel, rule abs_summable_in_conv_radius) fact+
+
+lemma not_summable_outside_conv_radius:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "ereal (norm z) > conv_radius f"
+  shows   "\<not>summable (\<lambda>n. f n * z ^ n)"
+proof (rule root_test_divergence)
+  def l \<equiv> "limsup (\<lambda>n. ereal (root n (norm (f n))))"
+  have "0 = limsup (\<lambda>n. 0)" by (simp add: Limsup_const)
+  also have "... \<le> l" unfolding l_def by (intro Limsup_mono) (simp_all add: real_root_ge_zero)
+  finally have l_nonneg: "l \<ge> 0" .
+  from assms have l_nz: "l \<noteq> 0" unfolding conv_radius_def l_def by auto
+
+  have "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) = l * ereal (norm z)"
+    unfolding l_def by (rule limsup_root_powser)
+  also have "... > 1"
+  proof (cases l)
+    assume "l = \<infinity>"
+    with assms conv_radius_nonneg[of f] show ?thesis
+      by (auto simp: zero_ereal_def[symmetric])
+  next
+    fix l' assume l': "l = ereal l'"
+    from l_nonneg l_nz have "1 = l * inverse l" by (auto simp: l' field_simps)
+    also from l_nz have "inverse l = conv_radius f" 
+      unfolding l_def conv_radius_def by auto
+    also from l' l_nz l_nonneg assms have "l * \<dots> < l * ereal (norm z)"
+      by (intro ereal_mult_strict_left_mono) (auto simp: l')
+    finally show ?thesis .
+  qed (insert l_nonneg, simp_all)
+  finally show "limsup (\<lambda>n. ereal (root n (norm (f n * z^n)))) > 1" .
+qed
+
+
+lemma conv_radius_geI:
+  assumes "summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
+  shows   "conv_radius f \<ge> norm z"
+  using not_summable_outside_conv_radius[of f z] assms by (force simp: not_le[symmetric])
+
+lemma conv_radius_leI:
+  assumes "\<not>summable (\<lambda>n. norm (f n * z ^ n :: 'a :: {banach, real_normed_div_algebra}))"
+  shows   "conv_radius f \<le> norm z"
+  using abs_summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
+
+lemma conv_radius_leI':
+  assumes "\<not>summable (\<lambda>n. f n * z ^ n :: 'a :: {banach, real_normed_div_algebra})"
+  shows   "conv_radius f \<le> norm z"
+  using summable_in_conv_radius[of z f] assms by (force simp: not_le[symmetric])
+
+lemma conv_radius_geI_ex:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f \<ge> R"
+proof (rule linorder_cases[of "conv_radius f" R])
+  assume R: "conv_radius f < R"
+  with conv_radius_nonneg[of f] obtain conv_radius' 
+    where [simp]: "conv_radius f = ereal conv_radius'"
+    by (cases "conv_radius f") simp_all
+  def r \<equiv> "if R = \<infinity> then conv_radius' + 1 else (real_of_ereal R + conv_radius') / 2"
+  from R conv_radius_nonneg[of f] have "0 < r \<and> ereal r < R \<and> ereal r > conv_radius f" 
+    unfolding r_def by (cases R) (auto simp: r_def field_simps)
+  with assms(1)[of r] obtain z where "norm z > conv_radius f" "summable (\<lambda>n. f n * z^n)" by auto
+  with not_summable_outside_conv_radius[of f z] show ?thesis by simp
+qed simp_all
+
+lemma conv_radius_geI_ex':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * of_real r^n)"
+  shows   "conv_radius f \<ge> R"
+proof (rule conv_radius_geI_ex)
+  fix r assume "0 < r" "ereal r < R"
+  with assms[of r] show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)"
+    by (intro exI[of _ "of_real r :: 'a"]) auto
+qed
+
+lemma conv_radius_leI_ex:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
+  shows   "conv_radius f \<le> R"
+proof (rule linorder_cases[of "conv_radius f" R])
+  assume R: "conv_radius f > R"
+  from R assms(1) obtain R' where R': "R = ereal R'" by (cases R) simp_all
+  def r \<equiv> "if conv_radius f = \<infinity> then R' + 1 else (R' + real_of_ereal (conv_radius f)) / 2"
+  from R conv_radius_nonneg[of f] have "r > R \<and> r < conv_radius f" unfolding r_def
+    by (cases "conv_radius f") (auto simp: r_def field_simps R')
+  with assms(1) assms(2)[of r] R' 
+    obtain z where "norm z < conv_radius f" "\<not>summable (\<lambda>n. norm (f n * z^n))" by auto
+  with abs_summable_in_conv_radius[of z f] show ?thesis by auto
+qed simp_all
+
+lemma conv_radius_leI_ex':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. f n * of_real r^n)"
+  shows   "conv_radius f \<le> R"
+proof (rule conv_radius_leI_ex)
+  fix r assume "0 < r" "ereal r > R"
+  with assms(2)[of r] show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))"
+    by (intro exI[of _ "of_real r :: 'a"]) (auto dest: summable_norm_cancel)
+qed fact+
+
+lemma conv_radius_eqI:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z^n))"
+  shows   "conv_radius f = R"
+  by (intro antisym conv_radius_geI_ex conv_radius_leI_ex assms)
+
+lemma conv_radius_eqI':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach, real_normed_div_algebra}"
+  assumes "R \<ge> 0"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r < R \<Longrightarrow> summable (\<lambda>n. f n * (of_real r)^n)"
+  assumes "\<And>r. 0 < r \<Longrightarrow> ereal r > R \<Longrightarrow> \<not>summable (\<lambda>n. norm (f n * (of_real r)^n))"
+  shows   "conv_radius f = R"
+proof (intro conv_radius_eqI[OF assms(1)])
+  fix r assume "0 < r" "ereal r < R" with assms(2)[OF this] 
+    show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z ^ n)" by force
+next
+  fix r assume "0 < r" "ereal r > R" with assms(3)[OF this] 
+    show "\<exists>z. norm z = r \<and> \<not>summable (\<lambda>n. norm (f n * z ^ n))" by force  
+qed
+
+lemma conv_radius_zeroI:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>z. z \<noteq> 0 \<Longrightarrow> \<not>summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = 0"
+proof (rule ccontr)
+  assume "conv_radius f \<noteq> 0"
+  with conv_radius_nonneg[of f] have pos: "conv_radius f > 0" by simp
+  def r \<equiv> "if conv_radius f = \<infinity> then 1 else real_of_ereal (conv_radius f) / 2"
+  from pos have r: "ereal r > 0 \<and> ereal r < conv_radius f" 
+    by (cases "conv_radius f") (simp_all add: r_def)
+  hence "summable (\<lambda>n. f n * of_real r ^ n)" by (intro summable_in_conv_radius) simp
+  moreover from r and assms[of "of_real r"] have "\<not>summable (\<lambda>n. f n * of_real r ^ n)" by simp
+  ultimately show False by contradiction
+qed
+
+lemma conv_radius_inftyI':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>r. r > c \<Longrightarrow> \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = \<infinity>"
+proof -
+  {
+    fix r :: real
+    have "max r (c + 1) > c" by (auto simp: max_def)
+    from assms[OF this] obtain z where "norm z = max r (c + 1)" "summable (\<lambda>n. f n * z^n)" by blast
+    from conv_radius_geI[OF this(2)] this(1) have "conv_radius f \<ge> r" by simp
+  }
+  from this[of "real_of_ereal (conv_radius f + 1)"] show "conv_radius f = \<infinity>"
+    by (cases "conv_radius f") simp_all
+qed
+
+lemma conv_radius_inftyI:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>r. \<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = \<infinity>"
+  using assms by (rule conv_radius_inftyI')
+
+lemma conv_radius_inftyI'':
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "\<And>z. summable (\<lambda>n. f n * z^n)"
+  shows   "conv_radius f = \<infinity>"
+proof (rule conv_radius_inftyI')
+  fix r :: real assume "r > 0"
+  with assms show "\<exists>z. norm z = r \<and> summable (\<lambda>n. f n * z^n)"
+    by (intro exI[of _ "of_real r"]) simp
+qed
+
+lemma conv_radius_ratio_limit_ereal:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes nz:  "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
+  assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c"
+proof (rule conv_radius_eqI')
+  show "c \<ge> 0" by (intro Lim_bounded2_ereal[OF lim]) simp_all
+next
+  fix r assume r: "0 < r" "ereal r < c"
+  let ?l = "liminf (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
+  have "?l = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
+    using r by (simp add: norm_mult norm_power divide_simps)
+  also from r have "\<dots> = liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
+    by (intro Liminf_ereal_mult_right) simp_all
+  also have "liminf (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
+    by (intro lim_imp_Liminf lim) simp
+  finally have l: "?l = c * ereal (inverse r)" by simp
+  from r have  l': "c * ereal (inverse r) > 1" by (cases c) (simp_all add: field_simps)
+  show "summable (\<lambda>n. f n * of_real r^n)"
+    by (rule summable_norm_cancel, rule ratio_test_convergence)
+       (insert r nz l l', auto elim!: eventually_mono)
+next
+  fix r assume r: "0 < r" "ereal r > c"
+  let ?l = "limsup (\<lambda>n. ereal (norm (f n * of_real r ^ n) / norm (f (Suc n) * of_real r ^ Suc n)))"
+  have "?l = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n)))) * ereal (inverse r))"
+    using r by (simp add: norm_mult norm_power divide_simps)
+  also from r have "\<dots> = limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) * ereal (inverse r)"
+    by (intro Limsup_ereal_mult_right) simp_all
+  also have "limsup (\<lambda>n. ereal (norm (f n) / (norm (f (Suc n))))) = c"
+    by (intro lim_imp_Limsup lim) simp
+  finally have l: "?l = c * ereal (inverse r)" by simp
+  from r have  l': "c * ereal (inverse r) < 1" by (cases c) (simp_all add: field_simps)
+  show "\<not>summable (\<lambda>n. norm (f n * of_real r^n))"
+    by (rule ratio_test_divergence) (insert r nz l l', auto elim!: eventually_mono)
+qed
+
+lemma conv_radius_ratio_limit_ereal_nonzero:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes nz:  "c \<noteq> 0"
+  assumes lim: "(\<lambda>n. ereal (norm (f n) / norm (f (Suc n)))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c"
+proof (rule conv_radius_ratio_limit_ereal[OF _ lim], rule ccontr)
+  assume "\<not>eventually (\<lambda>n. f n \<noteq> 0) sequentially"
+  hence "frequently (\<lambda>n. f n = 0) sequentially" by (simp add: frequently_def)
+  hence "frequently (\<lambda>n. ereal (norm (f n) / norm (f (Suc n))) = 0) sequentially"
+    by (force elim!: frequently_elim1)
+  hence "c = 0" by (intro limit_frequently_eq[OF _ _ lim]) auto
+  with nz show False by contradiction
+qed 
+
+lemma conv_radius_ratio_limit:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "c' = ereal c"
+  assumes nz:  "eventually (\<lambda>n. f n \<noteq> 0) sequentially"
+  assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c'"
+  using assms by (intro conv_radius_ratio_limit_ereal) simp_all
+  
+lemma conv_radius_ratio_limit_nonzero:
+  fixes f :: "nat \<Rightarrow> 'a :: {banach,real_normed_div_algebra}"
+  assumes "c' = ereal c"
+  assumes nz:  "c \<noteq> 0"
+  assumes lim: "(\<lambda>n. norm (f n) / norm (f (Suc n))) \<longlonglongrightarrow> c"
+  shows   "conv_radius f = c'"
+  using assms by (intro conv_radius_ratio_limit_ereal_nonzero) simp_all
+
+lemma conv_radius_mult_power: 
+  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
+  shows   "conv_radius (\<lambda>n. c ^ n * f n) = conv_radius f / ereal (norm c)"
+proof - 
+  have "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) =
+          limsup (\<lambda>n. ereal (norm c) * ereal (root n (norm (f n))))" 
+    using eventually_gt_at_top[of "0::nat"]
+    by (intro Limsup_eq) 
+       (auto elim!: eventually_mono simp: norm_mult norm_power real_root_mult real_root_power)
+  also have "\<dots> = ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))"
+    using assms by (subst Limsup_ereal_mult_left[symmetric]) simp_all
+  finally have A: "limsup (\<lambda>n. ereal (root n (norm (c ^ n * f n)))) = 
+                       ereal (norm c) * limsup (\<lambda>n. ereal (root n (norm (f n))))" .
+  show ?thesis using assms
+    apply (cases "limsup (\<lambda>n. ereal (root n (norm (f n)))) = 0")
+    apply (simp add: A conv_radius_def)
+    apply (unfold conv_radius_def A divide_ereal_def, simp add: mult.commute ereal_inverse_mult)
+    done
+qed
+
+lemma conv_radius_mult_power_right: 
+  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
+  shows   "conv_radius (\<lambda>n. f n * c ^ n) = conv_radius f / ereal (norm c)"
+  using conv_radius_mult_power[OF assms, of f]
+  unfolding conv_radius_def by (simp add: mult.commute norm_mult)
+
+lemma conv_radius_divide_power: 
+  assumes "c \<noteq> (0 :: 'a :: {real_normed_div_algebra,banach})"
+  shows   "conv_radius (\<lambda>n. f n / c^n) = conv_radius f * ereal (norm c)"
+proof - 
+  from assms have "inverse c \<noteq> 0" by simp
+  from conv_radius_mult_power_right[OF this, of f] show ?thesis
+    by (simp add: divide_inverse divide_ereal_def assms norm_inverse power_inverse)
+qed
+
+
+lemma conv_radius_add_ge: 
+  "min (conv_radius f) (conv_radius g) \<le> 
+       conv_radius (\<lambda>x. f x + g x :: 'a :: {banach,real_normed_div_algebra})"
+  by (rule conv_radius_geI_ex')
+     (auto simp: algebra_simps intro!: summable_add summable_in_conv_radius)
+
+lemma conv_radius_mult_ge:
+  fixes f g :: "nat \<Rightarrow> ('a :: {banach,real_normed_div_algebra})"
+  shows "conv_radius (\<lambda>x. \<Sum>i\<le>x. f i * g (x - i)) \<ge> min (conv_radius f) (conv_radius g)"
+proof (rule conv_radius_geI_ex')
+  fix r assume r: "r > 0" "ereal r < min (conv_radius f) (conv_radius g)"
+  from r have "summable (\<lambda>n. (\<Sum>i\<le>n. (f i * of_real r^i) * (g (n - i) * of_real r^(n - i))))"
+    by (intro summable_Cauchy_product abs_summable_in_conv_radius) simp_all
+  thus "summable (\<lambda>n. (\<Sum>i\<le>n. f i * g (n - i)) * of_real r ^ n)"
+    by (simp add: algebra_simps of_real_def scaleR_power power_add [symmetric] scaleR_setsum_right)
+qed
+
+end
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -5388,23 +5388,6 @@
     unfolding uniformly_continuous_on_def by blast
 qed
 
-lemma continuous_on_tendsto_compose:
-  assumes f_cont: "continuous_on s f"
-  assumes g: "(g \<longlongrightarrow> l) F"
-  assumes l: "l \<in> s"
-  assumes ev: "\<forall>\<^sub>F x in F. g x \<in> s"
-  shows "((\<lambda>x. f (g x)) \<longlongrightarrow> f l) F"
-proof -
-  from f_cont have f: "(f \<longlongrightarrow> f l) (at l within s)"
-    by (auto simp: l continuous_on)
-  have i: "((\<lambda>x. if g x = l then f l else f (g x)) \<longlongrightarrow> f l) F"
-    by (rule filterlim_If)
-      (auto intro!: filterlim_compose[OF f] eventually_conj tendsto_mono[OF _ g]
-        simp: filterlim_at eventually_inf_principal eventually_mono[OF ev])
-  show ?thesis
-    by (rule filterlim_cong[THEN iffD1[OF _ i]]) auto
-qed
-
 
 text\<open>The usual transformation theorems.\<close>
 
--- a/src/HOL/Real_Vector_Spaces.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Real_Vector_Spaces.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -191,7 +191,7 @@
   fixes x :: "'a::real_vector"
   shows "scaleR (-1) x = - x"
   using scaleR_minus_left [of 1 x] by simp
-
+  
 class real_algebra = real_vector + ring +
   assumes mult_scaleR_left [simp]: "scaleR a x * y = scaleR a (x * y)"
   and mult_scaleR_right [simp]: "x * scaleR a y = scaleR a (x * y)"
@@ -686,6 +686,10 @@
 
 class real_normed_algebra_1 = real_algebra_1 + real_normed_algebra +
   assumes norm_one [simp]: "norm 1 = 1"
+  
+lemma (in real_normed_algebra_1) scaleR_power [simp]: 
+  "(scaleR x y) ^ n = scaleR (x^n) (y^n)"
+  by (induction n) (simp_all add: scaleR_one scaleR_scaleR mult_ac)
 
 class real_normed_div_algebra = real_div_algebra + real_normed_vector +
   assumes norm_mult: "norm (x * y) = norm x * norm y"
--- a/src/HOL/Series.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Series.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -791,6 +791,12 @@
   using a b
   by (rule Cauchy_product_sums [THEN sums_unique])
 
+lemma summable_Cauchy_product:
+  assumes "summable (\<lambda>k. norm (a k :: 'a :: {real_normed_algebra,banach}))" 
+          "summable (\<lambda>k. norm (b k))"
+  shows   "summable (\<lambda>k. \<Sum>i\<le>k. a i * b (k - i))"
+  using Cauchy_product_sums[OF assms] by (simp add: sums_iff)  
+
 subsection \<open>Series on @{typ real}s\<close>
 
 lemma summable_norm_comparison_test: "\<exists>N. \<forall>n\<ge>N. norm (f n) \<le> g n \<Longrightarrow> summable g \<Longrightarrow> summable (\<lambda>n. norm (f n))"
--- a/src/HOL/Topological_Spaces.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Topological_Spaces.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -1621,6 +1621,23 @@
 
 lemma isCont_tendsto_compose: "isCont g l \<Longrightarrow> (f \<longlongrightarrow> l) F \<Longrightarrow> ((\<lambda>x. g (f x)) \<longlongrightarrow> g l) F"
   unfolding isCont_def by (rule tendsto_compose)
+  
+lemma continuous_on_tendsto_compose:
+  assumes f_cont: "continuous_on s f"
+  assumes g: "(g \<longlongrightarrow> l) F"
+  assumes l: "l \<in> s"
+  assumes ev: "\<forall>\<^sub>F x in F. g x \<in> s"
+  shows "((\<lambda>x. f (g x)) \<longlongrightarrow> f l) F"
+proof -
+  from f_cont l have f: "(f \<longlongrightarrow> f l) (at l within s)"
+    by (simp add: continuous_on_def)
+  have i: "((\<lambda>x. if g x = l then f l else f (g x)) \<longlongrightarrow> f l) F"
+    by (rule filterlim_If)
+       (auto intro!: filterlim_compose[OF f] eventually_conj tendsto_mono[OF _ g]
+             simp: filterlim_at eventually_inf_principal eventually_mono[OF ev])
+  show ?thesis
+    by (rule filterlim_cong[THEN iffD1[OF _ i]]) auto
+qed
 
 lemma continuous_within_compose3:
   "isCont g (f x) \<Longrightarrow> continuous (at x within s) f \<Longrightarrow> continuous (at x within s) (\<lambda>x. g (f x))"
--- a/src/HOL/Transcendental.thy	Sun Jan 03 21:45:34 2016 +0100
+++ b/src/HOL/Transcendental.thy	Mon Jan 04 17:45:36 2016 +0100
@@ -2354,6 +2354,11 @@
 lemma powr_numeral: "0 < x \<Longrightarrow> x powr (numeral n :: real) = x ^ (numeral n)"
   by (metis of_nat_numeral powr_realpow)
 
+lemma powr_real_of_int: 
+  "x > 0 \<Longrightarrow> x powr real_of_int n = (if n \<ge> 0 then x ^ nat n else inverse (x ^ nat (-n)))"
+  using powr_realpow[of x "nat n"] powr_realpow[of x "nat (-n)"]
+  by (auto simp: field_simps powr_minus)  
+
 lemma powr2_sqrt[simp]: "0 < x \<Longrightarrow> sqrt x powr 2 = x"
 by(simp add: powr_numeral)