To see a full list of all the modules go to:
open import Everything
For my bachelor thesis I am implementing categorical notions of partiality in agda using the agda-categories library. The repo for this project can be found here. This is an implementation of this paper by Sergey Goncharov: arxiv
The work takes place in an ambient category that fits our needs:
open import Category.Ambient using (Ambient)
We start out by formalizing Capretta’s Delay Monad in a categorical setting and then proving that it is strong and commutative.
open import Monad.Instance.Delay open import Monad.Instance.Delay.Strong open import Monad.Instance.Delay.Commutative
The next step is to quotient the delay monad by weak bisimilarity, but this quotiented structure is not monadic, so we postulate conditions under which it extends to a monad.
open import Monad.Instance.Delay.Quotienting
The goal of this last section is to formalize an equational lifting monad K that generalizes both the maybe monad and the delay monad.
To do so we first need to look at iteration structures, i.e. functor algebras with an iteration operator, these are called Elgot Algebras.
open import Algebra.Elgot
Afterwards we also introduce categories of iteration algebras with iteration preserving morphisms:
open import Category.Construction.ElgotAlgebras
Free Elgot algebras are free objects in the category of Elgot algebras, we will be needing a notion of stability for them:
open import Algebra.Elgot.Free open import Algebra.Elgot.StableIn a CCC stability follows directly
open import Algebra.Elgot.Properties
With this in hand we can now define KX as a free Elgot algebra and proof that under stability it is a strong monad.
open import Monad.Instance.K open import Monad.Instance.K.Strong
and with this we can show that K is and equational lifting monad, i.e. a commutative monad satisfying the equational lifting law:
open import Monad.Instance.K.Commutative open import Monad.Instance.K.EquationalLifting
and lastly we formalize the notion of pre-Elgot monad and show that K is the initial (strong) pre-Elgot monad.
open import Monad.PreElgot open import Category.Construction.PreElgotMonads open import Category.Construction.StrongPreElgotMonads open import Monad.Instance.K.PreElgot open import Monad.Instance.K.StrongPreElgot
Lastly we do a case study on the category of setoids.
First we look at the (quotiented delay monad) on setoids:
open import Monad.Instance.Setoids.Delay
and then we show that it is an instance of K on setoids:
open import Monad.Instance.Setoids.K