Dependent functions with finite support #
For a non-dependent version see Mathlib.Data.Finsupp.Defs
.
Notation #
This file introduces the notation Π₀ a, β a
as notation for DFinsupp β
, mirroring the α →₀ β
notation used for Finsupp
. This works for nested binders too, with Π₀ a b, γ a b
as notation
for DFinsupp (fun a ↦ DFinsupp (γ a))
.
Implementation notes #
The support is internally represented (in the primed DFinsupp.support'
) as a Multiset
that
represents a superset of the true support of the function, quotiented by the always-true relation so
that this does not impact equality. This approach has computational benefits over storing a
Finset
; it allows us to add together two finitely-supported functions without
having to evaluate the resulting function to recompute its support (which would required
decidability of b = 0
for b : β i
).
The true support of the function can still be recovered with DFinsupp.support
; but these
decidability obligations are now postponed to when the support is actually needed. As a consequence,
there are two ways to sum a DFinsupp
: with DFinsupp.sum
which works over an arbitrary function
but requires recomputation of the support and therefore a Decidable
argument; and with
DFinsupp.sumAddHom
which requires an additive morphism, using its properties to show that
summing over a superset of the support is sufficient.
Finsupp
takes an altogether different approach here; it uses Classical.Decidable
and declares
the Add
instance as noncomputable. This design difference is independent of the fact that
DFinsupp
is dependently-typed and Finsupp
is not; in future, we may want to align these two
definitions, or introduce two more definitions for the other combinations of decisions.
Evaluation at a point is an AddMonoidHom
. This is the finitely-supported version of
Pi.evalAddMonoidHom
.
Equations
- DFinsupp.evalAddMonoidHom i = (Pi.evalAddMonoidHom β i).comp DFinsupp.coeFnAddMonoidHom
Instances For
DFinsupp.prod f g
is the product of g i (f i)
over the support of f
.
Equations
- f.prod g = ∏ i ∈ f.support, g i (f i)
Instances For
sum f g
is the sum of g i (f i)
over the support of f
.
Equations
- f.sum g = ∑ i ∈ f.support, g i (f i)
Instances For
When summing over an AddMonoidHom
, the decidability assumption is not needed, and the result is
also an AddMonoidHom
.
Equations
- One or more equations did not get rendered due to their size.
Instances For
While we didn't need decidable instances to define it, we do to reduce it to a sum
The DFinsupp
version of Finsupp.liftAddHom
,
Equations
- DFinsupp.liftAddHom = { toFun := DFinsupp.sumAddHom, invFun := fun (F : (Π₀ (i : ι), β i) →+ γ) (i : ι) => F.comp (DFinsupp.singleAddHom β i), left_inv := ⋯, right_inv := ⋯, map_add' := ⋯ }
Instances For
The DFinsupp
version of Finsupp.liftAddHom_singleAddHom
,
The DFinsupp
version of Finsupp.liftAddHom_apply_single
,
The DFinsupp
version of Finsupp.liftAddHom_comp_single
,
The DFinsupp
version of Finsupp.comp_liftAddHom
,
Product and sum lemmas for bundled morphisms. #
In this section, we provide analogues of AddMonoidHom.map_sum
, AddMonoidHom.coe_finset_sum
,
and AddMonoidHom.finset_sum_apply
for DFinsupp.sum
and DFinsupp.sumAddHom
instead of
Finset.sum
.
We provide these for AddMonoidHom
, MonoidHom
, RingHom
, AddEquiv
, and MulEquiv
.
Lemmas for LinearMap
and LinearEquiv
are in another file.
The above lemmas, repeated for DFinsupp.sumAddHom
.