Sigma instances for additive and multiplicative actions #
This file defines instances for arbitrary sum of additive and multiplicative actions.
See also #
instance
Sigma.instIsScalarTowerOfSMul
{ι : Type u_1}
{M : Type u_2}
{N : Type u_3}
{α : ι → Type u_4}
[(i : ι) → SMul M (α i)]
[(i : ι) → SMul N (α i)]
[SMul M N]
[∀ (i : ι), IsScalarTower M N (α i)]
:
IsScalarTower M N ((i : ι) × α i)
Equations
- ⋯ = ⋯
instance
Sigma.instIsScalarTowerOfVAdd
{ι : Type u_1}
{M : Type u_2}
{N : Type u_3}
{α : ι → Type u_4}
[(i : ι) → VAdd M (α i)]
[(i : ι) → VAdd N (α i)]
[VAdd M N]
[∀ (i : ι), VAddAssocClass M N (α i)]
:
VAddAssocClass M N ((i : ι) × α i)
Equations
- ⋯ = ⋯
instance
Sigma.instSMulCommClass
{ι : Type u_1}
{M : Type u_2}
{N : Type u_3}
{α : ι → Type u_4}
[(i : ι) → SMul M (α i)]
[(i : ι) → SMul N (α i)]
[∀ (i : ι), SMulCommClass M N (α i)]
:
SMulCommClass M N ((i : ι) × α i)
Equations
- ⋯ = ⋯
instance
Sigma.instVAddCommClass
{ι : Type u_1}
{M : Type u_2}
{N : Type u_3}
{α : ι → Type u_4}
[(i : ι) → VAdd M (α i)]
[(i : ι) → VAdd N (α i)]
[∀ (i : ι), VAddCommClass M N (α i)]
:
VAddCommClass M N ((i : ι) × α i)
Equations
- ⋯ = ⋯
instance
Sigma.instIsCentralScalar
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
[(i : ι) → SMul M (α i)]
[(i : ι) → SMul Mᵐᵒᵖ (α i)]
[∀ (i : ι), IsCentralScalar M (α i)]
:
IsCentralScalar M ((i : ι) × α i)
Equations
- ⋯ = ⋯
instance
Sigma.instIsCentralVAdd
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
[(i : ι) → VAdd M (α i)]
[(i : ι) → VAdd Mᵃᵒᵖ (α i)]
[∀ (i : ι), IsCentralVAdd M (α i)]
:
IsCentralVAdd M ((i : ι) × α i)
Equations
- ⋯ = ⋯
theorem
Sigma.FaithfulSMul'
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
[(i : ι) → SMul M (α i)]
(i : ι)
[FaithfulSMul M (α i)]
:
FaithfulSMul M ((i : ι) × α i)
This is not an instance because i
becomes a metavariable.
theorem
Sigma.FaithfulVAdd'
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
[(i : ι) → VAdd M (α i)]
(i : ι)
[FaithfulVAdd M (α i)]
:
FaithfulVAdd M ((i : ι) × α i)
This is not an instance because i
becomes a metavariable.
instance
Sigma.instFaithfulSMulOfNonempty
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
[(i : ι) → SMul M (α i)]
[Nonempty ι]
[∀ (i : ι), FaithfulSMul M (α i)]
:
FaithfulSMul M ((i : ι) × α i)
Equations
- ⋯ = ⋯
instance
Sigma.instFaithfulVAddOfNonempty
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
[(i : ι) → VAdd M (α i)]
[Nonempty ι]
[∀ (i : ι), FaithfulVAdd M (α i)]
:
FaithfulVAdd M ((i : ι) × α i)
Equations
- ⋯ = ⋯
instance
Sigma.instMulAction
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
{m : Monoid M}
[(i : ι) → MulAction M (α i)]
:
MulAction M ((i : ι) × α i)
Equations
- Sigma.instMulAction = MulAction.mk ⋯ ⋯
instance
Sigma.instAddAction
{ι : Type u_1}
{M : Type u_2}
{α : ι → Type u_4}
{m : AddMonoid M}
[(i : ι) → AddAction M (α i)]
:
AddAction M ((i : ι) × α i)
Equations
- Sigma.instAddAction = AddAction.mk ⋯ ⋯