Dependent functions with finite support #
For a non-dependent version see data/finsupp.lean.
Equations
- dfinsupp.inhabited_pre ι β = {default := {to_fun := λ (i : ι), 0, pre_support := ∅, zero := _}}
Equations
- dfinsupp.pre.setoid ι β = {r := λ (x y : dfinsupp.pre ι β), ∀ (i : ι), x.to_fun i = y.to_fun i, iseqv := _}
Equations
- dfinsupp.fun_like = {coe := λ (f : Π₀ (i : ι), β i), quotient.lift_on f dfinsupp.pre.to_fun dfinsupp.fun_like._proof_1, coe_injective' := _}
Helper instance for when there are too many metavariables to apply fun_like.has_coe_to_fun
directly.
Equations
Deprecated. Use fun_like.ext_iff instead.
Deprecated. Use fun_like.coe_injective instead.
Equations
- dfinsupp.inhabited = {default := 0}
The composition of f : β₁ → β₂ and g : Π₀ i, β₁ i is
map_range f hf g : Π₀ i, β₂ i, well defined when f 0 = 0.
This preserves the structure on f, and exists in various bundled forms for when f is itself
bundled:
dfinsupp.map_range.add_monoid_homdfinsupp.map_range.add_equivdfinsupp.map_range.linear_mapdfinsupp.map_range.linear_equiv
Equations
- dfinsupp.map_range f hf = quotient.map (λ (x : dfinsupp.pre ι (λ (i : ι), β₁ i)), {to_fun := λ (i : ι), f i (x.to_fun i), pre_support := x.pre_support, zero := _}) _
Let f i be a binary operation β₁ i → β₂ i → β i such that f i 0 0 = 0.
Then zip_with f hf is a binary operation Π₀ i, β₁ i → Π₀ i, β₂ i → Π₀ i, β i.
Equations
- dfinsupp.zip_with f hf = quotient.map₂ (λ (x : dfinsupp.pre ι (λ (i : ι), β₁ i)) (y : dfinsupp.pre ι (λ (i : ι), β₂ i)), {to_fun := λ (i : ι), f i (x.to_fun i) (y.to_fun i), pre_support := x.pre_support + y.pre_support, zero := _}) _
Equations
- dfinsupp.has_add = {add := dfinsupp.zip_with (λ (_x : ι), has_add.add) dfinsupp.has_add._proof_1}
Equations
- dfinsupp.add_zero_class = function.injective.add_zero_class coe_fn dfinsupp.add_zero_class._proof_1 dfinsupp.add_zero_class._proof_2 dfinsupp.add_zero_class._proof_3
Note the general dfinsupp.has_scalar instance doesn't apply as ℕ is not distributive
unless β i's addition is commutative.
Equations
- dfinsupp.has_nat_scalar = {smul := λ (c : ℕ) (v : Π₀ (i : ι), β i), dfinsupp.map_range (λ (_x : ι), has_scalar.smul c) _ v}
Equations
- dfinsupp.add_monoid = function.injective.add_monoid coe_fn dfinsupp.add_monoid._proof_1 dfinsupp.add_monoid._proof_2 dfinsupp.add_monoid._proof_3 dfinsupp.add_monoid._proof_4
Coercion from a dfinsupp to a pi type is an add_monoid_hom.
Equations
- dfinsupp.coe_fn_add_monoid_hom = {to_fun := coe_fn dfinsupp.has_coe_to_fun, map_zero' := _, map_add' := _}
Evaluation at a point is an add_monoid_hom. This is the finitely-supported version of
pi.eval_add_monoid_hom.
Equations
Equations
- dfinsupp.add_comm_monoid = function.injective.add_comm_monoid coe_fn dfinsupp.add_comm_monoid._proof_1 dfinsupp.add_comm_monoid._proof_2 dfinsupp.add_comm_monoid._proof_3 dfinsupp.add_comm_monoid._proof_4
Equations
- dfinsupp.has_neg = {neg := λ (f : Π₀ (i : ι), β i), dfinsupp.map_range (λ (_x : ι), has_neg.neg) dfinsupp.has_neg._proof_1 f}
Equations
- dfinsupp.has_sub = {sub := dfinsupp.zip_with (λ (_x : ι), has_sub.sub) dfinsupp.has_sub._proof_1}
Note the general dfinsupp.has_scalar instance doesn't apply as ℤ is not distributive
unless β i's addition is commutative.
Equations
- dfinsupp.has_int_scalar = {smul := λ (c : ℤ) (v : Π₀ (i : ι), β i), dfinsupp.map_range (λ (_x : ι), has_scalar.smul c) _ v}
Equations
- dfinsupp.add_group = function.injective.add_group coe_fn dfinsupp.add_group._proof_1 dfinsupp.add_group._proof_2 dfinsupp.add_group._proof_3 dfinsupp.add_group._proof_4 dfinsupp.add_group._proof_5 dfinsupp.add_group._proof_6 dfinsupp.add_group._proof_7
Equations
- dfinsupp.add_comm_group = function.injective.add_comm_group coe_fn dfinsupp.add_comm_group._proof_1 dfinsupp.add_comm_group._proof_2 dfinsupp.add_comm_group._proof_3 dfinsupp.add_comm_group._proof_4 dfinsupp.add_comm_group._proof_5 dfinsupp.add_comm_group._proof_6 dfinsupp.add_comm_group._proof_7
Dependent functions with finite support inherit a semiring action from an action on each coordinate.
Equations
- dfinsupp.has_scalar = {smul := λ (c : γ) (v : Π₀ (i : ι), β i), dfinsupp.map_range (λ (_x : ι), has_scalar.smul c) _ v}
Dependent functions with finite support inherit a distrib_mul_action structure from such a
structure on each coordinate.
Equations
- dfinsupp.distrib_mul_action = function.injective.distrib_mul_action dfinsupp.coe_fn_add_monoid_hom dfinsupp.distrib_mul_action._proof_1 dfinsupp.distrib_mul_action._proof_2
Dependent functions with finite support inherit a module structure from such a structure on each coordinate.
Equations
- dfinsupp.module = {to_distrib_mul_action := {to_mul_action := distrib_mul_action.to_mul_action dfinsupp.distrib_mul_action, smul_add := _, smul_zero := _}, add_smul := _, zero_smul := _}
filter p f is the function which is f i if p i is true and 0 otherwise.
Equations
- dfinsupp.filter p = quotient.map (λ (x : dfinsupp.pre ι (λ (i : ι), β i)), {to_fun := λ (i : ι), ite (p i) (x.to_fun i) 0, pre_support := x.pre_support, zero := _}) _
dfinsupp.filter as an add_monoid_hom.
Equations
- dfinsupp.filter_add_monoid_hom β p = {to_fun := dfinsupp.filter p (λ (a : ι), _inst_2 a), map_zero' := _, map_add' := _}
dfinsupp.filter as a linear_map.
Equations
- dfinsupp.filter_linear_map γ β p = {to_fun := dfinsupp.filter p (λ (a : ι), _inst_4 a), map_add' := _, map_smul' := _}
subtype_domain p f is the restriction of the finitely supported function
f to the subtype p.
Equations
- dfinsupp.subtype_domain p = quotient.map (λ (x : dfinsupp.pre ι (λ (i : ι), β i)), {to_fun := λ (i : subtype p), x.to_fun ↑i, pre_support := multiset.map (λ (j : {x_1 // x_1 ∈ multiset.filter p x.pre_support}), ⟨↑j, _⟩) (multiset.filter p x.pre_support).attach, zero := _}) _
subtype_domain but as an add_monoid_hom.
Equations
- dfinsupp.subtype_domain_add_monoid_hom β p = {to_fun := dfinsupp.subtype_domain p (λ (a : ι), _inst_2 a), map_zero' := _, map_add' := _}
dfinsupp.subtype_domain as a linear_map.
Equations
- dfinsupp.subtype_domain_linear_map γ β p = {to_fun := dfinsupp.subtype_domain p (λ (a : ι), _inst_4 a), map_add' := _, map_smul' := _}
Create an element of Π₀ i, β i from a finset s and a function x
defined on this finset.
Equations
- dfinsupp.unique = {to_inhabited := {default := 0}, uniq := _}
Given fintype ι, equiv_fun_on_fintype is the equiv between Π₀ i, β i and Π i, β i.
(All dependent functions on a finite type are finitely supported.)
Equations
- dfinsupp.equiv_fun_on_fintype = {to_fun := coe_fn dfinsupp.has_coe_to_fun, inv_fun := λ (f : Π (i : ι), β i), ⟦{to_fun := f, pre_support := finset.univ.val, zero := _}⟧, left_inv := _, right_inv := _}
The function single i b : Π₀ i, β i sends i to b
and all other points to 0.
Equations
- dfinsupp.single i b = dfinsupp.mk {i} (λ (j : ↥↑{i}), _.rec_on b)
Like finsupp.single_eq_single_iff, but with a heq due to dependent types
Equality of sigma types is sufficient (but not necessary) to show equality of dfinsupps.
Redefine f i to be 0.
Equations
- dfinsupp.erase i = quotient.map (λ (x : dfinsupp.pre ι (λ (i : ι), β i)), {to_fun := λ (j : ι), ite (j = i) 0 (x.to_fun j), pre_support := x.pre_support, zero := _}) _
Replace the value of a Π₀ i, β i at a given point i : ι by a given value b : β i.
If b = 0, this amounts to removing i from the support.
Otherwise, i is added to it.
This is the (dependent) finitely-supported version of function.update.
Equations
- dfinsupp.update i f b = quotient.map (λ (x : dfinsupp.pre ι (λ (i : ι), β i)), {to_fun := function.update x.to_fun i b, pre_support := ite (b = 0) (x.pre_support.erase i) (i ::ₘ x.pre_support), zero := _}) _ f
dfinsupp.single as an add_monoid_hom.
Equations
- dfinsupp.single_add_hom β i = {to_fun := dfinsupp.single i, map_zero' := _, map_add' := _}
dfinsupp.erase as an add_monoid_hom.
Equations
- dfinsupp.erase_add_hom β i = {to_fun := dfinsupp.erase i, map_zero' := _, map_add' := _}
If two additive homomorphisms from Π₀ i, β i are equal on each single a b, then
they are equal.
If two additive homomorphisms from Π₀ i, β i are equal on each single a b, then
they are equal.
If s is a subset of ι then mk_add_group_hom s is the canonical additive
group homomorphism from $\prod_{i\in s}\beta_i$ to $\prod_{\mathtt{i : \iota}}\beta_i.$
Equations
- dfinsupp.mk_add_group_hom s = {to_fun := dfinsupp.mk s, map_zero' := _, map_add' := _}
Set {i | f x ≠ 0} as a finset.
Equations
- f.support = quotient.lift_on f (λ (x : dfinsupp.pre ι (λ (i : ι), β i)), finset.filter (λ (i : ι), x.to_fun i ≠ 0) x.pre_support.to_finset) dfinsupp.support._proof_1
Equations
- dfinsupp.decidable_zero = λ (f : Π₀ (i : ι), β i), decidable_of_iff (f.support = ∅) _
Reindexing (and possibly removing) terms of a dfinsupp.
Equations
- dfinsupp.comap_domain h hh = quotient.lift (λ (f : dfinsupp.pre ι (λ (i : ι), β i)), ⟦{to_fun := λ (x : κ), f.to_fun (h x), pre_support := (f.pre_support.to_finset.preimage h _).val, zero := _}⟧) _
A computable version of comap_domain when an explicit left inverse is provided.
Equations
- dfinsupp.comap_domain' h hh' = quotient.lift (λ (f : dfinsupp.pre ι (λ (i : ι), β i)), ⟦{to_fun := λ (x : κ), f.to_fun (h x), pre_support := multiset.map h' f.pre_support, zero := _}⟧) _
Reindexing terms of a dfinsupp.
This is the dfinsupp version of equiv.Pi_congr_left'.
Equations
- dfinsupp.equiv_congr_left h = {to_fun := dfinsupp.comap_domain' ⇑(h.symm) _, inv_fun := λ (f : Π₀ (k : κ), β (⇑(h.symm) k)), dfinsupp.map_range (λ (i : ι), ⇑(equiv.cast _)) _ (dfinsupp.comap_domain' ⇑h _ f), left_inv := _, right_inv := _}
Equations
Equations
Equations
The natural map between Π₀ (i : Σ i, α i), δ i.1 i.2 and Π₀ i (j : α i), δ i j.
The natural map between Π₀ i (j : α i), δ i j and Π₀ (i : Σ i, α i), δ i.1 i.2, inverse of
curry.
Equations
- f.sigma_uncurry = dfinsupp.mk (f.support.bUnion (λ (i : ι), finset.image (sigma.mk i) (⇑f i).support)) (λ (_x : ↥↑(f.support.bUnion (λ (i : ι), finset.image (sigma.mk i) (⇑f i).support))), dfinsupp.sigma_uncurry._match_1 f classical.prop_decidable _x)
- dfinsupp.sigma_uncurry._match_1 f _inst ⟨⟨i, j⟩, _x⟩ = ⇑(⇑f i) j
The natural bijection between Π₀ (i : Σ i, α i), δ i.1 i.2 and Π₀ i (j : α i), δ i j.
This is the dfinsupp version of equiv.Pi_curry.
Equations
- dfinsupp.sigma_curry_equiv = {to_fun := dfinsupp.sigma_curry _inst_1, inv_fun := dfinsupp.sigma_uncurry (λ (i : ι) (i_1 : α i), _inst_1 i i_1), left_inv := _, right_inv := _}
Adds a term to a dfinsupp, making a dfinsupp indexed by an option.
This is the dfinsupp version of option.rec.
Equations
- dfinsupp.extend_with a = quotient.lift (λ (f : dfinsupp.pre ι (λ (i : ι), α (some i))), ⟦{to_fun := option.rec a f.to_fun, pre_support := none ::ₘ multiset.map some f.pre_support, zero := _}⟧) _
Bijection obtained by separating the term of index none of a dfinsupp over option ι.
This is the dfinsupp version of equiv.pi_option_equiv_prod.
prod f g is the product of g i (f i) over the support of f.
sum f g is the sum of g i (f i) over the support of f.
When summing over an add_monoid_hom, the decidability assumption is not needed, and the result is
also an add_monoid_hom.
Equations
- dfinsupp.sum_add_hom φ = {to_fun := λ (f : Π₀ (i : ι), β i), quotient.lift_on f (λ (x : dfinsupp.pre ι (λ (i : ι), β i)), ∑ (i : ι) in x.pre_support.to_finset, ⇑(φ i) (x.to_fun i)) _, map_zero' := _, map_add' := _}
While we didn't need decidable instances to define it, we do to reduce it to a sum
The supremum of a family of commutative additive submonoids is equal to the range of
dfinsupp.sum_add_hom; that is, every element in the supr can be produced from taking a finite
number of non-zero elements of S i, coercing them to γ, and summing them.
The bounded supremum of a family of commutative additive submonoids is equal to the range of
dfinsupp.sum_add_hom composed with dfinsupp.filter_add_monoid_hom; that is, every element in the
bounded supr can be produced from taking a finite number of non-zero elements from the S i that
satisfy p i, coercing them to γ, and summing them.
A variant of add_submonoid.mem_supr_iff_exists_dfinsupp with the RHS fully unfolded.
The dfinsupp version of finsupp.lift_add_hom,
Equations
- dfinsupp.lift_add_hom = {to_fun := dfinsupp.sum_add_hom _inst_2, inv_fun := λ (F : (Π₀ (i : ι), β i) →+ γ) (i : ι), F.comp (dfinsupp.single_add_hom β i), left_inv := _, right_inv := _, map_add' := _}
The dfinsupp version of finsupp.lift_add_hom_single_add_hom,
The dfinsupp version of finsupp.lift_add_hom_apply_single,
The dfinsupp version of finsupp.lift_add_hom_comp_single,
The dfinsupp version of finsupp.comp_lift_add_hom,
Bundled versions of dfinsupp.map_range #
The names should match the equivalent bundled finsupp.map_range definitions.
dfinsupp.map_range as an add_monoid_hom.
Equations
- dfinsupp.map_range.add_monoid_hom f = {to_fun := dfinsupp.map_range (λ (i : ι) (x : β₁ i), ⇑(f i) x) _, map_zero' := _, map_add' := _}
dfinsupp.map_range.add_monoid_hom as an add_equiv.
Equations
- dfinsupp.map_range.add_equiv e = {to_fun := dfinsupp.map_range (λ (i : ι) (x : β₁ i), ⇑(e i) x) _, inv_fun := dfinsupp.map_range (λ (i : ι) (x : β₂ i), ⇑((e i).symm) x) _, left_inv := _, right_inv := _, map_add' := _}
Product and sum lemmas for bundled morphisms. #
In this section, we provide analogues of add_monoid_hom.map_sum, add_monoid_hom.coe_finset_sum,
and add_monoid_hom.finset_sum_apply for dfinsupp.sum and dfinsupp.sum_add_hom instead of
finset.sum.
We provide these for add_monoid_hom, monoid_hom, ring_hom, add_equiv, and mul_equiv.
Lemmas for linear_map and linear_equiv are in another file.
The above lemmas, repeated for dfinsupp.sum_add_hom.