Linear structures on function with finite support ι →₀ M
#
This file contains results on the R
-module structure on functions of finite support from a type
ι
to an R
-module M
, in particular in the case that R
is a field.
theorem
Finsupp.linearIndependent_single
{R : Type u_1}
{M : Type u_2}
{ι : Type u_3}
[Ring R]
[AddCommGroup M]
[Module R M]
{φ : ι → Type u_4}
{f : (ι : ι) → φ ι → M}
(hf : ∀ (i : ι), LinearIndependent R (f i))
:
LinearIndependent R fun (ix : (i : ι) × φ i) => single ix.fst (f ix.fst ix.snd)
def
Finsupp.basis
{R : Type u_1}
{M : Type u_2}
{ι : Type u_3}
[Semiring R]
[AddCommMonoid M]
[Module R M]
{φ : ι → Type u_4}
(b : (i : ι) → Basis (φ i) R M)
:
The basis on ι →₀ M
with basis vectors fun ⟨i, x⟩ ↦ single i (b i x)
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
The basis on ι →₀ M
with basis vectors fun i ↦ single i 1
.
Equations
- Finsupp.basisSingleOne = { repr := LinearEquiv.refl R (ι →₀ R) }
Instances For
@[simp]
@[simp]
noncomputable def
DFinsupp.basis
{ι : Type u_1}
{R : Type u_2}
{M : ι → Type u_3}
[Semiring R]
[(i : ι) → AddCommMonoid (M i)]
[(i : ι) → Module R (M i)]
{η : ι → Type u_4}
(b : (i : ι) → Basis (η i) R (M i))
:
The direct sum of free modules is free.
Note that while this is stated for DFinsupp
not DirectSum
, the types are defeq.
Equations
- DFinsupp.basis b = { repr := (DFinsupp.mapRange.linearEquiv fun (i : ι) => (b i).repr) ≪≫ₗ (sigmaFinsuppLequivDFinsupp R).symm }
Instances For
TODO: move this section to an earlier file.
theorem
Finset.sum_single_ite
{R : Type u_1}
{n : Type u_3}
[DecidableEq n]
[Semiring R]
[Fintype n]
(a : R)
(i : n)
:
@[simp]
theorem
Basis.equivFun_symm_single
{R : Type u_1}
{M : Type u_2}
{n : Type u_3}
[DecidableEq n]
[Semiring R]
[AddCommMonoid M]
[Module R M]
[Finite n]
(b : Basis n R M)
(i : n)
:
@[deprecated Basis.equivFun_symm_single (since := "2024-08-09")]
theorem
Basis.equivFun_symm_stdBasis
{R : Type u_1}
{M : Type u_2}
{n : Type u_3}
[DecidableEq n]
[Semiring R]
[AddCommMonoid M]
[Module R M]
[Finite n]
(b : Basis n R M)
(i : n)
: