Documentation

Mathlib.MeasureTheory.Group.Prod

Measure theory in the product of groups #

In this file we show properties about measure theory in products of measurable groups and properties of iterated integrals in measurable groups.

These lemmas show the uniqueness of left invariant measures on measurable groups, up to scaling. In this file we follow the proof and refer to the book Measure Theory by Paul Halmos.

The idea of the proof is to use the translation invariance of measures to prove μ(t) = c * μ(s) for two sets s and t, where c is a constant that does not depend on μ. Let e and f be the characteristic functions of s and t. Assume that μ and ν are left-invariant measures. Then the map (x, y) ↦ (y * x, x⁻¹) preserves the measure μ × ν, which means that

  ∫ x, ∫ y, h x y ∂ν ∂μ = ∫ x, ∫ y, h (y * x) x⁻¹ ∂ν ∂μ

If we apply this to h x y := e x * f y⁻¹ / ν ((fun h ↦ h * y⁻¹) ⁻¹' s), we can rewrite the RHS to μ(t), and the LHS to c * μ(s), where c = c(ν) does not depend on μ. Applying this to μ and to ν gives μ (t) / μ (s) = ν (t) / ν (s), which is the uniqueness up to scalar multiplication.

The proof in [Halmos] seems to contain an omission in §60 Th. A, see MeasureTheory.measure_lintegral_div_measure.

Note that this theory only applies in measurable groups, i.e., when multiplication and inversion are measurable. This is not the case in general in locally compact groups, or even in compact groups, when the topology is not second-countable. For arguments along the same line, but using continuous functions instead of measurable sets and working in the general locally compact setting, see the file MeasureTheory.Measure.Haar.Unique.lean.

The map (x, y) ↦ (x, x + y) as a MeasurableEquiv.

Equations
Instances For
    theorem MeasurableEquiv.shearAddRight.proof_2 (G : Type u_1) [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] [MeasurableNeg G] :
    Measurable fun (a : G × G) => (a.1, (Equiv.addLeft ((Equiv.refl G).symm a.1)).symm a.2)

    The map (x, y) ↦ (x, xy) as a MeasurableEquiv.

    Equations
    Instances For
      theorem MeasurableEquiv.shearSubRight.proof_2 (G : Type u_1) [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] :
      Measurable fun (a : G × G) => (a.1, (Equiv.subRight ((Equiv.refl G).symm a.1)).symm a.2)

      The map (x, y) ↦ (x, y - x) as a MeasurableEquiv with as inverse (x, y) ↦ (x, y + x).

      Equations
      Instances For

        The map (x, y) ↦ (x, y / x) as a MeasurableEquiv with as inverse (x, y) ↦ (x, yx)

        Equations
        Instances For

          The multiplicative shear mapping (x, y) ↦ (x, xy) preserves the measure μ × ν. This condition is part of the definition of a measurable group in [Halmos, §59]. There, the map in this lemma is called S.

          The map (x, y) ↦ (y, yx) sends the measure μ × ν to ν × μ. This is the map SR in [Halmos, §59]. S is the map (x, y) ↦ (x, xy) and R is Prod.swap.

          theorem MeasureTheory.measurable_measure_add_right {G : Type u_1} [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] (μ : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite μ] {s : Set G} (hs : MeasurableSet s) :
          Measurable fun (x : G) => μ ((fun (y : G) => y + x) ⁻¹' s)
          theorem MeasureTheory.measurable_measure_mul_right {G : Type u_1} [MeasurableSpace G] [Group G] [MeasurableMul₂ G] (μ : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite μ] {s : Set G} (hs : MeasurableSet s) :
          Measurable fun (x : G) => μ ((fun (y : G) => y * x) ⁻¹' s)

          The map (x, y) ↦ (x, x⁻¹y) is measure-preserving. This is the function S⁻¹ in [Halmos, §59], where S is the map (x, y) ↦ (x, xy).

          The map (x, y) ↦ (y, y⁻¹x) sends μ × ν to ν × μ. This is the function S⁻¹R in [Halmos, §59], where S is the map (x, y) ↦ (x, xy) and R is Prod.swap.

          The map (x, y) ↦ (yx, x⁻¹) is measure-preserving. This is the function S⁻¹RSR in [Halmos, §59], where S is the map (x, y) ↦ (x, xy) and R is Prod.swap.

          theorem MeasureTheory.lintegral_lintegral_add_neg {G : Type u_1} [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] [MeasurableNeg G] [MeasureTheory.Measure.IsAddLeftInvariant μ] [MeasureTheory.Measure.IsAddLeftInvariant ν] (f : GGENNReal) (hf : AEMeasurable (Function.uncurry f) (MeasureTheory.Measure.prod μ ν)) :
          ∫⁻ (x : G), ∫⁻ (y : G), f (y + x) (-x)νμ = ∫⁻ (x : G), ∫⁻ (y : G), f x yνμ
          theorem MeasureTheory.lintegral_lintegral_mul_inv {G : Type u_1} [MeasurableSpace G] [Group G] [MeasurableMul₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] [MeasurableInv G] [MeasureTheory.Measure.IsMulLeftInvariant μ] [MeasureTheory.Measure.IsMulLeftInvariant ν] (f : GGENNReal) (hf : AEMeasurable (Function.uncurry f) (MeasureTheory.Measure.prod μ ν)) :
          ∫⁻ (x : G), ∫⁻ (y : G), f (y * x) x⁻¹νμ = ∫⁻ (x : G), ∫⁻ (y : G), f x yνμ
          theorem MeasureTheory.measure_add_right_null {G : Type u_1} [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] (μ : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableNeg G] [MeasureTheory.Measure.IsAddLeftInvariant μ] (y : G) :
          μ ((fun (x : G) => x + y) ⁻¹' s) = 0 μ s = 0
          theorem MeasureTheory.measure_mul_right_null {G : Type u_1} [MeasurableSpace G] [Group G] [MeasurableMul₂ G] (μ : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableInv G] [MeasureTheory.Measure.IsMulLeftInvariant μ] (y : G) :
          μ ((fun (x : G) => x * y) ⁻¹' s) = 0 μ s = 0
          theorem MeasureTheory.measure_add_right_ne_zero {G : Type u_1} [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] (μ : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableNeg G] [MeasureTheory.Measure.IsAddLeftInvariant μ] (h2s : μ s 0) (y : G) :
          μ ((fun (x : G) => x + y) ⁻¹' s) 0
          theorem MeasureTheory.measure_mul_right_ne_zero {G : Type u_1} [MeasurableSpace G] [Group G] [MeasurableMul₂ G] (μ : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableInv G] [MeasureTheory.Measure.IsMulLeftInvariant μ] (h2s : μ s 0) (y : G) :
          μ ((fun (x : G) => x * y) ⁻¹' s) 0
          theorem MeasureTheory.measure_add_lintegral_eq {G : Type u_1} [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableNeg G] [MeasureTheory.Measure.IsAddLeftInvariant μ] [MeasureTheory.Measure.IsAddLeftInvariant ν] (sm : MeasurableSet s) (f : GENNReal) (hf : Measurable f) :
          μ s * ∫⁻ (y : G), f yν = ∫⁻ (x : G), ν ((fun (z : G) => z + x) ⁻¹' s) * f (-x)μ

          This is the computation performed in the proof of [Halmos, §60 Th. A].

          theorem MeasureTheory.measure_mul_lintegral_eq {G : Type u_1} [MeasurableSpace G] [Group G] [MeasurableMul₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableInv G] [MeasureTheory.Measure.IsMulLeftInvariant μ] [MeasureTheory.Measure.IsMulLeftInvariant ν] (sm : MeasurableSet s) (f : GENNReal) (hf : Measurable f) :
          μ s * ∫⁻ (y : G), f yν = ∫⁻ (x : G), ν ((fun (z : G) => z * x) ⁻¹' s) * f x⁻¹μ

          This is the computation performed in the proof of [Halmos, §60 Th. A].

          theorem MeasureTheory.measure_lintegral_sub_measure {G : Type u_1} [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableNeg G] [MeasureTheory.Measure.IsAddLeftInvariant μ] [MeasureTheory.Measure.IsAddLeftInvariant ν] (sm : MeasurableSet s) (h2s : ν s 0) (h3s : ν s ) (f : GENNReal) (hf : Measurable f) :
          μ s * ∫⁻ (y : G), f (-y) / ν ((fun (x : G) => x + -y) ⁻¹' s)ν = ∫⁻ (x : G), f xμ

          A technical lemma relating two different measures. This is basically [Halmos, §60 Th. A]. Note that if f is the characteristic function of a measurable set t this states that μ t = c * μ s for a constant c that does not depend on μ.

          Note: There is a gap in the last step of the proof in [Halmos]. In the last line, the equality g(-x) + ν(s - x) = f(x) holds if we can prove that 0 < ν(s - x) < ∞. The first inequality follows from §59, Th. D, but the second inequality is not justified. We prove this inequality for almost all x in MeasureTheory.ae_measure_preimage_add_right_lt_top_of_ne_zero.

          theorem MeasureTheory.measure_lintegral_div_measure {G : Type u_1} [MeasurableSpace G] [Group G] [MeasurableMul₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] {s : Set G} [MeasurableInv G] [MeasureTheory.Measure.IsMulLeftInvariant μ] [MeasureTheory.Measure.IsMulLeftInvariant ν] (sm : MeasurableSet s) (h2s : ν s 0) (h3s : ν s ) (f : GENNReal) (hf : Measurable f) :
          μ s * ∫⁻ (y : G), f y⁻¹ / ν ((fun (x : G) => x * y⁻¹) ⁻¹' s)ν = ∫⁻ (x : G), f xμ

          A technical lemma relating two different measures. This is basically [Halmos, §60 Th. A]. Note that if f is the characteristic function of a measurable set t this states that μ t = c * μ s for a constant c that does not depend on μ.

          Note: There is a gap in the last step of the proof in [Halmos]. In the last line, the equality g(x⁻¹)ν(sx⁻¹) = f(x) holds if we can prove that 0 < ν(sx⁻¹) < ∞. The first inequality follows from §59, Th. D, but the second inequality is not justified. We prove this inequality for almost all x in MeasureTheory.ae_measure_preimage_mul_right_lt_top_of_ne_zero.

          theorem MeasureTheory.measure_add_measure_eq {G : Type u_1} [MeasurableSpace G] [AddGroup G] [MeasurableAdd₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] [MeasurableNeg G] [MeasureTheory.Measure.IsAddLeftInvariant μ] [MeasureTheory.Measure.IsAddLeftInvariant ν] {s : Set G} {t : Set G} (hs : MeasurableSet s) (ht : MeasurableSet t) (h2s : ν s 0) (h3s : ν s ) :
          μ s * ν t = ν s * μ t
          theorem MeasureTheory.measure_mul_measure_eq {G : Type u_1} [MeasurableSpace G] [Group G] [MeasurableMul₂ G] (μ : MeasureTheory.Measure G) (ν : MeasureTheory.Measure G) [MeasureTheory.SigmaFinite ν] [MeasureTheory.SigmaFinite μ] [MeasurableInv G] [MeasureTheory.Measure.IsMulLeftInvariant μ] [MeasureTheory.Measure.IsMulLeftInvariant ν] {s : Set G} {t : Set G} (hs : MeasurableSet s) (ht : MeasurableSet t) (h2s : ν s 0) (h3s : ν s ) :
          μ s * ν t = ν s * μ t

          Left invariant Borel measures on an additive measurable group are unique (up to a scalar).

          Left invariant Borel measures on a measurable group are unique (up to a scalar).

          A left-invariant measure is quasi-preserved by right-addition. This should not be confused with (measurePreserving_add_right μ g).quasiMeasurePreserving.

          A left-invariant measure is quasi-preserved by right-multiplication. This should not be confused with (measurePreserving_mul_right μ g).quasiMeasurePreserving.

          A right-invariant measure is quasi-preserved by left-addition. This should not be confused with (measurePreserving_add_left μ g).quasiMeasurePreserving.

          A right-invariant measure is quasi-preserved by left-multiplication. This should not be confused with (measurePreserving_mul_left μ g).quasiMeasurePreserving.