Bochner integrals of convolutions #
This file contains results about the Bochner integrals of convolutions of measures.
These results are not placed in the main convolution file because we don't want to import Bochner integrals over there.
Main statements #
integrable_conv_iff
: A function is integrable with respect to the convolutionμ ∗ ν
iff the functiony ↦ f (x + y)
is integrable with respect toν
forμ
-almost everyx
and the functionx ↦ ∫ y, ‖f (x + y)‖ ∂ν
is integrable with respect toμ
.integral_conv
: iff
is integrable with respect to the convolutionμ ∗ ν
, then∫ x, f x ∂(μ ∗ ν) = ∫ x, ∫ y, f (x + y) ∂ν ∂μ
.
theorem
MeasureTheory.integrable_mconv_iff
{M : Type u_1}
{F : Type u_2}
[Monoid M]
{mM : MeasurableSpace M}
[MeasurableMul₂ M]
[NormedAddCommGroup F]
{μ ν : Measure M}
{f : M → F}
[SFinite ν]
(hf : AEStronglyMeasurable f (μ.mconv ν))
:
theorem
MeasureTheory.integrable_conv_iff
{M : Type u_1}
{F : Type u_2}
[AddMonoid M]
{mM : MeasurableSpace M}
[MeasurableAdd₂ M]
[NormedAddCommGroup F]
{μ ν : Measure M}
{f : M → F}
[SFinite ν]
(hf : AEStronglyMeasurable f (μ.conv ν))
:
theorem
MeasureTheory.integral_mconv
{M : Type u_1}
{F : Type u_2}
[Monoid M]
{mM : MeasurableSpace M}
[MeasurableMul₂ M]
[NormedAddCommGroup F]
{μ ν : Measure M}
{f : M → F}
[NormedSpace ℝ F]
[SFinite μ]
[SFinite ν]
(hf : Integrable f (μ.mconv ν))
:
theorem
MeasureTheory.integral_conv
{M : Type u_1}
{F : Type u_2}
[AddMonoid M]
{mM : MeasurableSpace M}
[MeasurableAdd₂ M]
[NormedAddCommGroup F]
{μ ν : Measure M}
{f : M → F}
[NormedSpace ℝ F]
[SFinite μ]
[SFinite ν]
(hf : Integrable f (μ.conv ν))
: