Incidence algebras #
Given a locally finite order α
the incidence algebra over α
is the type of functions from
non-empty intervals of α
to some algebraic codomain.
This algebra has a natural multiplication operation whereby the product of two such functions is defined on an interval by summing over all divisions into two subintervals the product of the values of the original pair of functions.
This structure allows us to interpret many natural invariants of the intervals (such as their cardinality) as elements of the incidence algebra. For instance the cardinality function, viewed as an element of the incidence algebra, is simply the square of the function that takes constant value one on all intervals. This constant function is called the zeta function, after its connection with the Riemann zeta function.
The incidence algebra is a good setting for proving many inclusion-exclusion type principles, these go under the name Möbius inversion, and are essentially due to the fact that the zeta function has a multiplicative inverse in the incidence algebra, an inductively definable function called the Möbius function that generalizes the Möbius function in number theory.
Main definitions #
1 : IncidenceAlgebra 𝕜 α
is the delta function, defined analogously to the identity matrix.f * g
is the incidence algebra product, defined analogously to the matrix product.IncidenceAlgebra.zeta
is the zeta function, defined analogously to the upper triangular matrix of ones.IncidenceAlgebra.mu
is the inverse of the zeta function.
Implementation notes #
One has to define mu
as either the left or right inverse of zeta
. We define it as the left
inverse, and prove it is also a right inverse by defining mu'
as the right inverse and using that
left and right inverses agree if they exist.
TODOs #
Here are some additions to this file that could be made in the future:
- Generalize the construction of
mu
to invert any element of the incidence algebraf
which hasf x x
a unit for allx
. - Give formulas for higher powers of zeta.
- A formula for the möbius function on a pi type similar to the one for products
- More examples / applications to different posets.
- Connection with Galois insertions
- Finsum version of Möbius inversion that holds even when an order doesn't have top/bot?
- Connect this theory to (infinite) matrices, giving maps of the incidence algebra to matrix rings
- Connect to the more advanced theory of arithmetic functions, and Dirichlet convolution.
References #
The 𝕜
-incidence algebra over α
.
- toFun : α → α → 𝕜
The underlying function of an element of the incidence algebra.
Do not use this function directly. Instead use the coercion coming from the
FunLike
instance.
Instances For
Equations
- IncidenceAlgebra.instFunLike = { coe := IncidenceAlgebra.toFun, coe_injective' := ⋯ }
Additive and multiplicative structure #
Equations
- IncidenceAlgebra.instZero = { zero := { toFun := fun (x x : α) => 0, eq_zero_of_not_le' := ⋯ } }
Equations
- IncidenceAlgebra.instInhabited = { default := 0 }
Equations
- IncidenceAlgebra.instAdd = { add := fun (f g : IncidenceAlgebra 𝕜 α) => { toFun := ⇑f + ⇑g, eq_zero_of_not_le' := ⋯ } }
Equations
- IncidenceAlgebra.instSmulZeroClassRight = SMulZeroClass.mk ⋯
Equations
- IncidenceAlgebra.instAddMonoid = Function.Injective.addMonoid (fun (f : IncidenceAlgebra 𝕜 α) => ⇑f) ⋯ ⋯ ⋯ ⋯
Equations
- IncidenceAlgebra.instAddCommMonoid = Function.Injective.addCommMonoid (fun (f : IncidenceAlgebra 𝕜 α) => ⇑f) ⋯ ⋯ ⋯ ⋯
Equations
- IncidenceAlgebra.instNeg = { neg := fun (f : IncidenceAlgebra 𝕜 α) => { toFun := -⇑f, eq_zero_of_not_le' := ⋯ } }
Equations
- IncidenceAlgebra.instSub = { sub := fun (f g : IncidenceAlgebra 𝕜 α) => { toFun := ⇑f - ⇑g, eq_zero_of_not_le' := ⋯ } }
Equations
- IncidenceAlgebra.instAddGroup = Function.Injective.addGroup (fun (f : IncidenceAlgebra 𝕜 α) => ⇑f) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
Equations
- IncidenceAlgebra.instAddCommGroup = Function.Injective.addCommGroup (fun (f : IncidenceAlgebra 𝕜 α) => ⇑f) ⋯ ⋯ ⋯ ⋯ ⋯ ⋯ ⋯
The unit incidence algebra is the delta function, whose entries are 0
except on the diagonal
where they are 1
.
The multiplication operation in incidence algebras is defined on an interval by summing over all divisions into two subintervals the product of the values of the original pair of functions.
Equations
- IncidenceAlgebra.instMul = { mul := fun (f g : IncidenceAlgebra 𝕜 α) => { toFun := fun (a b : α) => ∑ x ∈ Finset.Icc a b, f a x * g x b, eq_zero_of_not_le' := ⋯ } }
Equations
- IncidenceAlgebra.instNonUnitalNonAssocSemiring = NonUnitalNonAssocSemiring.mk ⋯ ⋯ ⋯ ⋯
Equations
- IncidenceAlgebra.instNonAssocSemiring = NonAssocSemiring.mk ⋯ ⋯ ⋯ ⋯
Equations
- IncidenceAlgebra.instSemiring = Semiring.mk ⋯ ⋯ ⋯ ⋯ npowRecAuto ⋯ ⋯
Scalar multiplication between incidence algebras #
Equations
- One or more equations did not get rendered due to their size.
Equations
- ⋯ = ⋯
Equations
- IncidenceAlgebra.smulWithZeroRight = Function.Injective.smulWithZero { toFun := DFunLike.coe, map_zero' := ⋯ } ⋯ ⋯
Equations
- IncidenceAlgebra.moduleRight = Function.Injective.module 𝕜 { toFun := DFunLike.coe, map_zero' := ⋯, map_add' := ⋯ } ⋯ ⋯
Equations
- IncidenceAlgebra.algebraRight = Algebra.mk { toFun := fun (c : 𝕜) => (algebraMap 𝕜 𝕝) c • 1, map_one' := ⋯, map_mul' := ⋯, map_zero' := ⋯, map_add' := ⋯ } ⋯ ⋯
The Lambda function #
The lambda function of the incidence algebra is the function that assigns 1
to every nonempty
interval of cardinality one or two.
Equations
- IncidenceAlgebra.lambda 𝕜 = { toFun := fun (a b : α) => if a ⩿ b then 1 else 0, eq_zero_of_not_le' := ⋯ }
Instances For
The Zeta and Möbius functions #
The zeta function of the incidence algebra is the function that assigns 1 to every nonempty interval, convolution with this function sums functions over intervals.
Equations
- IncidenceAlgebra.zeta 𝕜 = { toFun := fun (a b : α) => if a ≤ b then 1 else 0, eq_zero_of_not_le' := ⋯ }
Instances For
The Möbius function which inverts zeta
as an element of the incidence algebra.
Equations
- IncidenceAlgebra.mu 𝕜 = { toFun := IncidenceAlgebra.muFun 𝕜, eq_zero_of_not_le' := ⋯ }
Instances For
A general form of Möbius inversion. Based on lemma 2.1.2 of Incidence Algebras by Spiegel and O'Donnell.
A general form of Möbius inversion. Based on lemma 2.1.3 of Incidence Algebras by Spiegel and O'Donnell.
The cartesian product of two incidence algebras.
Equations
Instances For
This is a version of IncidenceAlgebra.prod_mul_prod
that works over non-commutative rings.
The Möbius function on a product order. Based on lemma 2.1.13 of Incidence Algebras by Spiegel and O'Donnell.