Kleene Algebras #
This file defines idempotent semirings and Kleene algebras, which are used extensively in the theory of computation.
An idempotent semiring is a semiring whose addition is idempotent. An idempotent semiring is
naturally a semilattice by setting a ≤ b
if a + b = b
.
A Kleene algebra is an idempotent semiring equipped with an additional unary operator ∗
, the
Kleene star.
Main declarations #
IdemSemiring
: Idempotent semiringIdemCommSemiring
: Idempotent commutative semiringKleeneAlgebra
: Kleene algebra
Notation #
a∗
is notation for kstar a
in locale Computability
.
References #
- [D. Kozen, A completeness theorem for Kleene algebras and the algebra of regular events] [Koz94]
- https://planetmath.org/idempotentsemiring
- https://encyclopediaofmath.org/wiki/Idempotent_semi-ring
- https://planetmath.org/kleene_algebra
TODO #
Instances for AddOpposite
, MulOpposite
, ULift
, Subsemiring
, Subring
, Subalgebra
.
Tags #
kleene algebra, idempotent semiring
An idempotent commutative semiring is a commutative semiring with the additional property that addition is idempotent.
Instances
The Kleene star operator on a Kleene algebra
Equations
- Computability.«term_∗» = Lean.ParserDescr.trailingNode `Computability.«term_∗» 1024 1024 (Lean.ParserDescr.symbol "∗")
Instances For
A Kleene Algebra is an idempotent semiring with an additional unary operator kstar
(for Kleene
star) that satisfies the following properties:
1 + a * a∗ ≤ a∗
1 + a∗ * a ≤ a∗
- If
a * c + b ≤ c
, thena∗ * b ≤ c
- If
c * a + b ≤ c
, thenb * a∗ ≤ c
- add : α → α → α
- zero : α
- mul : α → α → α
- one : α
- sup : α → α → α
- bot : α
- kstar : α → α
- mul_kstar_le_self (a b : α) : b * a ≤ b → b * KStar.kstar a ≤ b
- kstar_mul_le_self (a b : α) : a * b ≤ b → KStar.kstar a * b ≤ b
Instances
Equations
- IdemSemiring.toOrderBot = { bot := IdemSemiring.bot, bot_le := ⋯ }
Construct an idempotent semiring from an idempotent addition.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Alias of the reverse direction of add_eq_left_iff_le
.
Alias of the reverse direction of add_eq_right_iff_le
.
Equations
- IdemSemiring.toOrderedAddCommMonoid = { toAddCommMonoid := NonUnitalNonAssocSemiring.toAddCommMonoid, toPartialOrder := SemilatticeSup.toPartialOrder, add_le_add_left := ⋯ }
Equations
- Prod.instIdemSemiring = { toSemiring := Prod.instSemiring, toSemilatticeSup := Prod.instSemilatticeSup α β, add_eq_sup := ⋯, bot := ⊥, bot_le := ⋯ }
Equations
- Prod.instIdemCommSemiring = { toCommSemiring := Prod.instCommSemiring, toSemilatticeSup := IdemSemiring.toSemilatticeSup, add_eq_sup := ⋯, bot := IdemSemiring.bot, bot_le := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Equations
- Pi.instIdemSemiring = { toSemiring := Pi.semiring, toSemilatticeSup := Pi.instSemilatticeSup, add_eq_sup := ⋯, bot := ⊥, bot_le := ⋯ }
Equations
- Pi.instIdemCommSemiringForall = { toCommSemiring := Pi.commSemiring, toSemilatticeSup := IdemSemiring.toSemilatticeSup, add_eq_sup := ⋯, bot := IdemSemiring.bot, bot_le := ⋯ }
Equations
- One or more equations did not get rendered due to their size.
Pullback an IdemSemiring
instance along an injective function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pullback an IdemCommSemiring
instance along an injective function.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Pullback a KleeneAlgebra
instance along an injective function.
Equations
- One or more equations did not get rendered due to their size.