Documentation

Mathlib.MeasureTheory.Group.MeasurableEquiv

(Scalar) multiplication and (vector) addition as measurable equivalences #

In this file we define the following measurable equivalences:

We also deduce that the corresponding maps are measurable embeddings.

Tags #

measurable, equivalence, group action

def MeasurableEquiv.smul {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [Group G] [MulAction G α] [MeasurableConstSMul G α] (c : G) :
α ≃ᵐ α

If a group G acts on α by measurable maps, then each element c : G defines a measurable automorphism of α.

Equations
    Instances For
      def MeasurableEquiv.vadd {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [AddGroup G] [AddAction G α] [MeasurableConstVAdd G α] (c : G) :
      α ≃ᵐ α

      If an additive group G acts on α by measurable maps, then each element c : G defines a measurable automorphism of α.

      Equations
        Instances For
          @[simp]
          theorem MeasurableEquiv.smul_toEquiv {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [Group G] [MulAction G α] [MeasurableConstSMul G α] (c : G) :
          @[simp]
          theorem MeasurableEquiv.smul_apply {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [Group G] [MulAction G α] [MeasurableConstSMul G α] (c : G) :
          (smul c) = fun (x : α) => c x
          @[simp]
          theorem MeasurableEquiv.vadd_apply {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [AddGroup G] [AddAction G α] [MeasurableConstVAdd G α] (c : G) :
          (vadd c) = fun (x : α) => c +ᵥ x
          @[simp]
          theorem MeasurableEquiv.vadd_toEquiv {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [AddGroup G] [AddAction G α] [MeasurableConstVAdd G α] (c : G) :
          theorem measurableEmbedding_const_smul {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [Group G] [MulAction G α] [MeasurableConstSMul G α] (c : G) :
          MeasurableEmbedding fun (x : α) => c x
          theorem measurableEmbedding_const_vadd {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [AddGroup G] [AddAction G α] [MeasurableConstVAdd G α] (c : G) :
          MeasurableEmbedding fun (x : α) => c +ᵥ x
          @[simp]
          theorem MeasurableEquiv.symm_smul {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [Group G] [MulAction G α] [MeasurableConstSMul G α] (c : G) :
          @[simp]
          theorem MeasurableEquiv.symm_vadd {G : Type u_1} {α : Type u_3} [MeasurableSpace α] [AddGroup G] [AddAction G α] [MeasurableConstVAdd G α] (c : G) :
          (vadd c).symm = vadd (-c)
          def MeasurableEquiv.smul₀ {G₀ : Type u_2} {α : Type u_3} [MeasurableSpace α] [GroupWithZero G₀] [MulAction G₀ α] [MeasurableConstSMul G₀ α] (c : G₀) (hc : c 0) :
          α ≃ᵐ α

          If a group with zero G₀ acts on α by measurable maps, then each nonzero element c : G₀ defines a measurable automorphism of α

          Equations
            Instances For
              @[simp]
              theorem MeasurableEquiv.coe_smul₀ {G₀ : Type u_2} {α : Type u_3} [MeasurableSpace α] [GroupWithZero G₀] [MulAction G₀ α] [MeasurableConstSMul G₀ α] {c : G₀} (hc : c 0) :
              (smul₀ c hc) = fun (x : α) => c x
              @[simp]
              theorem MeasurableEquiv.symm_smul₀ {G₀ : Type u_2} {α : Type u_3} [MeasurableSpace α] [GroupWithZero G₀] [MulAction G₀ α] [MeasurableConstSMul G₀ α] {c : G₀} (hc : c 0) :
              (smul₀ c hc).symm = smul₀ c⁻¹
              theorem measurableEmbedding_const_smul₀ {G₀ : Type u_2} {α : Type u_3} [MeasurableSpace α] [GroupWithZero G₀] [MulAction G₀ α] [MeasurableConstSMul G₀ α] {c : G₀} (hc : c 0) :
              MeasurableEmbedding fun (x : α) => c x
              def MeasurableEquiv.mulLeft {G : Type u_1} [Group G] [MeasurableSpace G] [MeasurableMul G] (g : G) :
              G ≃ᵐ G

              If G is a group with measurable multiplication, then left multiplication by g : G is a measurable automorphism of G.

              Equations
                Instances For

                  If G is an additive group with measurable addition, then addition of g : G on the left is a measurable automorphism of G.

                  Equations
                    Instances For
                      @[simp]
                      theorem MeasurableEquiv.coe_mulLeft {G : Type u_1} [Group G] [MeasurableSpace G] [MeasurableMul G] (g : G) :
                      (mulLeft g) = fun (x : G) => g * x
                      @[simp]
                      theorem MeasurableEquiv.coe_addLeft {G : Type u_1} [AddGroup G] [MeasurableSpace G] [MeasurableAdd G] (g : G) :
                      (addLeft g) = fun (x : G) => g + x
                      @[simp]
                      theorem measurableEmbedding_mulLeft {G : Type u_1} [Group G] [MeasurableSpace G] [MeasurableMul G] (g : G) :
                      MeasurableEmbedding fun (x : G) => g * x
                      theorem measurableEmbedding_addLeft {G : Type u_1} [AddGroup G] [MeasurableSpace G] [MeasurableAdd G] (g : G) :
                      MeasurableEmbedding fun (x : G) => g + x

                      If G is a group with measurable multiplication, then right multiplication by g : G is a measurable automorphism of G.

                      Equations
                        Instances For

                          If G is an additive group with measurable addition, then addition of g : G on the right is a measurable automorphism of G.

                          Equations
                            Instances For
                              theorem measurableEmbedding_mulRight {G : Type u_1} [Group G] [MeasurableSpace G] [MeasurableMul G] (g : G) :
                              MeasurableEmbedding fun (x : G) => x * g
                              theorem measurableEmbedding_addRight {G : Type u_1} [AddGroup G] [MeasurableSpace G] [MeasurableAdd G] (g : G) :
                              MeasurableEmbedding fun (x : G) => x + g
                              @[simp]
                              theorem MeasurableEquiv.coe_mulRight {G : Type u_1} [Group G] [MeasurableSpace G] [MeasurableMul G] (g : G) :
                              (mulRight g) = fun (x : G) => x * g
                              @[simp]
                              theorem MeasurableEquiv.coe_addRight {G : Type u_1} [AddGroup G] [MeasurableSpace G] [MeasurableAdd G] (g : G) :
                              (addRight g) = fun (x : G) => x + g
                              def MeasurableEquiv.mulLeft₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] (g : G₀) (hg : g 0) :
                              G₀ ≃ᵐ G₀

                              If G₀ is a group with zero with measurable multiplication, then left multiplication by a nonzero element g : G₀ is a measurable automorphism of G₀.

                              Equations
                                Instances For
                                  theorem measurableEmbedding_mulLeft₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :
                                  MeasurableEmbedding fun (x : G₀) => g * x
                                  @[simp]
                                  theorem MeasurableEquiv.coe_mulLeft₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :
                                  (mulLeft₀ g hg) = fun (x : G₀) => g * x
                                  @[simp]
                                  theorem MeasurableEquiv.symm_mulLeft₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :
                                  @[simp]
                                  theorem MeasurableEquiv.toEquiv_mulLeft₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :
                                  def MeasurableEquiv.mulRight₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] (g : G₀) (hg : g 0) :
                                  G₀ ≃ᵐ G₀

                                  If G₀ is a group with zero with measurable multiplication, then right multiplication by a nonzero element g : G₀ is a measurable automorphism of G₀.

                                  Equations
                                    Instances For
                                      theorem measurableEmbedding_mulRight₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :
                                      MeasurableEmbedding fun (x : G₀) => x * g
                                      @[simp]
                                      theorem MeasurableEquiv.coe_mulRight₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :
                                      (mulRight₀ g hg) = fun (x : G₀) => x * g
                                      @[simp]
                                      theorem MeasurableEquiv.symm_mulRight₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :
                                      @[simp]
                                      theorem MeasurableEquiv.toEquiv_mulRight₀ {G₀ : Type u_2} [GroupWithZero G₀] [MeasurableSpace G₀] [MeasurableMul G₀] {g : G₀} (hg : g 0) :

                                      Inversion as a measurable automorphism of a group or group with zero.

                                      Equations
                                        Instances For

                                          Negation as a measurable automorphism of an additive group.

                                          Equations
                                            Instances For

                                              equiv.divRight as a MeasurableEquiv.

                                              Equations
                                                Instances For

                                                  equiv.subRight as a MeasurableEquiv

                                                  Equations
                                                    Instances For
                                                      theorem measurableEmbedding_divRight {G : Type u_1} [Group G] [MeasurableSpace G] [MeasurableMul G] (g : G) :
                                                      MeasurableEmbedding fun (x : G) => x / g
                                                      theorem measurableEmbedding_subRight {G : Type u_1} [AddGroup G] [MeasurableSpace G] [MeasurableAdd G] (g : G) :
                                                      MeasurableEmbedding fun (x : G) => x - g

                                                      equiv.divLeft as a MeasurableEquiv

                                                      Equations
                                                        Instances For

                                                          equiv.subLeft as a MeasurableEquiv

                                                          Equations
                                                            Instances For
                                                              theorem measurableEmbedding_divLeft {G : Type u_1} [Group G] [MeasurableSpace G] [MeasurableMul G] [MeasurableInv G] (g : G) :
                                                              MeasurableEmbedding fun (x : G) => g / x
                                                              theorem MeasureTheory.Measure.dmaSMul_apply {G : Type u_1} {A : Type u_2} [Group G] [MulAction G A] [MeasurableSpace A] [MeasurableConstSMul G A] (μ : Measure A) (g : Gᵈᵐᵃ) (s : Set A) :
                                                              (g μ) s = μ (DomMulAct.mk.symm g s)