Finsupp
s supported on a given submodule #
Finsupp.restrictDom
:Finsupp.filter
as a linear map toFinsupp.supported s
;Finsupp.supported R R s
and codomainSubmodule.span R (v '' s)
;Finsupp.supportedEquivFinsupp
: a linear equivalence between the functionsα →₀ M
supported ons
and the functionss →₀ M
;Finsupp.domLCongr
: aLinearEquiv
version ofFinsupp.domCongr
;Finsupp.congr
: if the setss
andt
are equivalent, thensupported M R s
is equivalent tosupported M R t
;
Tags #
function with finite support, module, linear algebra
def
Finsupp.supported
{α : Type u_1}
(M : Type u_2)
(R : Type u_5)
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
:
Finsupp.supported M R s
is the R
-submodule of all p : α →₀ M
such that p.support ⊆ s
.
Equations
Instances For
def
Finsupp.restrictDom
{α : Type u_1}
(M : Type u_2)
(R : Type u_5)
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
[DecidablePred fun (x : α) => x ∈ s]
:
Interpret Finsupp.filter s
as a linear map from α →₀ M
to supported M R s
.
Equations
- Finsupp.restrictDom M R s = LinearMap.codRestrict (Finsupp.supported M R s) { toFun := Finsupp.filter fun (x : α) => x ∈ s, map_add' := ⋯, map_smul' := ⋯ } ⋯
Instances For
@[simp]
theorem
Finsupp.restrictDom_apply
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
(l : α →₀ M)
[DecidablePred fun (x : α) => x ∈ s]
:
theorem
Finsupp.restrictDom_comp_subtype
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
[DecidablePred fun (x : α) => x ∈ s]
:
theorem
Finsupp.range_restrictDom
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
[DecidablePred fun (x : α) => x ∈ s]
:
theorem
Finsupp.disjoint_supported_supported_iff
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
[Nontrivial M]
{s t : Set α}
:
def
Finsupp.supportedEquivFinsupp
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
:
Interpret Finsupp.restrictSupportEquiv
as a linear equivalence between
supported M R s
and s →₀ M
.
Equations
Instances For
@[simp]
theorem
Finsupp.supportedEquivFinsupp_symm_apply_coe_toFun
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
(a✝ : ↑s →₀ M)
(a : α)
:
@[simp]
theorem
Finsupp.supportedEquivFinsupp_symm_apply_coe_support_val
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
(a✝ : ↑s →₀ M)
:
@[simp]
theorem
Finsupp.supportedEquivFinsupp_apply_support_val
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
(a✝ : ↥(supported M R s))
:
((supportedEquivFinsupp s) a✝).support.val = Multiset.map (fun (x : { x : α // x ∈ Finset.filter (fun (x : α) => x ∈ s) (↑a✝).support }) => ⟨↑x, ⋯⟩)
(Multiset.filter (fun (x : α) => x ∈ s) (↑a✝).support.val).attach
@[simp]
theorem
Finsupp.supportedEquivFinsupp_symm_apply_coe
{α : Type u_1}
{M : Type u_2}
{R : Type u_5}
[Semiring R]
[AddCommMonoid M]
[Module R M]
(s : Set α)
[DecidablePred fun (x : α) => x ∈ s]
(f : ↑s →₀ M)
:
theorem
Finsupp.supported_comap_lmapDomain
{α : Type u_1}
(M : Type u_2)
(R : Type u_5)
[Semiring R]
[AddCommMonoid M]
[Module R M]
{α' : Type u_7}
(f : α → α')
(s : Set α')
:
theorem
Finsupp.lmapDomain_supported
{α : Type u_1}
(M : Type u_2)
(R : Type u_5)
[Semiring R]
[AddCommMonoid M]
[Module R M]
{α' : Type u_7}
(f : α → α')
(s : Set α)
:
theorem
Finsupp.lmapDomain_disjoint_ker
{α : Type u_1}
(M : Type u_2)
(R : Type u_5)
[Semiring R]
[AddCommMonoid M]
[Module R M]
{α' : Type u_7}
(f : α → α')
{s : Set α}
(H : ∀ a ∈ s, ∀ b ∈ s, f a = f b → a = b)
:
Disjoint (supported M R s) (LinearMap.ker (lmapDomain M R f))