Fourier transform of the Gaussian #
We prove that the Fourier transform of the Gaussian function is another Gaussian:
integral_cexp_quadratic
: general formula for∫ (x : ℝ), exp (b * x ^ 2 + c * x + d)
fourierIntegral_gaussian
: for all complexb
andt
with0 < re b
, we have∫ x:ℝ, exp (I * t * x) * exp (-b * x^2) = (π / b) ^ (1 / 2) * exp (-t ^ 2 / (4 * b))
.fourierIntegral_gaussian_pi
: a variant withb
andt
scaled to give a more symmetric statement, and formulated in terms of the Fourier transform operator𝓕
.
We also give versions of these formulas in finite-dimensional inner product spaces, see
integral_cexp_neg_mul_sq_norm_add
and fourierIntegral_gaussian_innerProductSpace
.
Fourier integral of Gaussian functions #
The integral of the Gaussian function over the vertical edges of a rectangle
with vertices at (±T, 0)
and (±T, c)
.
Equations
Instances For
theorem
GaussianFourier.tendsto_verticalIntegral
{b : ℂ}
(hb : 0 < b.re)
(c : ℝ)
:
Filter.Tendsto (GaussianFourier.verticalIntegral b c) Filter.atTop (nhds 0)
theorem
GaussianFourier.integrable_cexp_neg_mul_sq_add_real_mul_I
{b : ℂ}
(hb : 0 < b.re)
(c : ℝ)
:
MeasureTheory.Integrable (fun (x : ℝ) => Complex.exp (-b * (↑x + ↑c * Complex.I) ^ 2)) MeasureTheory.volume
theorem
integrable_cexp_quadratic'
{b : ℂ}
(hb : b.re < 0)
(c : ℂ)
(d : ℂ)
:
MeasureTheory.Integrable (fun (x : ℝ) => Complex.exp (b * ↑x ^ 2 + c * ↑x + d)) MeasureTheory.volume
theorem
integrable_cexp_quadratic
{b : ℂ}
(hb : 0 < b.re)
(c : ℂ)
(d : ℂ)
:
MeasureTheory.Integrable (fun (x : ℝ) => Complex.exp (-b * ↑x ^ 2 + c * ↑x + d)) MeasureTheory.volume
@[deprecated fourierIntegral_gaussian_pi']
Alias of fourierIntegral_gaussian_pi'
.
@[deprecated fourierIntegral_gaussian_pi]
Alias of fourierIntegral_gaussian_pi
.
theorem
GaussianFourier.integrable_cexp_neg_sum_mul_add
{ι : Type u_2}
[Fintype ι]
{b : ι → ℂ}
(hb : ∀ (i : ι), 0 < (b i).re)
(c : ι → ℂ)
:
MeasureTheory.Integrable (fun (v : ι → ℝ) => Complex.exp (-∑ i : ι, b i * ↑(v i) ^ 2 + ∑ i : ι, c i * ↑(v i)))
MeasureTheory.volume
theorem
GaussianFourier.integrable_cexp_neg_mul_sum_add
{b : ℂ}
{ι : Type u_2}
[Fintype ι]
(hb : 0 < b.re)
(c : ι → ℂ)
:
MeasureTheory.Integrable (fun (v : ι → ℝ) => Complex.exp (-b * ∑ i : ι, ↑(v i) ^ 2 + ∑ i : ι, c i * ↑(v i)))
MeasureTheory.volume
theorem
GaussianFourier.integrable_cexp_neg_mul_sq_norm_add_of_euclideanSpace
{b : ℂ}
{ι : Type u_2}
[Fintype ι]
(hb : 0 < b.re)
(c : ℂ)
(w : EuclideanSpace ℝ ι)
:
MeasureTheory.Integrable (fun (v : EuclideanSpace ℝ ι) => Complex.exp (-b * ↑‖v‖ ^ 2 + c * ↑(inner w v)))
MeasureTheory.volume
theorem
GaussianFourier.integrable_cexp_neg_mul_sq_norm_add
{b : ℂ}
{V : Type u_1}
[NormedAddCommGroup V]
[InnerProductSpace ℝ V]
[FiniteDimensional ℝ V]
[MeasurableSpace V]
[BorelSpace V]
(hb : 0 < b.re)
(c : ℂ)
(w : V)
:
MeasureTheory.Integrable (fun (v : V) => Complex.exp (-b * ↑‖v‖ ^ 2 + c * ↑(inner w v))) MeasureTheory.volume
In a real inner product space, the complex exponential of minus the square of the norm plus a scalar product is integrable. Useful when discussing the Fourier transform of a Gaussian.
theorem
GaussianFourier.integral_cexp_neg_mul_sq_norm_add_of_euclideanSpace
{b : ℂ}
{ι : Type u_2}
[Fintype ι]
(hb : 0 < b.re)
(c : ℂ)
(w : EuclideanSpace ℝ ι)
:
theorem
GaussianFourier.integral_cexp_neg_mul_sq_norm_add
{b : ℂ}
{V : Type u_1}
[NormedAddCommGroup V]
[InnerProductSpace ℝ V]
[FiniteDimensional ℝ V]
[MeasurableSpace V]
[BorelSpace V]
(hb : 0 < b.re)
(c : ℂ)
(w : V)
:
theorem
GaussianFourier.integral_cexp_neg_mul_sq_norm
{b : ℂ}
{V : Type u_1}
[NormedAddCommGroup V]
[InnerProductSpace ℝ V]
[FiniteDimensional ℝ V]
[MeasurableSpace V]
[BorelSpace V]
(hb : 0 < b.re)
:
theorem
GaussianFourier.integral_rexp_neg_mul_sq_norm
{V : Type u_1}
[NormedAddCommGroup V]
[InnerProductSpace ℝ V]
[FiniteDimensional ℝ V]
[MeasurableSpace V]
[BorelSpace V]
{b : ℝ}
(hb : 0 < b)
:
theorem
fourierIntegral_gaussian_innerProductSpace'
{b : ℂ}
{V : Type u_1}
[NormedAddCommGroup V]
[InnerProductSpace ℝ V]
[FiniteDimensional ℝ V]
[MeasurableSpace V]
[BorelSpace V]
(hb : 0 < b.re)
(x : V)
(w : V)
:
theorem
fourierIntegral_gaussian_innerProductSpace
{b : ℂ}
{V : Type u_1}
[NormedAddCommGroup V]
[InnerProductSpace ℝ V]
[FiniteDimensional ℝ V]
[MeasurableSpace V]
[BorelSpace V]
(hb : 0 < b.re)
(w : V)
: