Documentation

Mathlib.Algebra.GroupWithZero.Action.Units

Multiplicative actions with zero on and by #

This file provides the multiplicative actions with zero of a unit on a type α, SMul Mˣ α, in the presence of SMulWithZero M α, with the obvious definition stated in Units.smul_def.

Additionally, a MulDistribMulAction G M for some group G satisfying some additional properties admits a MulDistribMulAction G Mˣ structure, again with the obvious definition stated in Units.coe_smul. This instance uses a primed name.

See also #

Action of the units of M on a type α #

instance AddUnits.instVAdd_1 {M : Type u_2} {α : Type u_3} [AddMonoid M] [VAdd M α] :
VAdd (AddUnits M) α
Equations
  • AddUnits.instVAdd_1 = { vadd := fun (m : AddUnits M) (a : α) => m +ᵥ a }
instance Units.instSMul_1 {M : Type u_2} {α : Type u_3} [Monoid M] [SMul M α] :
SMul Mˣ α
Equations
  • Units.instSMul_1 = { smul := fun (m : Mˣ) (a : α) => m a }
instance Units.instSMulZeroClass {M : Type u_2} {α : Type u_3} [Monoid M] [Zero α] [SMulZeroClass M α] :
Equations
instance Units.instDistribSMulUnits {M : Type u_2} {α : Type u_3} [Monoid M] [AddZeroClass α] [DistribSMul M α] :
Equations
instance Units.instDistribMulAction {M : Type u_2} {α : Type u_3} [Monoid M] [AddMonoid α] [DistribMulAction M α] :
Equations
Equations

Action of a group G on units of M #

A stronger form of Units.mul_action'.

Equations
@[simp]
theorem IsUnit.smul_eq_zero {G : Type u_1} {M : Type u_2} [Monoid G] [AddMonoid M] [DistribMulAction G M] {u : G} {x : M} (hu : IsUnit u) :
u x = 0 x = 0