Type of functions with finite support
For any type α
and a type β
with zero, we define the type finsupp α β
of finitely supported
functions from α
to β
, i.e. the functions which are zero everywhere on α
except on a finite
set. We write this in infix notation as α →₀ β
.
Functions with finite support provide the basis for the following concrete instances:
ℕ →₀ α
: Polynomials (whereα
is a ring)(σ →₀ ℕ) →₀ α
: Multivariate Polynomials (againα
is a ring, andσ
are variable names)α →₀ ℕ
: Multisetsα →₀ ℤ
: Abelian groups freely generated byα
β →₀ α
: Linear combinations overβ
whereα
is the scalar ring
Most of the theory assumes that the range is a commutative monoid. This gives us the big sum
operator as a powerful way to construct finsupp
elements.
A general piece of advice is to not use α →₀ β
directly, as the type class setup might not be a
good fit. Defining a copy and selecting the instances that are best suited for the application works
better.
Implementation notes
This file is a noncomputable theory
and uses classical logic throughout.
Notation
This file defines α →₀ β
as notation for finsupp α β
.
finsupp α β
, denoted α →₀ β
, is the type of functions f : α → β
such that
f x = 0
for all but finitely many x
.
Basic declarations about finsupp
Equations
- finsupp.has_coe_to_fun = {F := λ (_x : α →₀ β), α → β, coe := finsupp.to_fun _inst_1}
Equations
- finsupp.has_zero = {zero := {support := ∅, to_fun := λ (_x : α), 0, mem_support_to_fun := _}}
Equations
- finsupp.inhabited = {default := 0}
Given fintype α
, equiv_fun_on_fintype
is the equiv
between α →₀ β
and α → β
.
(All functions on a finite type are finitely supported.)
Equations
- finsupp.equiv_fun_on_fintype = {to_fun := λ (f : α →₀ β) (a : α), ⇑f a, inv_fun := λ (f : α → β), {support := finset.filter (λ (a : α), f a ≠ 0) finset.univ, to_fun := f, mem_support_to_fun := _}, left_inv := _, right_inv := _}
Declarations about single
single a b
is the finitely supported function which has
value b
at a
and zero otherwise.
Declarations about on_finset
on_finset s f hf
is the finsupp function representing f
restricted to the finset s
.
The function needs to be 0
outside of s
. Use this when the set needs to be filtered anyways,
otherwise a better set representation is often available.
Equations
- finsupp.on_finset s f hf = {support := finset.filter (λ (a : α), f a ≠ 0) s, to_fun := f, mem_support_to_fun := _}
Declarations about map_range
The composition of f : β₁ → β₂
and g : α →₀ β₁
is
map_range f hf g : α →₀ β₂
, well-defined when f 0 = 0
.
Equations
- finsupp.map_range f hf g = finsupp.on_finset g.support (f ∘ ⇑g) _
Declarations about emb_domain
Given f : α₁ ↪ α₂
and v : α₁ →₀ β
, emb_domain f v : α₂ →₀ β
is the finitely supported function whose value at f a : α₂
is v a
.
For a b : α₂
outside the range of f
, it is zero.
Equations
- finsupp.emb_domain f v = {support := finset.map f v.support, to_fun := λ (a₂ : α₂), dite (a₂ ∈ finset.map f v.support) (λ (h : a₂ ∈ finset.map f v.support), ⇑v (finset.choose (λ (a₁ : α₁), ⇑f a₁ = a₂) v.support _)) (λ (h : a₂ ∉ finset.map f v.support), 0), mem_support_to_fun := _}
Declarations about zip_with
zip_with f hf g₁ g₂
is the finitely supported function satisfying
zip_with f hf g₁ g₂ a = f (g₁ a) (g₂ a)
, and it is well-defined when f 0 0 = 0
.
Equations
- finsupp.zip_with f hf g₁ g₂ = finsupp.on_finset (g₁.support ∪ g₂.support) (λ (a : α), f (⇑g₁ a) (⇑g₂ a)) _
Declarations about erase
erase a f
is the finitely supported function equal to f
except at a
where it is equal to
0
.
Declarations about sum
and prod
In most of this section, the domain β
is assumed to be an add_monoid
.
A restatement of prod_ite_eq
with the equality test reversed.
A restatement of sum_ite_eq
with the equality test reversed.
If g
maps a second argument of 0 to 0, summing it over the
result of on_finset
is the same as summing it over the original
finset
.
Equations
- finsupp.has_add = {add := finsupp.zip_with has_add.add finsupp.has_add._proof_1}
Equations
- finsupp.add_monoid = {add := has_add.add finsupp.has_add, add_assoc := _, zero := 0, zero_add := _, add_zero := _}
Evaluation of a function f : α →₀ β
at a point as an additive monoid homomorphism.
Equations
- finsupp.nat_sub = {sub := finsupp.zip_with (λ (m n : ℕ), m - n) finsupp.nat_sub._proof_1}
Equations
- finsupp.add_comm_monoid = {add := add_monoid.add finsupp.add_monoid, add_assoc := _, zero := add_monoid.zero finsupp.add_monoid, zero_add := _, add_zero := _, add_comm := _}
Equations
- finsupp.add_group = {add := add_monoid.add finsupp.add_monoid, add_assoc := _, zero := add_monoid.zero finsupp.add_monoid, zero_add := _, add_zero := _, neg := finsupp.map_range has_neg.neg neg_zero, add_left_neg := _}
Equations
- finsupp.add_comm_group = {add := add_group.add finsupp.add_group, add_assoc := _, zero := add_group.zero finsupp.add_group, zero_add := _, add_zero := _, neg := add_group.neg finsupp.add_group, add_left_neg := _, add_comm := _}
Composition with a fixed additive homomorphism is itself an additive homomorphism on functions.
Equations
- finsupp.map_range.add_monoid_hom f = {to_fun := finsupp.map_range ⇑f _, map_zero' := _, map_add' := _}
Declarations about map_domain
Given f : α₁ → α₂
and v : α₁ →₀ β
, map_domain f v : α₂ →₀ β
is the finitely supported function whose value at a : α₂
is the sum
of v x
over all x
such that f x = a
.
Equations
- finsupp.map_domain f v = v.sum (λ (a : α₁), finsupp.single (f a))
Declarations about comap_domain
Given f : α₁ → α₂
, l : α₂ →₀ γ
and a proof hf
that f
is injective on
the preimage of l.support
, comap_domain f l hf
is the finitely supported function
from α₁
to γ
given by composing l
with f
.
Equations
- finsupp.comap_domain f l hf = {support := l.support.preimage f hf, to_fun := λ (a : α₁), ⇑l (f a), mem_support_to_fun := _}
Declarations about filter
filter p f
is the function which is f a
if p a
is true and 0 otherwise.
Equations
- finsupp.filter p f = finsupp.on_finset f.support (λ (a : α), ite (p a) (⇑f a) 0) _
Declarations about frange
frange f
is the image of f
on the support of f
.
Equations
- f.frange = finset.image ⇑f f.support
Declarations about subtype_domain
subtype_domain p f
is the restriction of the finitely supported function
f
to the subtype p
.
Equations
- finsupp.subtype_domain p f = {support := finset.subtype p f.support, to_fun := ⇑f ∘ subtype.val, mem_support_to_fun := _}
Equations
Equations
- _ = _
Given f : α →₀ ℕ
, f.to_multiset
is the multiset with multiplicities given by the values of
f
on the elements of α
.
Equations
- f.to_multiset = f.sum (λ (a : α) (n : ℕ), n •ℕ {a})
Equations
Given m : multiset α
, of_multiset m
is the finitely supported function from α
to ℕ
given by the multiplicities of the elements of α
in m
.
Equations
- finsupp.of_multiset m = finsupp.on_finset m.to_finset (λ (a : α), multiset.count a m) _
equiv_multiset
defines an equiv
between finitely supported functions
from α
to ℕ
and multisets on α
.
Equations
- finsupp.equiv_multiset = {to_fun := finsupp.to_multiset α, inv_fun := finsupp.of_multiset α, left_inv := _, right_inv := _}
Declarations about curry
and uncurry
Given a finitely supported function f
from a product type α × β
to γ
,
curry f
is the "curried" finitely supported function from α
to the type of
finitely supported functions from β
to γ
.
Equations
- f.curry = f.sum (λ (p : α × β) (c : γ), finsupp.single p.fst (finsupp.single p.snd c))
Given a finitely supported function f
from α
to the type of
finitely supported functions from β
to γ
,
uncurry f
is the "uncurried" finitely supported function from α × β
to γ
.
finsupp_prod_equiv
defines the equiv
between ((α × β) →₀ γ)
and (α →₀ (β →₀ γ))
given by
currying and uncurrying.
Equations
- finsupp.finsupp_prod_equiv = {to_fun := finsupp.curry _inst_1, inv_fun := finsupp.uncurry _inst_1, left_inv := _, right_inv := _}
Scalar multiplication by a group element g, given by precomposition with the action of g⁻¹ on the domain.
Equations
- finsupp.comap_has_scalar = {smul := λ (g : γ) (f : α →₀ β), finsupp.comap_domain (λ (a : α), g⁻¹ • a) f _}
Scalar multiplication by a group element, given by precomposition with the action of g⁻¹ on the domain, is multiplicative in g.
Equations
- finsupp.comap_mul_action = {to_has_scalar := finsupp.comap_has_scalar _inst_3, one_smul := _, mul_smul := _}
Scalar multiplication by a group element, given by precomposition with the action of g⁻¹ on the domain, is additive in the second argument.
Equations
- finsupp.comap_distrib_mul_action = {to_mul_action := finsupp.comap_mul_action _inst_3, smul_add := _, smul_zero := _}
Scalar multiplication by a group element on finitely supported functions on a group, given by precomposition with the action of g⁻¹.
Equations
- finsupp.has_scalar = {smul := λ (a : γ) (v : α →₀ β), finsupp.map_range (has_scalar.smul a) _ v}
Equations
- finsupp.semimodule α β = {to_distrib_mul_action := {to_mul_action := {to_has_scalar := {smul := has_scalar.smul finsupp.has_scalar}, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}, add_smul := _, zero_smul := _}
Evaluation at point as a linear map. This version assumes that the codomain is a semimodule
over some semiring. See also leval
.
Evaluation at point as a linear map. This version assumes that the codomain is a semiring.
Equations
- finsupp.leval a = finsupp.leval' β a
Given an add_comm_monoid β
and s : set α
, restrict_support_equiv s β
is the equiv
between the subtype of finitely supported functions with support contained in s
and
the type of finitely supported functions from s
.
Equations
- finsupp.restrict_support_equiv s β = {to_fun := λ (f : {f // ↑(f.support) ⊆ s}), finsupp.subtype_domain (λ (x : α), x ∈ s) f.val, inv_fun := λ (f : ↥s →₀ β), ⟨finsupp.map_domain subtype.val f, _⟩, left_inv := _, right_inv := _}
Given add_comm_monoid β
and e : α₁ ≃ α₂
, dom_congr e
is the corresponding equiv
between
α₁ →₀ β
and α₂ →₀ β
.
Equations
- finsupp.dom_congr e = {to_fun := finsupp.map_domain ⇑e, inv_fun := finsupp.map_domain ⇑(e.symm), left_inv := _, right_inv := _}
Declarations about sigma types
Given l
, a finitely supported function from the sigma type Σ (i : ι), αs i
to β
and
an index element i : ι
, split l i
is the i
th component of l
,
a finitely supported function from as i
to β
.
Equations
- l.split i = finsupp.comap_domain (sigma.mk i) l _
Given l
, a finitely supported function from the sigma type Σ (i : ι), αs i
to β
,
split_support l
is the finset of indices in ι
that appear in the support of l
.
Equations
Given l
, a finitely supported function from the sigma type Σ i, αs i
to β
and
an ι
-indexed family g
of functions from (αs i →₀ β)
to γ
, split_comp
defines a
finitely supported function from the index type ι
to γ
given by composing g i
with
split l i
.
Equations
- l.split_comp g hg = {support := l.split_support, to_fun := λ (i : ι), g i (l.split i), mem_support_to_fun := _}
Given a multiset s
, s.to_finsupp
returns the finitely supported function on ℕ
given by
the multiplicities of the elements of s
.
Equations
- s.to_finsupp = {support := s.to_finset, to_fun := λ (a : α), multiset.count a s, mem_support_to_fun := _}
Equations
Declarations about order(ed) instances on finsupp
Equations
- finsupp.partial_order = {le := preorder.le finsupp.preorder, lt := preorder.lt finsupp.preorder, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _}
Equations
Equations
Equations
- finsupp.ordered_cancel_add_comm_monoid = {add := add_comm_monoid.add finsupp.add_comm_monoid, add_assoc := _, zero := add_comm_monoid.zero finsupp.add_comm_monoid, zero_add := _, add_zero := _, add_comm := _, add_left_cancel := _, add_right_cancel := _, le := partial_order.le finsupp.partial_order, lt := partial_order.lt finsupp.partial_order, le_refl := _, le_trans := _, lt_iff_le_not_le := _, le_antisymm := _, add_le_add_left := _, le_of_add_le_add_left := _}
The order on σ →₀ ℕ
is well-founded.
Equations
- finsupp.decidable_le σ = λ (m n : σ →₀ ℕ), _.mpr finset.decidable_dforall_finset
The finsupp
counterpart of multiset.antidiagonal
: the antidiagonal of
s : σ →₀ ℕ
consists of all pairs (t₁, t₂) : (σ →₀ ℕ) × (σ →₀ ℕ)
such that t₁ + t₂ = s
.
The finitely supported function antidiagonal s
is equal to the multiplicities of these pairs.