This is the Base.Algebras.Basic module of the Agda Universal Algebra Library.
{-# OPTIONS --without-K --exact-split --safe #-} open import Overture using ( π ; π₯ ; Signature ) module Base.Algebras.Basic {π : Signature π π₯ } where -- Imports from the Agda (Builtin) and the Agda Standard Library -------------- open import Agda.Primitive using () renaming ( Set to Type ; lzero to ββ ) open import Data.Product using ( _,_ ; _Γ_ ; Ξ£-syntax ) open import Level using ( Level ; _β_ ; suc ) open import Relation.Binary using ( IsEquivalence ) renaming ( Rel to BinRel ) open import Relation.Unary using ( _β_ ; Pred ) -- Imports from the Agda Universal Algebra Library ---------------------------- open import Overture using ( β£_β£ ; β₯_β₯ ; Op ) open import Base.Relations using ( _|:_ ; _|:pred_ ; Rel ; compatible-Rel ) using ( REL ; compatible-REL ) private variable Ξ± Ξ² Ο : Level
Our first goal is to develop a working vocabulary and formal library for classical
(single-sorted, set-based) universal algebra. In this section we define the main
objects of study. An algebraic structure (or algebra) in the signature
π = (πΉ, Ο)
is denoted by π¨ = (A, Fᴬ)
and consists of
A
:= a nonempty set (or type), called the domain (or carrier or
universe) of the algebra;Fᴬ := { fᴬ β£ f β F, : (Οf β A) β A }
, a collection of operations on π¨
;π¨
.Note that to each operation symbol f β πΉ
corresponds an operation
fᴬ
on π¨
of arity Οf
; we call such fᴬ
the interpretation of the symbol
f
in the algebra π¨
. We call an algebra in the signature π
an π
-algebra.
An algebra is called finite if it has a finite domain, and is called trivial
if its universe is a singleton. Given two algebras π¨
and π©
, we say that π©
is a reduct of π¨
if both algebras have the same domain and π©
can be obtained
from π¨
by simply removing some of the operations.
Recall, we defined the type Signature π π₯
above as the dependent pair type
Ξ£ F κ Type π , (F β Type π₯)
, and the type Op
of operation symbols is the
function type Op I A = (I β A) β A
(see Base.Relations.Discrete).
For a fixed signature π : Signature π π₯
and universe level Ξ±
, we define the
type of algebras in the signature π
(or type of π
-algebras) with domain
type Type Ξ±
as follows.
Algebra : (Ξ± : Level) β Type (π β π₯ β suc Ξ±) Algebra Ξ± = Ξ£[ A β Type Ξ± ] -- the domain β (f : β£ π β£) β Op A (β₯ π β₯ f) -- the basic operations
It would be more precise to refer to inhabitants of this type as β-algebras because their domains can be of arbitrary type and need not be truncated at some level and, in particular, need to be a set. (See Base.Equality.Truncation.)
We might take this opportunity to define the type of 0-algebras, that is, algebras whose domains are sets, which is probably closer to what most of us think of when doing informal universal algebra. However, in the agda-algebras library we will only need to know that the domains of certain algebras are sets in a few places, so it seems preferable to work with general (β-)algebras throughout and then explicitly postulate additional axioms (e.g., uniquness of identity proofs if and only if necessary.
A popular way to represent algebraic structures in type theory is with record
types. The Sigma type defined above provides an equivalent alternative that we
happen to prefer and we use it throughout the library, both for consistency and
because of its direct connection to the existential quantifier of logic. Recall
that the type Ξ£ x κ X , P x
represents the proposition, βthere exists x
in X
such that P x
holds;β in symbols, β x β X , P x
. Indeed, an inhabitant of Ξ£
x κ X , P x
is a pair (x , p)
such that x
inhabits X
and p
is a proof of
P x
. In other terms, the pair (x , p)
is a witness and proof of the
proposition β x β X , P x
.
Nonetheless, for those who prefer to represent algebraic structures in type theory using records, we offer the following definition (which is equivalent to the Sigma type formulation).
record algebra (Ξ± : Level) : Type(suc(π β π₯ β Ξ±)) where constructor mkalg field carrier : Type Ξ± opsymbol : (f : β£ π β£) β ((β₯ π β₯ f) β carrier) β carrier
This representation of algebras as inhabitants of a record type is equivalent to the representation using Sigma types in the sense that a bi-implication between the two representations is obvious.
open algebra algebraβAlgebra : algebra Ξ± β Algebra Ξ± algebraβAlgebra π¨ = (carrier π¨ , opsymbol π¨) Algebraβalgebra : Algebra Ξ± β algebra Ξ± Algebraβalgebra π¨ = mkalg β£ π¨ β£ β₯ π¨ β₯
We now define a convenient shorthand for the interpretation of an operation symbol. This looks more similar to the standard notation one finds in the literature as compared to the double bar notation we started with, so we will use this new notation almost exclusively in the remaining modules of the agda-algebras library.
_Μ_ : (π : β£ π β£)(π¨ : Algebra Ξ±) β (β₯ π β₯ π β β£ π¨ β£) β β£ π¨ β£ π Μ π¨ = Ξ» π β (β₯ π¨ β₯ π) π
So, if π : β£ π β£
is an operation symbol in the signature π
, and if
π : β₯ π β₯ π β β£ π¨ β£
is a tuple of the appropriate arity, then (π Μ π¨) π
denotes the operation π
interpreted in π¨
and evaluated at π
.
Occasionally we will be given an algebra and we just need to know the universe level of its domain. The following utility function provides this.
Level-of-Alg : {Ξ± : Level} β Algebra Ξ± β Level Level-of-Alg {Ξ± = Ξ±} _ = π β π₯ β suc Ξ± Level-of-Carrier : {Ξ± : Level} β Algebra Ξ± β Level Level-of-Carrier {Ξ± = Ξ±} _ = Ξ±
Recall, in the section on level lifting and lowering, we described the difficulties one may encounter when working with a noncumulative universe hierarchy. We made a promise to provide some domain-specific level lifting and level lowering methods. Here we fulfill this promise by supplying a couple of bespoke tools designed specifically for our operation and algebra types.
open Level Lift-alg-op : {I : Type π₯} {A : Type Ξ±} β Op A I β (Ξ² : Level) β Op (Lift Ξ² A) I Lift-alg-op f Ξ² = Ξ» x β lift (f (Ξ» i β lower (x i))) Lift-Alg : Algebra Ξ± β (Ξ² : Level) β Algebra (Ξ± β Ξ²) Lift-Alg π¨ Ξ² = Lift Ξ² β£ π¨ β£ , (Ξ» (π : β£ π β£) β Lift-alg-op (π Μ π¨) Ξ²) open algebra Lift-algebra : algebra Ξ± β (Ξ² : Level) β algebra (Ξ± β Ξ²) Lift-algebra π¨ Ξ² = mkalg (Lift Ξ² (carrier π¨)) (Ξ» (f : β£ π β£) β Lift-alg-op ((opsymbol π¨) f) Ξ²)
What makes the Lift-Alg
type so useful for resolving type level errors involving
algebras is the nice properties it possesses. Indeed, the agda-algebras
library contains formal proofs of the following facts.
Lift-Alg
is a homomorphism
(see Base.Homomorphisms.Basic)Lift-Alg
is an algebraic invariant
(see Base.Homomorphisms.Isomorphisms)Lift-Alg
of a subalgebra is a subalgebra
(see Base.Subalgebras.Subalgebras)Lift-Alg
preserves identities)
(see Base.Varieties.EquationalLogic)We now define the function compatible
so that, if π¨
denotes an algebra and R
a binary relation, then compatible π¨ R
will represent the assertion that R
is
compatible with all basic operations of π¨
. The formal definition is immediate
since all the work is done by the relation |:
, which we defined above (see
Base.Relations.Discrete).
compatible : (π¨ : Algebra Ξ±) β BinRel β£ π¨ β£ Ο β Type (π β π₯ β Ξ± β Ο) compatible π¨ R = β π β (π Μ π¨) |: R compatible-pred : (π¨ : Algebra Ξ±) β Pred (β£ π¨ β£ Γ β£ π¨ β£)Ο β Type (π β π₯ β Ξ± β Ο) compatible-pred π¨ P = β π β (π Μ π¨) |:pred P
Recall, the |:
type was defined in Base.Relations.Discrete module.
In the Base.Relations.Continuous module, we defined a function called
compatible-Rel
to represent the assertion that a given continuous relation is
compatible with a given operation. With that, it is easy to define a function,
which we call compatible-Rel-alg
, representing compatibility of a continuous
relation with all operations of an algebra. Similarly, we define the analogous
compatible-REL-alg
function for the (even more general) type of dependent
relations.
module _ {I : Type π₯} where compatible-Rel-alg : (π¨ : Algebra Ξ±) β Rel β£ π¨ β£ I{Ο} β Type(π β Ξ± β π₯ β Ο) compatible-Rel-alg π¨ R = β (π : β£ π β£ ) β compatible-Rel (π Μ π¨) R compatible-REL-alg : (π : I β Algebra Ξ±) β REL I (Ξ» i β β£ π i β£) {Ο} β Type _ compatible-REL-alg π R = β ( π : β£ π β£ ) β compatible-REL (Ξ» i β π Μ (π i)) R