Miscellaneous definitions, lemmas, and constructions using finsupp #
Main declarations #
finsupp.graph: the finset of input and output pairs with non-zero outputs.alist.lookup_finsupp: converts an association list into a finitely supported function viaalist.lookup, sending absent keys to zero.finsupp.map_range.equiv:finsupp.map_rangeas an equiv.finsupp.map_domain: maps the domain of afinsuppby a function and by summing.finsupp.comap_domain: postcomposition of afinsuppwith a function injective on the preimage of its support.finsupp.some: restrict a finitely supported function onoption αto a finitely supported function onα.finsupp.filter:filter p fis the finitely supported function that isf aifp ais true and 0 otherwise.finsupp.frange: the image of a finitely supported function on its support.finsupp.subtype_domain: the restriction of a finitely supported functionfto a subtype.
Implementation notes #
This file is a noncomputable theory and uses classical logic throughout.
TODO #
-
This file is currently ~1600 lines long and is quite a miscellany of definitions and lemmas, so it should be divided into smaller pieces.
-
Expand the list of definitions and important lemmas to the module docstring.
Declarations about graph #
The graph of a finitely supported function over its support, i.e. the finset of input and output pairs with non-zero outputs.
Declarations about alist.lookup_finsupp #
Converts an association list into a finitely supported function via alist.lookup, sending
absent keys to zero.
Equations
- l.lookup_finsupp = {support := (list.filter (λ (x : Σ (x : α), M), x.snd ≠ 0) l.entries).keys.to_finset, to_fun := λ (a : α), (alist.lookup a l).get_or_else 0, mem_support_to_fun := _}
Alias of alist.lookup_finsupp_to_fun.
Declarations about map_range #
finsupp.map_range as an equiv.
Equations
- finsupp.map_range.equiv f hf hf' = {to_fun := finsupp.map_range ⇑f hf, inv_fun := finsupp.map_range ⇑(f.symm) hf', left_inv := _, right_inv := _}
Composition with a fixed zero-preserving homomorphism is itself an zero-preserving homomorphism on functions.
Equations
- finsupp.map_range.zero_hom f = {to_fun := finsupp.map_range ⇑f _, map_zero' := _}
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' := _}
finsupp.map_range.add_monoid_hom as an equiv.
Equations
- finsupp.map_range.add_equiv f = {to_fun := finsupp.map_range ⇑f _, inv_fun := finsupp.map_range ⇑(f.symm) _, left_inv := _, right_inv := _, map_add' := _}
Declarations about equiv_congr_left #
Given f : α ≃ β, we can map l : α →₀ M to equiv_map_domain f l : β →₀ M (computably)
by mapping the support forwards and the function backwards.
Equations
- finsupp.equiv_map_domain f l = {support := finset.map f.to_embedding l.support, to_fun := λ (a : β), ⇑l (⇑(f.symm) a), mem_support_to_fun := _}
Given f : α ≃ β, the finitely supported function spaces are also in bijection:
(α →₀ M) ≃ (β →₀ M).
This is the finitely-supported version of equiv.Pi_congr_left.
Equations
- finsupp.equiv_congr_left f = {to_fun := finsupp.equiv_map_domain f, inv_fun := finsupp.equiv_map_domain f.symm, left_inv := _, right_inv := _}
Declarations about map_domain #
Given f : α → β and v : α →₀ M, map_domain f v : β →₀ M
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))
finsupp.map_domain is an add_monoid_hom.
Equations
- finsupp.map_domain.add_monoid_hom f = {to_fun := finsupp.map_domain f, map_zero' := _, map_add' := _}
A version of sum_map_domain_index that takes a bundled add_monoid_hom,
rather than separate linearity hypotheses.
When f is an embedding we have an embedding (α →₀ ℕ) ↪ (β →₀ ℕ) given by map_domain.
Equations
- finsupp.map_domain_embedding f = {to_fun := finsupp.map_domain ⇑f, inj' := _}
When g preserves addition, map_range and map_domain commute.
Declarations about comap_domain #
Given f : α → β, l : β →₀ M 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 M 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 := _}
Note the hif argument is needed for this to work in rw.
A version of finsupp.comap_domain_add that's easier to use.
finsupp.comap_domain is an add_monoid_hom.
Equations
- finsupp.comap_domain.add_monoid_hom hf = {to_fun := λ (x : β →₀ M), finsupp.comap_domain f x _, map_zero' := _, map_add' := _}
Restrict a finitely supported function on option α to a finitely supported function on α.
Equations
- f.some = finsupp.comap_domain option.some f _
filter p f is the finitely supported function that is f a if p a is true and 0 otherwise.
Equations
- finsupp.filter p f = {support := finset.filter (λ (a : α), p a) f.support, to_fun := λ (a : α), ite (p a) (⇑f a) 0, mem_support_to_fun := _}
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 subtype p.
Equations
- finsupp.subtype_domain p f = {support := finset.subtype p f.support, to_fun := ⇑f ∘ coe, mem_support_to_fun := _}
subtype_domain but as an add_monoid_hom.
Equations
- finsupp.subtype_domain_add_monoid_hom = {to_fun := finsupp.subtype_domain p, map_zero' := _, map_add' := _}
finsupp.filter as an add_monoid_hom.
Equations
- finsupp.filter_add_hom p = {to_fun := finsupp.filter p, map_zero' := _, map_add' := _}
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 : M), 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 M,
uncurry f is the "uncurried" finitely supported function from α × β to M.
finsupp_prod_equiv defines the equiv between ((α × β) →₀ M) and (α →₀ (β →₀ M)) 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 := _}
finsupp.sum_elim f g maps inl x to f x and inr y to g y.
The equivalence between (α ⊕ β) →₀ γ and (α →₀ γ) × (β →₀ γ).
This is the finsupp version of equiv.sum_arrow_equiv_prod_arrow.
The additive equivalence between (α ⊕ β) →₀ M and (α →₀ M) × (β →₀ M).
This is the finsupp version of equiv.sum_arrow_equiv_prod_arrow.
Equations
Declarations about scalar multiplication #
Scalar multiplication acting on the domain.
This is not an instance as it would conflict with the action on the range.
See the instance_diamonds test for examples of such conflicts.
Equations
- finsupp.comap_has_smul = {smul := λ (g : G), finsupp.map_domain (has_smul.smul g)}
finsupp.comap_has_smul is multiplicative
Equations
- finsupp.comap_mul_action = {to_has_smul := finsupp.comap_has_smul _inst_3, one_smul := _, mul_smul := _}
finsupp.comap_has_smul is distributive
Equations
- finsupp.comap_distrib_mul_action = {to_mul_action := finsupp.comap_mul_action _inst_3, smul_add := _, smul_zero := _}
When G is a group, finsupp.comap_has_smul acts by precomposition with the action of g⁻¹.
Equations
- finsupp.has_smul = {smul := λ (a : R) (v : α →₀ M), finsupp.map_range (has_smul.smul a) _ v}
Throughout this section, some monoid and semiring arguments are specified with {} instead of
[]. See note [implicit instance arguments].
Equations
- finsupp.distrib_mul_action α M = {to_mul_action := {to_has_smul := {smul := has_smul.smul finsupp.has_smul}, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}
Equations
- finsupp.module α M = {to_distrib_mul_action := {to_mul_action := {to_has_smul := {smul := has_smul.smul finsupp.has_smul}, one_smul := _, mul_smul := _}, smul_add := _, smul_zero := _}, add_smul := _, zero_smul := _}
A version of finsupp.comap_domain_smul that's easier to use.
A version of finsupp.sum_smul_index' for bundled additive maps.
finsupp.single as a distrib_mul_action_hom.
See also finsupp.lsingle for the version as a linear map.
Equations
- finsupp.distrib_mul_action_hom.single a = {to_fun := (finsupp.single_add_hom a).to_fun, map_smul' := _, map_zero' := _, map_add' := _}
The finsupp version of pi.unique.
Equations
- finsupp.unique_of_right = finsupp.unique_of_right._proof_2.unique
The finsupp version of pi.unique_of_is_empty.
Equations
- finsupp.unique_of_left = finsupp.unique_of_left._proof_2.unique
Given an add_comm_monoid M and s : set α, restrict_support_equiv s M 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 M = {to_fun := λ (f : {f // ↑(f.support) ⊆ s}), finsupp.subtype_domain (λ (x : α), x ∈ s) f.val, inv_fun := λ (f : ↥s →₀ M), ⟨finsupp.map_domain subtype.val f, _⟩, left_inv := _, right_inv := _}
Given add_comm_monoid M and e : α ≃ β, dom_congr e is the corresponding equiv between
α →₀ M and β →₀ M.
This is finsupp.equiv_congr_left as an add_equiv.
Equations
- finsupp.dom_congr e = {to_fun := finsupp.equiv_map_domain e, inv_fun := finsupp.equiv_map_domain e.symm, left_inv := _, right_inv := _, map_add' := _}
Declarations about sigma types #
Given l, a finitely supported function from the sigma type Σ (i : ι), αs i to M and
an index element i : ι, split l i is the ith component of l,
a finitely supported function from as i to M.
This is the finsupp version of sigma.curry.
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 := _}
On a fintype η, finsupp.split is an equivalence between (Σ (j : η), ιs j) →₀ α
and Π j, (ιs j →₀ α).
This is the finsupp version of equiv.Pi_curry.
Equations
- finsupp.sigma_finsupp_equiv_pi_finsupp = {to_fun := finsupp.split _inst_3, inv_fun := λ (f : Π (j : η), ιs j →₀ α), finsupp.on_finset (finset.univ.sigma (λ (j : η), (f j).support)) (λ (ji : Σ (j : η), ιs j), ⇑(f ji.fst) ji.snd) _, left_inv := _, right_inv := _}
On a fintype η, finsupp.split is an additive equivalence between
(Σ (j : η), ιs j) →₀ α and Π j, (ιs j →₀ α).
This is the add_equiv version of finsupp.sigma_finsupp_equiv_pi_finsupp.