theorem
MeasureTheory.setIntegral_integral_swap
{α : Type u_1}
{β : Type u_2}
{E : Type u_3}
[MeasurableSpace α]
[MeasurableSpace β]
{μ : Measure α}
{ν : Measure β}
[NormedAddCommGroup E]
[SigmaFinite ν]
[NormedSpace ℝ E]
[SigmaFinite μ]
(f : α → β → E)
{s : Set α}
{t : Set β}
(hf : IntegrableOn (Function.uncurry f) (s ×ˢ t) (μ.prod ν))
:
theorem
IntervalIntegral.integral_eq_integral_of_support_subset_Icc
{a b : ℝ}
{μ : MeasureTheory.Measure ℝ}
[MeasureTheory.NoAtoms μ]
{E : Type u_2}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
[CompleteSpace E]
{f : ℝ → E}
(h : Function.support f ⊆ Set.Icc a b)
:
theorem
SetIntegral.integral_eq_integral_inter_of_support_subset
{μ : MeasureTheory.Measure ℝ}
{E : Type u_2}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{s t : Set ℝ}
{f : ℝ → E}
(h : Function.support f ⊆ t)
(ht : MeasurableSet t)
:
theorem
SetIntegral.integral_eq_integral_inter_of_support_subset_Icc
{a b : ℝ}
{μ : MeasureTheory.Measure ℝ}
{E : Type u_2}
[NormedAddCommGroup E]
[NormedSpace ℝ E]
{s : Set ℝ}
{f : ℝ → E}
(h : Function.support f ⊆ Set.Icc a b)
(hs : Set.Icc a b ⊆ s)
:
theorem
Filter.TendstoAtZero_of_support_in_Icc
{𝕂 : Type u_1}
[RCLike 𝕂]
{a b : ℝ}
(f : ℝ → 𝕂)
(ha : 0 < a)
(fSupp : Function.support f ⊆ Set.Icc a b)
:
Tendsto f (nhdsWithin 0 (Set.Ioi 0)) (nhds 0)
theorem
Filter.BigO_zero_atZero_of_support_in_Icc
{𝕂 : Type u_1}
[RCLike 𝕂]
{a b : ℝ}
(f : ℝ → 𝕂)
(ha : 0 < a)
(fSupp : Function.support f ⊆ Set.Icc a b)
:
theorem
PartialIntegration
(f g : ℝ → ℂ)
(fDiff : DifferentiableOn ℝ f (Set.Ioi 0))
(gDiff : DifferentiableOn ℝ g (Set.Ioi 0))
(fDerivgInt : MeasureTheory.IntegrableOn (f * deriv g) (Set.Ioi 0) MeasureTheory.volume)
(gDerivfInt : MeasureTheory.IntegrableOn (deriv f * g) (Set.Ioi 0) MeasureTheory.volume)
(lim_at_zero : Filter.Tendsto (f * g) (nhdsWithin 0 (Set.Ioi 0)) (nhds 0))
(lim_at_inf : Filter.Tendsto (f * g) Filter.atTop (nhds 0))
:
Need differentiability, and decay at 0
and ∞
theorem
PartialIntegration_of_support_in_Icc
{a b : ℝ}
(f g : ℝ → ℂ)
(ha : 0 < a)
(h : a ≤ b)
(fSupp : Function.support f ⊆ Set.Icc a b)
(fDiff : DifferentiableOn ℝ f (Set.Ioi 0))
(gDiff : DifferentiableOn ℝ g (Set.Ioi 0))
(fderivCont : ContinuousOn (deriv f) (Set.Ioi 0))
(gderivCont : ContinuousOn (deriv g) (Set.Ioi 0))
:
Equations
- mellintransform = Lean.ParserDescr.node `mellintransform 1024 (Lean.ParserDescr.symbol "𝓜")
Instances For
Equations
- MellinInverseTransform F σ x = VerticalIntegral' (fun (s : ℂ) => ↑x ^ (-s) * F s) σ
Instances For
theorem
MellinInversion
(σ : ℝ)
{f : ℝ → ℂ}
{x : ℝ}
(hx : 0 < x)
(hf : MellinConvergent f ↑σ)
(hFf : Complex.VerticalIntegrable (mellin f) σ MeasureTheory.volume)
(hfx : ContinuousAt f x)
:
theorem
support_MellinConvolution_subsets
{𝕂 : Type u_1}
[RCLike 𝕂]
{f g : ℝ → 𝕂}
{A B : Set ℝ}
(hf : Function.support f ⊆ A)
(hg : Function.support g ⊆ B)
:
theorem
MellinConvolutionTransform
(f g : ℝ → ℂ)
(s : ℂ)
(hf :
MeasureTheory.IntegrableOn (Function.uncurry fun (x y : ℝ) => f y * g (x / y) / ↑y * ↑x ^ (s - 1))
(Set.Ioi 0 ×ˢ Set.Ioi 0) MeasureTheory.volume)
:
theorem
DeltaSpikeContinuous
{ν : ℝ → ℝ}
{ε : ℝ}
(εpos : 0 < ε)
(diffν : ContDiff ℝ 1 ν)
:
Continuous fun (x : ℝ) => DeltaSpike ν ε x
theorem
DeltaSpikeOfRealContinuous
{ν : ℝ → ℝ}
{ε : ℝ}
(εpos : 0 < ε)
(diffν : ContDiff ℝ 1 ν)
:
Continuous fun (x : ℝ) => ↑(DeltaSpike ν ε x)
theorem
MellinOfDeltaSpike
(ν : ℝ → ℝ)
{ε : ℝ}
(εpos : ε > 0)
(s : ℂ)
:
MellinTransform (fun (x : ℝ) => ↑(DeltaSpike ν ε x)) s = MellinTransform (fun (x : ℝ) => ↑(ν x)) (↑ε * s)
theorem
Smooth1ContinuousAt
{SmoothingF : ℝ → ℝ}
(diffSmoothingF : ContDiff ℝ 1 SmoothingF)
(SmoothingFpos : ∀ x > 0, 0 ≤ SmoothingF x)
(suppSmoothingF : Function.support SmoothingF ⊆ Set.Icc (1 / 2) 2)
{ε : ℝ}
(εpos : 0 < ε)
{y : ℝ}
(ypos : 0 < y)
:
ContinuousAt (fun (x : ℝ) => Smooth1 SmoothingF ε x) y
theorem
Smooth1MellinConvergent
{Ψ : ℝ → ℝ}
{ε : ℝ}
(diffΨ : ContDiff ℝ 1 Ψ)
(suppΨ : Function.support Ψ ⊆ Set.Icc (1 / 2) 2)
(hε : ε ∈ Set.Ioo 0 1)
(Ψnonneg : ∀ x > 0, 0 ≤ Ψ x)
(mass_one : ∫ (x : ℝ) in Set.Ioi 0, Ψ x / x = 1)
{s : ℂ}
(hs : 0 < s.re)
:
MellinConvergent (fun (x : ℝ) => ↑(Smooth1 Ψ ε x)) s
theorem
Smooth1MellinDifferentiable
{Ψ : ℝ → ℝ}
{ε : ℝ}
(diffΨ : ContDiff ℝ 1 Ψ)
(suppΨ : Function.support Ψ ⊆ Set.Icc (1 / 2) 2)
(hε : ε ∈ Set.Ioo 0 1)
(Ψnonneg : ∀ x > 0, 0 ≤ Ψ x)
(mass_one : ∫ (x : ℝ) in Set.Ioi 0, Ψ x / x = 1)
{s : ℂ}
(hs : 0 < s.re)
:
DifferentiableAt ℂ (MellinTransform fun (x : ℝ) => ↑(Smooth1 Ψ ε x)) s