open import Cat.Functor.Adjoint
open import Cat.Functor.Base
open import Cat.Prelude

open Functor
open _=>_

module Cat.Diagram.Monad {o h : _} (C : Precategory o h) where

import Cat.Reasoning C as C

Monads🔗

A monad on a category C\ca{C} is one way of categorifying the concept of monoid. Specifically, rather than living in a monoidal category, a monad lives in a bicategory. Here, we concern ourselves with the case of monads in the bicategory of categories, so that we may say: A monad is an endofunctor MM, equipped with a unit natural transformation IdM\id{Id} \To M, and a multiplication (MM)M(M \circ M) \To M.

record Monad : Type (o  h) where
  no-eta-equality
  field
    M    : Functor C C
    unit : Id => M
    mult : (M F∘ M) => M

  module unit = _=>_ unit
  module mult = _=>_ mult

  M₀ = F₀ M
  M₁ = F₁ M
  M-id = F-id M
  M-∘ = F-∘ M

Furthermore, these natural transformations must satisfy identity and associativity laws exactly analogous to those of a monoid.

  field
    left-ident  :  {x}  mult.η x C.∘ M₁ (unit.η x)  C.id
    right-ident :  {x}  mult.η x C.∘ unit.η (M₀ x)  C.id

  field
    mult-assoc :  {x}  mult.η x C.∘ M₁ (mult.η x)  mult.η x C.∘ mult.η (M₀ x)

Algebras over a monad🔗

One way of interpreting a monad MM is as giving a signature for an algebraic theory. For instance, the free monoid monad describes the signature for the theory of monoids, and the free group monad describes the theory of groups.

Under this light, an algebra over a monad is a way of evaluating the abstract operations given by a monadic expression to a concrete value. Formally, an algebra for MM is given by a choice of object AA and a morphism ν:M(A)A\nu : M(A) \to A.

record Algebra-on (M : Monad) (ob : C.Ob) : Type (o  h) where
  no-eta-equality
  open Monad M

  field
    ν : C.Hom (M₀ ob) ob

This morphism must satisfy equations categorifying those which define a monoid action. If we think of MM as specifying a signature of effects, then v-unit says that the unit has no effects, and v-mult says that, given two layers M(M(A))M(M(A)), it doesn’t matter whether you first join then evaluate, or evaluate twice.

    ν-unit : ν C.∘ unit.η ob  C.id
    ν-mult : ν C.∘ M₁ ν  ν C.∘ mult.η ob

Algebra : Monad  Type (o  h)
Algebra M = Σ (Algebra-on M)

Eilenberg-Moore Category🔗

If we take a monad MM as the signature of an (algebraic) theory, and MM-algebras as giving models of that theory, then we can ask (like with everything in category theory): Are there maps between interpretations? The answer (as always!) is yes: An algebra homomorphism is a map of the underlying objects which “commutes with the algebras”.

record Algebra-hom (M : Monad) (X Y : Algebra M) : Type (o  h) where
  no-eta-equality
  private
    module X = Algebra-on (X .snd)
    module Y = Algebra-on (Y .snd)

  open Monad M

  field
    morphism : C.Hom (X .fst) (Y .fst)
    commutes : morphism C.∘ X.ν  Y.ν C.∘ M₁ morphism

open Algebra-hom

We can be more specific about “commuting with the algebras” by drawing a square: A map m:XYm : X \to Y in the ambient category is a homomorphism of MM-algebras when the square below commutes.

Since commutes is an identification between morphisms, it inhabits a proposition (because Hom-sets are sets), equality of algebra homomorphisms only depends on an equality of their underlying morphisms.

Algebra-hom-path
  : {M : Monad} {X Y : Algebra M} {F G : Algebra-hom M X Y}
   morphism F  morphism G
   F  G
Algebra-hom-path x i .morphism = x i
Algebra-hom-path {M = M} {X} {Y} {F} {G} x i .commutes =
  is-prop→pathp  i  C.Hom-set _ _ (x i C.∘ X .snd .Algebra-on.ν)
                                     (Y .snd .Algebra-on.ν C.∘ Monad.M₁ M (x i)))
    (F .commutes) (G .commutes) i

Since the square we drew above commutes for the identity morphism, and we can show that the composite of two algebra homomorphisms is an algebra homomorphism, they assemble into a category: The Eilenberg-Moore category of MM.

module _ (M : Monad) where
  private
    module M = Monad M
  open M hiding (M)
  open Precategory
  open Algebra-on

  Eilenberg-Moore : Precategory _ _
  Eilenberg-Moore .Ob = Algebra M
  Eilenberg-Moore .Hom X Y = Algebra-hom M X Y

Defining the identity and composition maps is mostly an exercise in categorical yoga:

  Eilenberg-Moore .id {o , x} = record
    { morphism = C.id
    ; commutes = C.id C.∘ ν x     ≡⟨ C.id-comm-sym 
                 ν x C.∘ C.id     ≡⟨ ap (C._∘_ _) (sym M-id) 
                 ν x C.∘ M₁ C.id  
    }

  Eilenberg-Moore ._∘_ {_ , x} {_ , y} {_ , z} F G = record
    { morphism = morphism F C.∘ morphism G
    ; commutes = (morphism F C.∘ morphism G) C.∘ ν x            ≡⟨ C.extendr (commutes G) 
                  (morphism F C.∘ ν y) C.∘ M₁ (morphism G)       ≡⟨ ap₂ C._∘_ (commutes F) refl 
                  (ν z C.∘ M₁ (morphism F)) C.∘ M₁ (morphism G)  ≡⟨ C.pullr (sym (M-∘ _ _)) 
                  ν z C.∘ M₁ (morphism F C.∘ morphism G)         
    }
Because we have characterised equality of algebra homomorphisms as equality of their underlying maps, the Eilenberg-Moore category inherits the identity and associativity laws from its underlying category.
  Eilenberg-Moore .idr f = Algebra-hom-path (C.idr (morphism f))
  Eilenberg-Moore .idl f = Algebra-hom-path (C.idl (morphism f))
  Eilenberg-Moore .assoc f g h = Algebra-hom-path (C.assoc _ _ _)
  Eilenberg-Moore .Hom-set X Y = is-hlevel≃ 2 (Iso→Equiv eqv e⁻¹) (hlevel 2)
    where open C.HLevel-instance

By projecting the underlying object of the algebras, and the underlying morphisms of the homomorphisms between them, we can define a functor from Eilenberg-Moore back to the underlying category:

  Forget : Functor Eilenberg-Moore C
  Forget .F₀ = fst
  Forget .F₁ = Algebra-hom.morphism
  Forget .F-id = refl
  Forget .F-∘ f g = refl

The lemma Algebra-hom-path says exactly that this functor is faithful.

  Forget-is-faithful : is-faithful Forget
  Forget-is-faithful = Algebra-hom-path

Free Algebras🔗

In exactly the same way that we may construct a free group by taking the inhabitants of some set XX as generating the “words” of a group, we can, given an object AA of the underlying category, build a free MM-algebra on AA. Keeping with our interpretation of monads as logical signatures, this is the syntactic model of MM, with a set of “neutrals” chosen from the object AA.

This construction is a lot simpler to do in generality than in any specific case: We can always turn AA into an MM-algebra by taking the underlying object to be M(A)M(A), and the algebra map to be the monadic multiplication; The associativity and unit laws of the monad itself become those of the MM-action.

  Free : Functor C Eilenberg-Moore
  Free .F₀ A = M₀ A ,
    record
      { ν      = mult.η A
      ; ν-mult = mult-assoc
      ; ν-unit = right-ident
      }

The construction of free MM-algebras is furthermore functorial on the underlying objects; Since the monadic multiplication is a natural transformation MMMM\circ M \To M, the naturality condition (drawn below) doubles as showing that the functorial action of MM can be taken as an algebraic action:

  Free .F₁ f = record
    { morphism = M₁ f
    ; commutes = sym (mult.is-natural _ _ _)
    }
  Free .F-id = Algebra-hom-path M-id
  Free .F-∘ f g = Algebra-hom-path (M-∘ f g)

This is a free construction in the precise sense of the word: it’s the left adjoint to the functor Forget, so in particular it provides a systematic, universal way of mapping from C\ca{C} to CM\ca{C}^M.

  open _⊣_

  Free⊣Forget : Free  Forget
  Free⊣Forget .unit = NT M.unit.η M.unit.is-natural
  Free⊣Forget .counit .η x =
    record { morphism = x .snd .ν
           ; commutes = sym (x .snd .ν-mult)
           }
  Free⊣Forget .counit .is-natural x y f =
    Algebra-hom-path (sym (commutes f))
  Free⊣Forget .zig = Algebra-hom-path left-ident
  Free⊣Forget .zag {x} = x .snd .ν-unit