↑ Top


Basic definitions

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

Algebras

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

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.

Algebras as record types

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 ∣ 𝑨 ∣ βˆ₯ 𝑨 βˆ₯

Operation interpretation syntax

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 π‘Ž.

The universe level of an algebra

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 {Ξ± = Ξ±} _ = Ξ±

Level lifting algebra types

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.

Compatibility of binary relations

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.

Compatibility of continuous relations

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

↑ Base.Algebras Base.Algebras.Products β†’