Sets invariant to a MulAction #
In this file we define SubMulAction R M; a subset of a MulAction R M which is closed with
respect to scalar multiplication.
For most uses, typically Submodule R M is more powerful.
Main definitions #
- SubMulAction.mulAction- the- MulAction R Mtransferred to the subtype.
- SubMulAction.mulAction'- the- MulAction S Mtransferred to the subtype when- IsScalarTower S R M.
- SubMulAction.isScalarTower- the- IsScalarTower S R Mtransferred to the subtype.
- SubMulAction.inclusion— the inclusion of a submulaction, as an equivariant map
Tags #
submodule, mul_action
SMulMemClass S R M says S is a type of subsets s ≤ M that are closed under the
scalar action of R on M.
Note that only R is marked as an outParam here, since M is supplied by the SetLike
class instead.
- Multiplication by a scalar on an element of the set remains in the set. 
Instances
VAddMemClass S R M says S is a type of subsets s ≤ M that are closed under the
additive action of R on M.
Note that only R is marked as an outParam here, since M is supplied by the SetLike
class instead.
- Addition by a scalar with an element of the set remains in the set. 
Instances
Not registered as an instance because R is an outParam in SMulMemClass S R M.
Not registered as an instance because R is an outParam in SMulMemClass S R M.
A subset closed under the additive action inherits that action.
Equations
- SetLike.vadd s = { vadd := fun (r : R) (x : ↥s) => ⟨r +ᵥ ↑x, ⋯⟩ }
A subset closed under the scalar action inherits that action.
Equations
- SetLike.smul s = { smul := fun (r : R) (x : ↥s) => ⟨r • ↑x, ⋯⟩ }
This can't be an instance because Lean wouldn't know how to find N, but we can still use
this to manually derive SMulMemClass on specific types.
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
A subset closed under the additive action inherits that action.
Equations
- SetLike.vadd' s = { vadd := fun (r : M) (x : ↥s) => ⟨r +ᵥ ↑x, ⋯⟩ }
A subset closed under the scalar action inherits that action.
Equations
- SetLike.smul' s = { smul := fun (r : M) (x : ↥s) => ⟨r • ↑x, ⋯⟩ }
A SubMulAction is a set which is closed under scalar multiplication.
- carrier : Set MThe underlying set of a SubMulAction.
- The carrier set is closed under scalar multiplication. 
Instances For
The carrier set is closed under scalar multiplication.
Equations
- SubMulAction.instSetLike = { coe := SubMulAction.carrier, coe_injective' := ⋯ }
Equations
- ⋯ = ⋯
Copy of a sub_mul_action with a new carrier equal to the old one. Useful to fix definitional
equalities.
Equations
- p.copy s hs = { carrier := s, smul_mem' := ⋯ }
Instances For
Embedding of a submodule p to the ambient space M.
Equations
- p.subtype = { toFun := Subtype.val, map_smul' := ⋯ }
Instances For
A SubMulAction of a MulAction is a MulAction.
Equations
- SubMulAction.SMulMemClass.toMulAction S' = Function.Injective.mulAction Subtype.val ⋯ ⋯
The natural MulActionHom over R from a SubMulAction of M to M.
Equations
- SubMulAction.SMulMemClass.subtype S' = { toFun := Subtype.val, map_smul' := ⋯ }
Instances For
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
Equations
- ⋯ = ⋯
If the scalar product forms a MulAction, then the subset inherits this action
Equations
- p.mulAction' = MulAction.mk ⋯ ⋯
Equations
- p.mulAction = p.mulAction'
Orbits in a SubMulAction coincide with orbits in the ambient space.
Stabilizers in monoid SubMulAction coincide with stabilizers in the ambient space
Stabilizers in group SubMulAction coincide with stabilizers in the ambient space
If the scalar product forms a Module, and the SubMulAction is not ⊥, then the
subset inherits the zero.
Equations
- p.instZeroSubtypeMemOfNonempty = { zero := ⟨0, ⋯⟩ }
The inclusion of a SubMulAction into the ambient set, as an equivariant map
Equations
- s.inclusion = { toFun := Subtype.val, map_smul' := ⋯ }