mathlib3 documentation

analysis.special_functions.gamma.basic

The Gamma function #

THIS FILE IS SYNCHRONIZED WITH MATHLIB4. Any changes to this file require a corresponding PR to mathlib4.

This file defines the Γ function (of a real or complex variable s). We define this by Euler's integral Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1) in the range where this integral converges (i.e., for 0 < s in the real case, and 0 < re s in the complex case).

We show that this integral satisfies Γ(1) = 1 and Γ(s + 1) = s * Γ(s); hence we can define Γ(s) for all s as the unique function satisfying this recurrence and agreeing with Euler's integral in the convergence range. (If s = -n for n ∈ ℕ, then the function is undefined, and we set it to be 0 by convention.)

Gamma function: main statements (complex case) #

Gamma function: main statements (real case) #

Tags #

Gamma

theorem real.Gamma_integrand_is_o (s : ) :
(λ (x : ), rexp (-x) * x ^ s) =o[filter.at_top] λ (x : ), rexp (-(1 / 2) * x)

Asymptotic bound for the Γ function integrand.

The Euler integral for the Γ function converges for positive real s.

The integral defining the Γ function converges for complex s with 0 < re s.

This is proved by reduction to the real case.

noncomputable def complex.Gamma_integral (s : ) :

Euler's integral for the Γ function (of a complex variable s), defined as ∫ x in Ioi 0, exp (-x) * x ^ (s - 1).

See complex.Gamma_integral_convergent for a proof of the convergence of the integral for 0 < re s.

Equations

Now we establish the recurrence relation Γ(s + 1) = s * Γ(s) using integration by parts.

noncomputable def complex.partial_Gamma (s : ) (X : ) :

The indefinite version of the Γ function, Γ(s, X) = ∫ x ∈ 0..X, exp(-x) x ^ (s - 1).

Equations
theorem complex.partial_Gamma_add_one {s : } (hs : 0 < s.re) {X : } (hX : 0 X) :
(s + 1).partial_Gamma X = s * s.partial_Gamma X - (rexp (-X)) * X ^ s

The recurrence relation for the indefinite version of the Γ function.

theorem complex.Gamma_integral_add_one {s : } (hs : 0 < s.re) :

The recurrence relation for the Γ integral.

Now we define Γ(s) on the whole complex plane, by recursion.

noncomputable def complex.Gamma_aux  :

The nth function in this family is Γ(s) if -n < s.re, and junk otherwise.

Equations
theorem complex.Gamma_aux_recurrence1 (s : ) (n : ) (h1 : -s.re < n) :
theorem complex.Gamma_aux_recurrence2 (s : ) (n : ) (h1 : -s.re < n) :
noncomputable def complex.Gamma (s : ) :

The Γ function (of a complex variable s).

Equations
theorem complex.Gamma_add_one (s : ) (h2 : s 0) :

The recurrence relation for the Γ function.

At 0 the Gamma function is undefined; by convention we assign it the value 0.

At -n for n ∈ ℕ, the Gamma function is undefined; by convention we assign it the value 0.

Now check that the Γ function is differentiable, wherever this makes sense.

Rewrite the Gamma integral as an example of a Mellin transform.

The derivative of the Γ integral, at any s ∈ ℂ with 1 < re s, is given by the Melllin transform of log t * exp (-t).

theorem complex.differentiable_at_Gamma_aux (s : ) (n : ) (h1 : 1 - s.re < n) (h2 : (m : ), s -m) :

At s = 0, the Gamma function has a simple pole with residue 1.

noncomputable def real.Gamma (s : ) :

The Γ function (of a real variable s).

Equations
theorem real.Gamma_eq_integral {s : } (hs : 0 < s) :
real.Gamma s = (x : ) in set.Ioi 0, rexp (-x) * x ^ (s - 1)
theorem real.Gamma_add_one {s : } (hs : s 0) :
theorem real.Gamma_one  :
theorem real.Gamma_zero  :

At 0 the Gamma function is undefined; by convention we assign it the value 0.

At -n for n ∈ ℕ, the Gamma function is undefined; by convention we assign it the value 0.

theorem real.Gamma_pos_of_pos {s : } (hs : 0 < s) :
theorem real.Gamma_ne_zero {s : } (hs : (m : ), s -m) :

The Gamma function does not vanish on (except at non-positive integers, where the function is mathematically undefined and we set it to 0 by convention).

theorem real.Gamma_eq_zero_iff (s : ) :
real.Gamma s = 0 (m : ), s = -m