Documentation

Mathlib.Algebra.Group.Action.Opposite

Scalar actions on and by Mᵐᵒᵖ #

This file defines the actions on the opposite type SMul R Mᵐᵒᵖ, and actions by the opposite type, SMul Rᵐᵒᵖ M.

Note that MulOpposite.smul is provided in an earlier file as it is needed to provide the AddMonoid.nsmul and AddCommGroup.zsmul fields.

Notation #

With open scoped RightActions, this provides:

Actions on the opposite type #

Actions on the opposite type just act on the underlying type.

instance MulOpposite.instMulAction {M : Type u_1} {α : Type u_3} [Monoid M] [MulAction M α] :
Equations
instance AddOpposite.instAddAction {M : Type u_1} {α : Type u_3} [AddMonoid M] [AddAction M α] :
Equations
theorem MulOpposite.instIsScalarTower {M : Type u_1} {N : Type u_2} {α : Type u_3} [SMul M N] [SMul M α] [SMul N α] [IsScalarTower M N α] :
theorem AddOpposite.instIsScalarTower {M : Type u_1} {N : Type u_2} {α : Type u_3} [VAdd M N] [VAdd M α] [VAdd N α] [VAddAssocClass M N α] :
theorem MulOpposite.instSMulCommClass {M : Type u_1} {N : Type u_2} {α : Type u_3} [SMul M α] [SMul N α] [SMulCommClass M N α] :
theorem AddOpposite.instVAddCommClass {M : Type u_1} {N : Type u_2} {α : Type u_3} [VAdd M α] [VAdd N α] [VAddCommClass M N α] :
theorem MulOpposite.op_smul_eq_op_smul_op {M : Type u_1} {α : Type u_3} [SMul M α] [SMul Mᵐᵒᵖ α] [IsCentralScalar M α] (r : M) (a : α) :
theorem AddOpposite.op_vadd_eq_op_vadd_op {M : Type u_1} {α : Type u_3} [VAdd M α] [VAdd Mᵃᵒᵖ α] [IsCentralVAdd M α] (r : M) (a : α) :

Right actions #

In this section we establish SMul αᵐᵒᵖ β as the canonical spelling of right scalar multiplication of β by α, and provide convenient notations.

Pretty printer defined by notation3 command.

Equations
  • One or more equations did not get rendered due to their size.
Instances For

    With open scoped RightActions, an alternative symbol for left actions, r • m.

    In lemma names this is still called smul.

    Equations
    Instances For

      Pretty printer defined by notation3 command.

      Equations
      • One or more equations did not get rendered due to their size.
      Instances For

        With open scoped RightActions, a shorthand for right actions, op r • m.

        In lemma names this is still called op_smul.

        Equations
        Instances For

          With open scoped RightActions, an alternative symbol for left actions, r • m.

          In lemma names this is still called vadd.

          Equations
          Instances For

            Pretty printer defined by notation3 command.

            Equations
            • One or more equations did not get rendered due to their size.
            Instances For

              With open scoped RightActions, a shorthand for right actions, op r +ᵥ m.

              In lemma names this is still called op_vadd.

              Equations
              Instances For

                Pretty printer defined by notation3 command.

                Equations
                • One or more equations did not get rendered due to their size.
                Instances For
                  theorem op_smul_op_smul {α : Type u_3} {β : Type u_4} [Monoid α] [MulAction αᵐᵒᵖ β] (b : β) (a₁ a₂ : α) :
                  theorem op_vadd_op_vadd {α : Type u_3} {β : Type u_4} [AddMonoid α] [AddAction αᵃᵒᵖ β] (b : β) (a₁ a₂ : α) :
                  theorem op_smul_mul {α : Type u_3} {β : Type u_4} [Monoid α] [MulAction αᵐᵒᵖ β] (b : β) (a₁ a₂ : α) :
                  theorem op_vadd_add {α : Type u_3} {β : Type u_4} [AddMonoid α] [AddAction αᵃᵒᵖ β] (b : β) (a₁ a₂ : α) :

                  Actions by the opposite type (right actions) #

                  In Mul.toSMul in another file, we define the left action a₁ • a₂ = a₁ * a₂. For the multiplicative opposite, we define MulOpposite.op a₁ • a₂ = a₂ * a₁, with the multiplication reversed.

                  instance Mul.toHasOppositeSMul {α : Type u_3} [Mul α] :

                  Like Mul.toSMul, but multiplies on the right.

                  See also Monoid.toOppositeMulAction and MonoidWithZero.toOppositeMulActionWithZero.

                  Equations
                  instance Add.toHasOppositeVAdd {α : Type u_3} [Add α] :

                  Like Add.toVAdd, but adds on the right.

                  See also AddMonoid.toOppositeAddAction.

                  Equations
                  theorem op_smul_eq_mul {α : Type u_3} [Mul α] {a a' : α} :
                  MulOpposite.op a a' = a' * a
                  theorem op_vadd_eq_add {α : Type u_3} [Add α] {a a' : α} :
                  @[simp]
                  theorem MulOpposite.smul_eq_mul_unop {α : Type u_3} [Mul α] {a : αᵐᵒᵖ} {a' : α} :
                  @[simp]
                  theorem AddOpposite.vadd_eq_add_unop {α : Type u_3} [Add α] {a : αᵃᵒᵖ} {a' : α} :

                  The right regular action of a group on itself is transitive.

                  The right regular action of an additive group on itself is transitive.

                  Like Monoid.toMulAction, but multiplies on the right.

                  Equations

                  Like AddMonoid.toAddAction, but adds on the right.

                  Equations
                  theorem IsScalarTower.opposite_mid {M : Type u_5} {N : Type u_6} [Mul N] [SMul M N] [SMulCommClass M N N] :
                  theorem VAddAssocClass.opposite_mid {M : Type u_5} {N : Type u_6} [Add N] [VAdd M N] [VAddCommClass M N N] :
                  theorem SMulCommClass.opposite_mid {M : Type u_5} {N : Type u_6} [Mul N] [SMul M N] [IsScalarTower M N N] :
                  theorem VAddCommClass.opposite_mid {M : Type u_5} {N : Type u_6} [Add N] [VAdd M N] [VAddAssocClass M N N] :