Bases and scalar multiplication #
This file defines the scalar multiplication of bases by multiplying each basis vector.
instance
Basis.instSMul
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G M]
[SMulCommClass G R M]
:
The action on a Basis
by acting on each element.
See also Basis.unitsSMul
and Basis.groupSMul
, for the cases when a different action is applied
to each basis element.
Equations
- Basis.instSMul = { smul := fun (g : G) (b : Basis ι R M) => b.map (DistribMulAction.toLinearEquiv R M g) }
@[simp]
theorem
Basis.smul_apply
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G M]
[SMulCommClass G R M]
(g : G)
(b : Basis ι R M)
(i : ι)
:
theorem
Basis.coe_smul
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G M]
[SMulCommClass G R M]
(g : G)
(b : Basis ι R M)
:
@[simp]
theorem
Basis.repr_smul
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G M]
[SMulCommClass G R M]
(g : G)
(b : Basis ι R M)
:
instance
Basis.instMulAction
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G M]
[SMulCommClass G R M]
:
Equations
- Basis.instMulAction = Function.Injective.mulAction (fun (f : Basis ι R M) => ⇑f) ⋯ ⋯
instance
Basis.instSMulCommClass
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
{G' : Type u_8}
[Group G]
[Group G']
[DistribMulAction G M]
[DistribMulAction G' M]
[SMulCommClass G R M]
[SMulCommClass G' R M]
[SMulCommClass G G' M]
:
SMulCommClass G G' (Basis ι R M)
instance
Basis.instIsScalarTower
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
{G' : Type u_8}
[Group G]
[Group G']
[DistribMulAction G M]
[DistribMulAction G' M]
[SMulCommClass G R M]
[SMulCommClass G' R M]
[SMul G G']
[IsScalarTower G G' M]
:
IsScalarTower G G' (Basis ι R M)
theorem
Basis.groupSMul_span_eq_top
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G R]
[DistribMulAction G M]
[IsScalarTower G R M]
{v : ι → M}
(hv : Submodule.span R (Set.range v) = ⊤)
{w : ι → G}
:
def
Basis.groupSMul
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G R]
[DistribMulAction G M]
[IsScalarTower G R M]
[SMulCommClass G R M]
(v : Basis ι R M)
(w : ι → G)
:
Basis ι R M
Given a basis v
and a map w
such that for all i
, w i
are elements of a group,
groupSMul
provides the basis corresponding to w • v
.
Instances For
theorem
Basis.groupSMul_apply
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{G : Type u_7}
[Group G]
[DistribMulAction G R]
[DistribMulAction G M]
[IsScalarTower G R M]
[SMulCommClass G R M]
{v : Basis ι R M}
{w : ι → G}
(i : ι)
:
theorem
Basis.units_smul_span_eq_top
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{v : ι → M}
(hv : Submodule.span R (Set.range v) = ⊤)
{w : ι → Rˣ}
:
theorem
Basis.isUnitSMul_apply
{ι : Type u_1}
{R : Type u_3}
{M : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{v : Basis ι R M}
{w : ι → R}
(hw : ∀ (i : ι), IsUnit (w i))
(i : ι)
:
theorem
Basis.repr_isUnitSMul
{ι : Type u_1}
{R₂ : Type u_4}
{M : Type u_5}
[AddCommMonoid M]
[CommSemiring R₂]
[Module R₂ M]
{v : Basis ι R₂ M}
{w : ι → R₂}
(hw : ∀ (i : ι), IsUnit (w i))
(x : M)
(i : ι)
: