Documentation

Mathlib.Analysis.SpecialFunctions.MulExpNegMulSqIntegral

Properties of the integral of mulExpNegMulSq #

The mapping mulExpNegMulSq can be used to transform a function g : E → ℝ into a bounded function mulExpNegMulSq ε ∘ g : E → ℝ = fun x => g x * Real.exp (-ε * g x * g x). This file contains results on the integral of mulExpNegMulSq g ε with respect to a finite measure P.

Lemmas #

Main Result #

dist_integral_mulExpNegMulSq_comp_le: For a subalgebra of functions A, if for any g ∈ A the integral with respect to two finite measures P, P' coincide, then the difference of the integrals of mulExpNegMulSq ε ∘ g with respect to P, P' is bounded by 6 * sqrt ε. This will be a key ingredient in the proof of theorem ext_of_forall_mem_subalgebra_integral_eq (future work), where it is shown that a subalgebra of functions that separates points separates finite measures.

The integral of mulExpNegMulSq ε ∘ g with respect to a finite measure P converges to the integral of g, as ε → 0 from above.

theorem tendsto_integral_mul_one_plus_inv_smul_sq_pow {E : Type u_1} [TopologicalSpace E] [MeasurableSpace E] [BorelSpace E] {P : MeasureTheory.Measure E} [MeasureTheory.IsFiniteMeasure P] {ε : } (g : BoundedContinuousFunction E ) (hε : 0 < ε) :
Filter.Tendsto (fun (n : ) => (x : E), (g * (1 + (↑n)⁻¹ -(ε g * g)) ^ n) x P) Filter.atTop (nhds ( (x : E), ε.mulExpNegMulSq (g x) P))

The integral of mulExpNegMulSq ε ∘ g with respect to a finite measure P can be approximated by the integral of the sequence approximating the exponential function.

theorem integral_mulExpNegMulSq_comp_eq {E : Type u_1} [TopologicalSpace E] [MeasurableSpace E] [BorelSpace E] {P : MeasureTheory.Measure E} [MeasureTheory.IsFiniteMeasure P] {ε : } {P' : MeasureTheory.Measure E} [MeasureTheory.IsFiniteMeasure P'] {A : Subalgebra C(E, )} (hε : 0 < ε) (hbound : gA, ∃ (C : ), ∀ (x y : E), dist (g x) (g y) C) (heq : gA, (x : E), g x P = (x : E), g x P') {g : C(E, )} (hgA : g A) :
(x : E), ε.mulExpNegMulSq (g x) P = (x : E), ε.mulExpNegMulSq (g x) P'
theorem abs_integral_sub_setIntegral_mulExpNegMulSq_comp_lt {E : Type u_1} [TopologicalSpace E] [MeasurableSpace E] [BorelSpace E] {P : MeasureTheory.Measure E} [MeasureTheory.IsFiniteMeasure P] {ε : } (f : C(E, )) {K : Set E} (hK : MeasurableSet K) (hε : 0 < ε) (hKP : P K < ε.toNNReal) :
| (x : E), ε.mulExpNegMulSq (f x) P - (x : E) in K, ε.mulExpNegMulSq (f x) P| < ε
theorem abs_setIntegral_mulExpNegMulSq_comp_sub_le_mul_measure {E : Type u_1} [TopologicalSpace E] [MeasurableSpace E] [BorelSpace E] {P : MeasureTheory.Measure E} [MeasureTheory.IsFiniteMeasure P] {ε : } {K : Set E} (hK : IsCompact K) (hKmeas : MeasurableSet K) (f g : C(E, )) {δ : } (hε : 0 < ε) (hfg : xK, |g x - f x| < δ) :
| (x : E) in K, ε.mulExpNegMulSq (g x) P - (x : E) in K, ε.mulExpNegMulSq (f x) P| δ * (P K).toReal
theorem dist_integral_mulExpNegMulSq_comp_le {ε : } {E : Type u_2} [MeasurableSpace E] [PseudoEMetricSpace E] [BorelSpace E] [CompleteSpace E] [SecondCountableTopology E] {P P' : MeasureTheory.Measure E} [MeasureTheory.IsFiniteMeasure P] [MeasureTheory.IsFiniteMeasure P'] (f : BoundedContinuousFunction E ) {A : Subalgebra C(E, )} (hA : A.SeparatesPoints) (hbound : gA, ∃ (C : ), ∀ (x y : E), dist (g x) (g y) C) (heq : gA, (x : E), g x P = (x : E), g x P') (hε : 0 < ε) :
| (x : E), ε.mulExpNegMulSq (f x) P - (x : E), ε.mulExpNegMulSq (f x) P'| 6 * ε

If for any g ∈ A the integrals with respect to two finite measures P, P' coincide, then the difference of the integrals of mulExpNegMulSq ε ∘ g with respect to P, P' is bounded by 6 * sqrt ε.