Definitions of group actions #
This file defines a hierarchy of group action type-classes on top of the previously defined
notation classes SMul and its additive version VAdd:
- SMulZeroClassis a typeclass for an action that preserves zero
- DistribSMul M Ais a typeclass for an action on an additive monoid (- AddZeroClass) that preserves addition and zero
- DistribMulAction M Ais a typeclass for an action of a multiplicative monoid on an additive monoid such that- a • (b + c) = a • b + a • cand- a • 0 = 0.
The hierarchy is extended further by Module, defined elsewhere.
Notation #
- a • bis used as notation for- SMul.smul a b.
Implementation details #
This file should avoid depending on other parts of GroupTheory, to avoid import cycles.
More sophisticated lemmas belong in GroupTheory.GroupAction.
Tags #
group action
Pullback a zero-preserving scalar multiplication along an injective zero-preserving map. See note [reducible non-instances].
Equations
- Function.Injective.smulZeroClass f hf smul = { smul := fun (x1 : M) (x2 : B) => x1 • x2, smul_zero := ⋯ }
Instances For
Pushforward a zero-preserving scalar multiplication along a zero-preserving map. See note [reducible non-instances].
Equations
- f.smulZeroClass smul = { toSMul := inst✝, smul_zero := ⋯ }
Instances For
Push forward the multiplication of R on M along a compatible surjective map f : R → S.
See also Function.Surjective.distribMulActionLeft.
Equations
- Function.Surjective.smulZeroClassLeft f hf hsmul = { smul := fun (x1 : S) (x2 : M) => x1 • x2, smul_zero := ⋯ }
Instances For
Compose a SMulZeroClass with a function, with scalar multiplication f r' • m.
See note [reducible non-instances].
Equations
- SMulZeroClass.compFun A f = { smul := SMul.comp.smul f, smul_zero := ⋯ }
Instances For
Each element of the scalars defines a zero-preserving map.
Equations
- SMulZeroClass.toZeroHom A x = { toFun := fun (x_1 : A) => x • x_1, map_zero' := ⋯ }
Instances For
SMulWithZero is a class consisting of a Type M₀ with 0 ∈ M₀ and a scalar multiplication
of M₀ on a Type A with 0, such that the equality r • m = 0 holds if at least one among r
or m equals 0.
- smul : M₀ → A → A
- Scalar multiplication by the scalar - 0is- 0.
Instances
Equations
- MulZeroClass.toSMulWithZero M₀ = { smul := fun (x1 x2 : M₀) => x1 * x2, smul_zero := ⋯, zero_smul := ⋯ }
Like MulZeroClass.toSMulWithZero, but multiplies on the right.
Pullback a SMulWithZero structure along an injective zero-preserving homomorphism.
Equations
- Function.Injective.smulWithZero f hf smul = { smul := fun (x1 : M₀) (x2 : A') => x1 • x2, smul_zero := ⋯, zero_smul := ⋯ }
Instances For
Pushforward a SMulWithZero structure along a surjective zero-preserving homomorphism.
Equations
- Function.Surjective.smulWithZero f hf smul = { smul := fun (x1 : M₀) (x2 : A') => x1 • x2, smul_zero := ⋯, zero_smul := ⋯ }
Instances For
Compose a SMulWithZero with a ZeroHom, with action f r' • m
Equations
- SMulWithZero.compHom A f = { smul := fun (x1 : M₀') (x2 : A) => f x1 • x2, smul_zero := ⋯, zero_smul := ⋯ }
Instances For
Equations
- AddMonoid.natSMulWithZero = { toSMul := AddMonoid.toNatSMul, smul_zero := ⋯, zero_smul := ⋯ }
Equations
- AddGroup.intSMulWithZero = { toSMul := SubNegMonoid.toZSMul, smul_zero := ⋯, zero_smul := ⋯ }
An action of a monoid with zero M₀ on a Type A, also with 0, extends MulAction and
is compatible with 0 (both in M₀ and in A), with 1 ∈ M₀, and with associativity of
multiplication on the monoid A.
- smul : M₀ → A → A
- Scalar multiplication by any element send - 0to- 0.
- Scalar multiplication by the scalar - 0is- 0.
Instances
Equations
- MulActionWithZero.toSMulWithZero M₀ A = { toSMul := m.toSMul, smul_zero := ⋯, zero_smul := ⋯ }
See also Semiring.toModule
Equations
- MonoidWithZero.toMulActionWithZero M₀ = { toSMul := (MulZeroClass.toSMulWithZero M₀).toSMul, one_smul := ⋯, mul_smul := ⋯, smul_zero := ⋯, zero_smul := ⋯ }
Like MonoidWithZero.toMulActionWithZero, but multiplies on the right. See also
Semiring.toOppositeModule
Equations
- MonoidWithZero.toOppositeMulActionWithZero M₀ = { toSMul := (MulZeroClass.toOppositeSMulWithZero M₀).toSMul, one_smul := ⋯, mul_smul := ⋯, smul_zero := ⋯, zero_smul := ⋯ }
Pullback a MulActionWithZero structure along an injective zero-preserving homomorphism.
Equations
- Function.Injective.mulActionWithZero f hf smul = { toMulAction := Function.Injective.mulAction (⇑f) hf smul, smul_zero := ⋯, zero_smul := ⋯ }
Instances For
Pushforward a MulActionWithZero structure along a surjective zero-preserving homomorphism.
Equations
- Function.Surjective.mulActionWithZero f hf smul = { toMulAction := Function.Surjective.mulAction (⇑f) hf smul, smul_zero := ⋯, zero_smul := ⋯ }
Instances For
Compose a MulActionWithZero with a MonoidWithZeroHom, with action f r' • m
Equations
- MulActionWithZero.compHom A f = { toSMul := (SMulWithZero.compHom A ↑f).toSMul, one_smul := ⋯, mul_smul := ⋯, smul_zero := ⋯, zero_smul := ⋯ }
Instances For
Typeclass for scalar multiplication that preserves 0 and + on the right.
This is exactly DistribMulAction without the MulAction part.
- smul : M → A → A
- Scalar multiplication distributes across addition 
Instances
Pullback a distributive scalar multiplication along an injective additive monoid homomorphism. See note [reducible non-instances].
Equations
- Function.Injective.distribSMul f hf smul = { toSMulZeroClass := Function.Injective.smulZeroClass (↑f) hf smul, smul_add := ⋯ }
Instances For
Pushforward a distributive scalar multiplication along a surjective additive monoid homomorphism. See note [reducible non-instances].
Equations
- Function.Surjective.distribSMul f hf smul = { toSMulZeroClass := (↑f).smulZeroClass smul, smul_add := ⋯ }
Instances For
Push forward the multiplication of R on M along a compatible surjective map f : R → S.
See also Function.Surjective.distribMulActionLeft.
Equations
- Function.Surjective.distribSMulLeft f hf hsmul = { toSMulZeroClass := Function.Surjective.smulZeroClassLeft f hf hsmul, smul_add := ⋯ }
Instances For
Compose a DistribSMul with a function, with scalar multiplication f r' • m.
See note [reducible non-instances].
Equations
- DistribSMul.compFun A f = { toSMulZeroClass := SMulZeroClass.compFun A f, smul_add := ⋯ }
Instances For
Each element of the scalars defines an additive monoid homomorphism.
Equations
- DistribSMul.toAddMonoidHom A x = { toFun := fun (x_1 : A) => x • x_1, map_zero' := ⋯, map_add' := ⋯ }
Instances For
Typeclass for multiplicative actions on additive structures.
For example, if G is a group (with group law written as multiplication) and A is an
abelian group (with group law written as addition), then to give A a G-module
structure (for example, to use the theory of group cohomology) is to say [DistribMulAction G A].
Note in that we do not use the Module typeclass for G-modules, as the Module typclass
is for modules over a ring rather than a group.
Mathematically, DistribMulAction G A is equivalent to giving A the structure of
a ℤ[G]-module.
- smul : M → A → A
- Multiplying - 0by a scalar gives- 0
- Scalar multiplication distributes across addition 
Instances
Equations
- DistribMulAction.toDistribSMul = { toSMul := inst✝.toSMul, smul_zero := ⋯, smul_add := ⋯ }
Since Lean 3 does not have definitional eta for structures, we have to make sure
that the definition of DistribMulAction.toDistribSMul was done correctly,
and the two paths from DistribMulAction to SMul are indeed definitionally equal.
Pullback a distributive multiplicative action along an injective additive monoid homomorphism. See note [reducible non-instances].
Equations
- Function.Injective.distribMulAction f hf smul = { toSMul := (Function.Injective.distribSMul f hf smul).toSMul, one_smul := ⋯, mul_smul := ⋯, smul_zero := ⋯, smul_add := ⋯ }
Instances For
Pushforward a distributive multiplicative action along a surjective additive monoid homomorphism. See note [reducible non-instances].
Equations
- Function.Surjective.distribMulAction f hf smul = { toSMul := (Function.Surjective.distribSMul f hf smul).toSMul, one_smul := ⋯, mul_smul := ⋯, smul_zero := ⋯, smul_add := ⋯ }
Instances For
Each element of the monoid defines an additive monoid homomorphism.
Equations
Instances For
Each element of the monoid defines an additive monoid homomorphism.
Equations
- DistribMulAction.toAddMonoidEnd M A = { toFun := DistribMulAction.toAddMonoidHom A, map_one' := ⋯, map_mul' := ⋯ }