├── .gitignore ├── Isomorphism.agda ├── Categories ├── Sets.agda ├── Families.agda ├── Initial.agda ├── Terminal.agda ├── CoProducts.agda ├── PushOuts.agda ├── Products.agda └── Setoids.agda ├── FunctorCat.agda ├── StrongArrows.agda ├── Monoids.agda ├── RMonads ├── CatofRMonads.agda ├── RKleisli.agda ├── CatofRAdj │ ├── TermRAdjObj.agda │ ├── TermRAdj.agda │ ├── InitRAdj.agda │ └── TermRAdjHom.agda ├── RKleisli │ ├── Adjunction.agda │ └── Functors.agda ├── Restriction.agda ├── REM │ ├── Functors.agda │ └── Adjunction.agda ├── Modules.agda ├── SpecialCase.agda ├── RMonadMorphs.agda ├── REM.agda └── CatofRAdj.agda ├── StrongRMonads.agda ├── Monads ├── MonadMorphs.agda ├── CatofAdj │ ├── TermAdj.agda │ ├── TermAdjUniq.agda │ ├── TermAdjHom.agda │ ├── TermAdjObj.agda │ └── InitAdj.agda ├── Kleisli.agda ├── Kleisli │ ├── Functors.agda │ └── Adjunction.agda ├── EM │ ├── Functors.agda │ └── Adjunction.agda ├── EM.agda └── CatofAdj.agda ├── Categories.agda ├── Functors ├── FullyFaithful.agda └── Fin.agda ├── LICENSE ├── Monads.agda ├── Adjunctions.agda ├── YonedaLemma.agda ├── DayConvolution.agda ├── RAdjunctions.agda ├── Vectors.agda ├── Functors.agda ├── README.md ├── Adjunctions └── Adj2Mon.agda ├── RAdjunctions └── RAdj2RMon.agda ├── RMonads.agda ├── Everything.agda ├── WellTypedTermsModel.agda ├── Naturals.agda ├── WeakArrows.agda ├── MonoidalCat.agda ├── Library.agda ├── Lawvere.agda ├── WellScopedTermsModel.agda ├── WellTypedTermsNBEModel.agda ├── WellScopedTerms.agda ├── WellTypedTerms.agda └── Delay.agda /.gitignore: -------------------------------------------------------------------------------- 1 | *.agdai 2 | *.agda~ -------------------------------------------------------------------------------- /Isomorphism.agda: -------------------------------------------------------------------------------- 1 | module Isomorphism where 2 | 3 | open import Library 4 | 5 | record Iso {a b}(A : Set a)(B : Set b) : Set (a ⊔ b) where 6 | field fun : A → B 7 | inv : B → A 8 | law1 : ∀ b → fun (inv b) ≅ b 9 | law2 : ∀ a → inv (fun a) ≅ a 10 | -------------------------------------------------------------------------------- /Categories/Sets.agda: -------------------------------------------------------------------------------- 1 | module Categories.Sets where 2 | 3 | open import Library 4 | open import Categories 5 | 6 | Sets : ∀{l} → Cat 7 | Sets {l} = record{ 8 | Obj = Set l; 9 | Hom = λ X Y → X → Y; 10 | iden = id; 11 | comp = λ f g → f ∘ g; 12 | idl = refl; 13 | idr = refl; 14 | ass = refl} 15 | -------------------------------------------------------------------------------- /Categories/Families.agda: -------------------------------------------------------------------------------- 1 | module Categories.Families where 2 | 3 | open import Library 4 | open import Categories 5 | 6 | open Cat 7 | 8 | Fam : Set → Cat 9 | Fam I = record { 10 | Obj = I → Set; 11 | Hom = λ A B → ∀ {i} → A i → B i; 12 | iden = id; 13 | comp = λ f g → f ∘ g; 14 | idl = refl; 15 | idr = refl; 16 | ass = refl} 17 | -------------------------------------------------------------------------------- /FunctorCat.agda: -------------------------------------------------------------------------------- 1 | module FunctorCat where 2 | 3 | open import Categories 4 | open import Functors 5 | open import Naturals 6 | 7 | FunctorCat : ∀{a b c d} → Cat {a}{b} → Cat {c}{d} → Cat 8 | FunctorCat C D = record{ 9 | Obj = Fun C D; 10 | Hom = NatT; 11 | iden = idNat; 12 | comp = compNat; 13 | idl = idlNat; 14 | idr = idrNat; 15 | ass = λ{_}{_}{_}{_}{α}{β}{η} → assNat {α = α}{β}{η}} 16 | -------------------------------------------------------------------------------- /Categories/Initial.agda: -------------------------------------------------------------------------------- 1 | module Categories.Initial where 2 | 3 | open import Library 4 | open import Categories 5 | open import Categories.Sets 6 | open Cat 7 | 8 | record Init {a b} (C : Cat {a}{b})(I : Obj C) : Set (a ⊔ b) where 9 | constructor init 10 | field i : ∀{X} → Hom C I X 11 | law : ∀{X}{f : Hom C I X} → i {X} ≅ f 12 | 13 | ZeroSet : Init Sets ⊥ 14 | ZeroSet = record {i = λ(); law = ext λ()} 15 | -------------------------------------------------------------------------------- /Categories/Terminal.agda: -------------------------------------------------------------------------------- 1 | module Categories.Terminal where 2 | 3 | open import Library 4 | open import Categories 5 | open import Categories.Sets 6 | open Cat 7 | 8 | record Term {a b} (C : Cat {a}{b})(T : Obj C) : Set (a ⊔ b) where 9 | constructor term 10 | field t : ∀{X} → Hom C X T 11 | law : ∀{X}{f : Hom C X T} → t {X} ≅ f 12 | 13 | OneSet : Term Sets ⊤ 14 | OneSet = record {t = λ _ → _; law = ext (λ _ → refl)} 15 | -------------------------------------------------------------------------------- /StrongArrows.agda: -------------------------------------------------------------------------------- 1 | module StrongArrows where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | open import MonoidalCat 7 | open import WeakArrows 8 | 9 | record SArrow {l m}(J : Monoidal {l}{m}) : Set (lsuc (l ⊔ m)) where 10 | constructor sarrow 11 | open Monoidal J 12 | open Cat C 13 | open Fun 14 | field A : Arrow C 15 | open Arrow A 16 | field fst' : ∀{X X' Y} -> R X' X -> R (OMap ⊗ (X' , Y)) (OMap ⊗ (X , Y)) 17 | -- laws later 18 | 19 | -------------------------------------------------------------------------------- /Monoids.agda: -------------------------------------------------------------------------------- 1 | module Monoids where 2 | 3 | open import Library 4 | 5 | record Monoid {a} : Set (lsuc a) where 6 | field S : Set a 7 | ε : S 8 | _•_ : S → S → S 9 | lid : ∀{m} → ε • m ≅ m 10 | rid : ∀{m} → m • ε ≅ m 11 | ass : ∀{m n o} → (m • n) • o ≅ m • (n • o) 12 | 13 | infix 10 _•_ 14 | 15 | Nat+Mon : Monoid 16 | Nat+Mon = record { 17 | S = ℕ; 18 | ε = zero; 19 | _•_ = _+_; 20 | lid = refl; 21 | rid = ≡-to-≅ $ +-right-identity _; 22 | ass = λ{m} → ≡-to-≅ $ +-assoc m _ _} 23 | 24 | -------------------------------------------------------------------------------- /RMonads/CatofRMonads.agda: -------------------------------------------------------------------------------- 1 | module RMonads.CatofRMonads where 2 | 3 | open import Categories 4 | open import Functors 5 | open import RMonads 6 | open import RMonads.RMonadMorphs 7 | 8 | CatofRMonads : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}(J : Fun C D) → Cat 9 | CatofRMonads J = record 10 | { Obj = RMonad J 11 | ; Hom = RMonadMorph 12 | ; iden = IdRMonadMorph _ 13 | ; comp = CompRMonadMorph 14 | ; idl = idl _ 15 | ; idr = idr _ 16 | ; ass = λ{_ _ _ _ f g h} → ass f g h 17 | } 18 | -------------------------------------------------------------------------------- /StrongRMonads.agda: -------------------------------------------------------------------------------- 1 | module StrongRMonads where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | open import MonoidalCat 7 | open import RMonads 8 | 9 | record SRMonad {a b c d}{M : Monoidal {a}{b}}{M' : Monoidal {c}{d}} 10 | (J : MonoidalFun M M') : Set (a ⊔ b ⊔ c ⊔ d) where 11 | constructor srmonad 12 | open MonoidalFun J 13 | field RM : RMonad F 14 | open RMonad RM 15 | open Cat 16 | open Monoidal 17 | open Fun 18 | field st : ∀{X Y} -> Hom (C M') (OMap (⊗ M') (T X , OMap F Y)) 19 | (T (OMap (⊗ M) (X , Y))) 20 | -- laws later 21 | 22 | -------------------------------------------------------------------------------- /Monads/MonadMorphs.agda: -------------------------------------------------------------------------------- 1 | module Monads.MonadMorphs where 2 | 3 | open import Library 4 | open import Functors 5 | open import Categories 6 | open import Monads 7 | 8 | open Fun 9 | open Monad 10 | 11 | record MonadMorph {a b}{C : Cat {a}{b}}(M M' : Monad C) : Set (a ⊔ b) where 12 | constructor monadmorph 13 | open Cat C 14 | field morph : ∀ {X} → Hom (T M X) (T M' X) 15 | lawη : ∀ {X} → comp morph (η M {X}) ≅ η M' {X} 16 | lawbind : ∀ {X Y}{k : Hom X (T M Y)} → 17 | comp (morph {Y}) (bind M k) 18 | ≅ 19 | comp (bind M' (comp (morph {Y}) k)) (morph {X}) 20 | 21 | -------------------------------------------------------------------------------- /Monads/CatofAdj/TermAdj.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.CatofAdj.TermAdj {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open import Library 7 | open import Functors 8 | open import Monads.CatofAdj M 9 | open import Categories.Terminal 10 | open import Monads.CatofAdj.TermAdjObj M 11 | open import Monads.CatofAdj.TermAdjHom M 12 | open import Monads.CatofAdj.TermAdjUniq M 13 | 14 | EMIsTerm : Term CatofAdj EMObj 15 | EMIsTerm = record { 16 | t = λ {A} → EMHom A; 17 | law = λ {A} {V} → 18 | HomAdjEq _ _ (FunctorEq _ _ (omaplem A V) 19 | (iext λ _ → iext λ _ → ext $ hmaplem A V))} 20 | -------------------------------------------------------------------------------- /Categories.agda: -------------------------------------------------------------------------------- 1 | module Categories where 2 | 3 | open import Library 4 | 5 | record Cat {a b} : Set (lsuc (a ⊔ b)) where 6 | field Obj : Set a 7 | Hom : Obj → Obj → Set b 8 | iden : ∀{X} → Hom X X 9 | comp : ∀{X Y Z} → Hom Y Z → Hom X Y → Hom X Z 10 | idl : ∀{X Y}{f : Hom X Y} → comp iden f ≅ f 11 | idr : ∀{X Y}{f : Hom X Y} → comp f iden ≅ f 12 | ass : ∀{W X Y Z}{f : Hom Y Z}{g : Hom X Y}{h : Hom W X} → 13 | comp (comp f g) h ≅ comp f (comp g h) 14 | open Cat 15 | 16 | _Op : ∀{a b} → Cat {a}{b} → Cat 17 | C Op = record{ 18 | Obj = Obj C; 19 | Hom = λ X Y → Hom C Y X; 20 | iden = iden C; 21 | comp = λ f g → comp C g f; 22 | idl = idr C; 23 | idr = idl C; 24 | ass = sym (ass C)} 25 | -------------------------------------------------------------------------------- /Monads/Kleisli.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.Kleisli {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open import Library 7 | open Cat C 8 | open Monad M 9 | 10 | Kl : Cat 11 | Kl = record{ 12 | Obj = Obj; 13 | Hom = λ X Y → Hom X (T Y); 14 | iden = η; 15 | comp = λ f g → comp (bind f) g; 16 | idl = λ{X}{Y}{f} → 17 | proof 18 | comp (bind η) f 19 | ≅⟨ cong (λ g → comp g f) law1 ⟩ 20 | comp iden f 21 | ≅⟨ idl ⟩ 22 | f 23 | ∎; 24 | idr = law2; 25 | ass = λ{_}{_}{_}{_}{f}{g}{h} → 26 | proof 27 | comp (bind (comp (bind f) g)) h 28 | ≅⟨ cong (λ f → comp f h) law3 ⟩ 29 | comp (comp (bind f) (bind g)) h 30 | ≅⟨ ass ⟩ 31 | comp (bind f) (comp (bind g) h) 32 | ∎} 33 | -------------------------------------------------------------------------------- /Monads/Kleisli/Functors.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.Kleisli.Functors {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open import Library 7 | open import Functors 8 | open import Monads.Kleisli M 9 | 10 | open Cat C 11 | open Fun 12 | open Monad M 13 | 14 | KlL : Fun C Kl 15 | KlL = record{ 16 | OMap = id; 17 | HMap = comp η; 18 | fid = idr; 19 | fcomp = λ{_}{_}{_}{f}{g} → 20 | proof 21 | comp η (comp f g) 22 | ≅⟨ sym ass ⟩ 23 | comp (comp η f) g 24 | ≅⟨ cong (λ f → comp f g) (sym law2) ⟩ 25 | comp (comp (bind (comp η f)) η) g 26 | ≅⟨ ass ⟩ 27 | comp (bind (comp η f)) (comp η g) 28 | ∎} 29 | 30 | KlR : Fun Kl C 31 | KlR = record{ 32 | OMap = T; 33 | HMap = bind; 34 | fid = law1; 35 | fcomp = law3} 36 | -------------------------------------------------------------------------------- /Categories/CoProducts.agda: -------------------------------------------------------------------------------- 1 | module Categories.CoProducts where 2 | 3 | open import Data.Sum hiding ([_,_]) 4 | open import Library hiding (_+_ ; _,_) 5 | open import Categories 6 | 7 | record CoProd {l m}(C : Cat {l}{m}) : Set (m ⊔ l) where 8 | open Cat C 9 | field _+_ : Obj -> Obj -> Obj 10 | inl : ∀{A B} -> Hom A (A + B) 11 | inr : ∀{A B} -> Hom B (A + B) 12 | [_,_] : ∀{A B C} -> Hom A C -> Hom B C -> Hom (A + B) C 13 | law1 : ∀{A B C}(f : Hom A C)(g : Hom B C) → 14 | comp [ f , g ] inl ≅ f 15 | law2 : ∀{A B C}(f : Hom A C)(g : Hom B C) → 16 | comp [ f , g ] inr ≅ g 17 | law3 : ∀{A B C}(f : Hom A C)(g : Hom B C) 18 | (h : Hom (A + B) C) → 19 | comp h inl ≅ f → comp h inr ≅ g → h ≅ [ f , g ] 20 | -------------------------------------------------------------------------------- /Functors/FullyFaithful.agda: -------------------------------------------------------------------------------- 1 | module Functors.FullyFaithful where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | open import Naturals hiding (Iso) 7 | open import Isomorphism 8 | 9 | open Cat 10 | open Fun 11 | 12 | FullyFaithful : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 13 | (F : Fun C D) → Set (a ⊔ b ⊔ d) 14 | FullyFaithful {C = C}{D = D} F = 15 | ∀ (X Y : Obj C) → Iso (Hom D (OMap F X) (OMap F Y)) (Hom C X Y) 16 | 17 | Full : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 18 | (F : Fun C D) → Set _ 19 | Full {C = C}{D = D} F = ∀ {X Y}(h : Hom D (OMap F X) (OMap F Y)) → 20 | Σ (Hom C X Y) \ f -> HMap F f ≅ h 21 | 22 | Faithful : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 23 | (F : Fun C D) → Set _ 24 | Faithful {C = C}{D = D} F = ∀{X Y}{f g : Hom C X Y} → HMap F f ≅ HMap F g → f ≅ g 25 | -------------------------------------------------------------------------------- /RMonads/RKleisli.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | import RMonads 4 | 5 | module RMonads.RKleisli 6 | {a b c d} 7 | {C : Cat {a}{b}} 8 | {D : Cat {c}{d}} 9 | {J : Fun C D} 10 | (M : RMonads.RMonad J) where 11 | 12 | open import Library 13 | open RMonads.RMonad M 14 | open Cat 15 | open Fun 16 | 17 | Kl : Cat 18 | Kl = record{ 19 | Obj = Obj C; 20 | Hom = λ X Y → Hom D (OMap J X) (T Y); 21 | iden = η; 22 | comp = λ f g → comp D (bind f) g; 23 | idl = λ{X}{Y}{f} → 24 | proof 25 | comp D (bind η) f 26 | ≅⟨ cong (λ g → comp D g f) law1 ⟩ 27 | comp D (iden D) f 28 | ≅⟨ idl D ⟩ 29 | f 30 | ∎; 31 | idr = law2; 32 | ass = λ{_ _ _ _ f g h} → 33 | proof 34 | comp D (bind (comp D (bind f) g)) h 35 | ≅⟨ cong (λ f → comp D f h) law3 ⟩ 36 | comp D (comp D (bind f) (bind g)) h 37 | ≅⟨ ass D ⟩ 38 | comp D (bind f) (comp D (bind g) h) 39 | ∎} 40 | -------------------------------------------------------------------------------- /Monads/Kleisli/Adjunction.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.Kleisli.Adjunction {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open Cat C 7 | open Monad M 8 | 9 | open import Library 10 | open import Functors 11 | open import Monads.Kleisli M 12 | open import Adjunctions 13 | open import Monads.Kleisli.Functors M 14 | open Fun 15 | 16 | KlAdj : Adj C Kl 17 | KlAdj = record { 18 | L = KlL; 19 | R = KlR; 20 | left = id; 21 | right = id; 22 | lawa = λ _ → refl; 23 | lawb = λ _ → refl; 24 | natleft = lem; 25 | natright = lem} 26 | where 27 | lem = λ {X}{X'}{Y}{Y'} (f : Hom X' X)(g : Hom Y (T Y')) h → 28 | proof 29 | comp (bind g) (comp h f) 30 | ≅⟨ cong (λ h → comp (bind g) (comp h f)) (sym law2) ⟩ 31 | comp (bind g) (comp (comp (bind h) η) f) 32 | ≅⟨ cong (comp (bind g)) ass ⟩ 33 | comp (bind g) (comp (bind h) (comp η f)) 34 | ∎ 35 | -------------------------------------------------------------------------------- /RMonads/CatofRAdj/TermRAdjObj.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | open import RMonads 4 | 5 | module RMonads.CatofRAdj.TermRAdjObj {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 6 | {J : Fun C D}(M : RMonad J) where 7 | 8 | open import Library 9 | open import Naturals 10 | open import RAdjunctions 11 | open import RMonads.CatofRAdj M 12 | open import Categories.Terminal 13 | open import RMonads.REM M 14 | open import RMonads.REM.Adjunction M 15 | open import RAdjunctions.RAdj2RMon 16 | 17 | open Cat 18 | open Fun 19 | open NatT 20 | open RAdj 21 | 22 | lemX : R REMAdj ○ L REMAdj ≅ TFun M 23 | lemX = FunctorEq _ _ refl refl 24 | 25 | EMObj : Obj CatofAdj 26 | EMObj = record { 27 | E = EM; 28 | adj = REMAdj; 29 | law = lemX; 30 | ηlaw = idl D; 31 | bindlaw = λ{X}{Y}{f} → 32 | cong bind 33 | (stripsubst (Hom D (OMap J X)) 34 | f 35 | (fcong Y (cong OMap (sym lemX))))} where open RMonad M 36 | -------------------------------------------------------------------------------- /Categories/PushOuts.agda: -------------------------------------------------------------------------------- 1 | module Categories.PushOuts where 2 | 3 | open import Library 4 | open import Categories 5 | 6 | open Cat 7 | 8 | record Square {a}{b}{C : Cat {a}{b}}{X Y Z}(f : Hom C Z X)(g : Hom C Z Y) : Set (a ⊔ b) where 9 | constructor square 10 | field W : Obj C 11 | h : Hom C X W 12 | k : Hom C Y W 13 | scom : comp C h f ≅ comp C k g 14 | 15 | record SqMap {a}{b}{C : Cat {a}{b}}{X Y Z : Obj C}{f : Hom C Z X}{g : Hom C Z Y} 16 | (sq sq' : Square {a}{b}{C} f g) : Set (a ⊔ b) where 17 | constructor sqmap 18 | open Square 19 | field sqMor : Hom C (W sq) (W sq') 20 | leftTr : comp C sqMor (h sq) ≅ h sq' 21 | rightTr : comp C sqMor (k sq) ≅ k sq' 22 | open SqMap 23 | 24 | record PushOut {a}{b}{C : Cat {a}{b}}{X Y Z}(f : Hom C Z X)(g : Hom C Z Y) : Set (a ⊔ b) where 25 | constructor pushout 26 | field sq : Square {a}{b}{C} f g 27 | uniqPush : (sq' : Square f g) → Σ (SqMap sq sq') 28 | \ u → (u' : SqMap sq sq') → sqMor u ≅ sqMor u' 29 | 30 | -------------------------------------------------------------------------------- /RMonads/RKleisli/Adjunction.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | open import RMonads 4 | 5 | module RMonads.RKleisli.Adjunction {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 6 | {J : Fun C D}(M : RMonad J) where 7 | 8 | open import Library 9 | open import RMonads.RKleisli M 10 | open import RAdjunctions 11 | open import RMonads.RKleisli.Functors M 12 | open Cat 13 | open Fun 14 | open RMonad M 15 | 16 | KlAdj : RAdj J Kl 17 | KlAdj = record{ 18 | L = RKlL; 19 | R = RKlR; 20 | left = id; 21 | right = id; 22 | lawa = λ _ → refl; 23 | lawb = λ _ → refl; 24 | natleft = lem; 25 | natright = lem} 26 | where 27 | lem = λ {X}{X'}{Y}{Y'} (f : Hom C X' X) (g : Hom D (OMap J Y) (T Y')) h → 28 | proof 29 | comp D (bind g) (comp D h (HMap J f)) 30 | ≅⟨ cong (λ h₁ → comp D (bind g) (comp D h₁ (HMap J f))) (sym law2) ⟩ 31 | comp D (bind g) (comp D (comp D (bind h) η) (HMap J f)) 32 | ≅⟨ cong (comp D (bind g)) (ass D) ⟩ 33 | comp D (bind g) (comp D (bind h) (comp D η (HMap J f))) 34 | ∎ 35 | -------------------------------------------------------------------------------- /LICENSE: -------------------------------------------------------------------------------- 1 | MIT License 2 | 3 | Copyright (c) 2019 James Chapman 4 | 5 | Permission is hereby granted, free of charge, to any person obtaining a copy 6 | of this software and associated documentation files (the "Software"), to deal 7 | in the Software without restriction, including without limitation the rights 8 | to use, copy, modify, merge, publish, distribute, sublicense, and/or sell 9 | copies of the Software, and to permit persons to whom the Software is 10 | furnished to do so, subject to the following conditions: 11 | 12 | The above copyright notice and this permission notice shall be included in all 13 | copies or substantial portions of the Software. 14 | 15 | THE SOFTWARE IS PROVIDED "AS IS", WITHOUT WARRANTY OF ANY KIND, EXPRESS OR 16 | IMPLIED, INCLUDING BUT NOT LIMITED TO THE WARRANTIES OF MERCHANTABILITY, 17 | FITNESS FOR A PARTICULAR PURPOSE AND NONINFRINGEMENT. IN NO EVENT SHALL THE 18 | AUTHORS OR COPYRIGHT HOLDERS BE LIABLE FOR ANY CLAIM, DAMAGES OR OTHER 19 | LIABILITY, WHETHER IN AN ACTION OF CONTRACT, TORT OR OTHERWISE, ARISING FROM, 20 | OUT OF OR IN CONNECTION WITH THE SOFTWARE OR THE USE OR OTHER DEALINGS IN THE 21 | SOFTWARE. 22 | -------------------------------------------------------------------------------- /Categories/Products.agda: -------------------------------------------------------------------------------- 1 | module Categories.Products where 2 | 3 | open import Library renaming (_×_ to Pair) hiding ([_]) 4 | open import Categories 5 | 6 | _×_ : ∀{l l' m m'}(C : Cat {l}{m})(D : Cat {l'}{m'}) -> Cat 7 | C × D = record 8 | { Obj = Pair (Obj C) (Obj D) 9 | ; Hom = λ {(X , X') (Y , Y') -> Pair (Hom C X Y) (Hom D X' Y')} 10 | ; iden = iden C , iden D 11 | ; comp = λ {(f , f') (g , g') -> comp C f g , comp D f' g'} 12 | ; idl = cong₂ _,_ (idl C) (idl D) 13 | ; idr = cong₂ _,_ (idr C) (idr D) 14 | ; ass = cong₂ _,_ (ass C) (ass D) 15 | } where open Cat 16 | 17 | infixr 50 _×_ 18 | 19 | record Prod {l}{m}(C : Cat {l}{m})(A : Cat.Obj C)(B : Cat.Obj C) : Set (l ⊔ m) 20 | where 21 | constructor prod 22 | open Cat C 23 | field Pr : Obj 24 | π₁ : Hom Pr A 25 | π₂ : Hom Pr B 26 | [_,_] : ∀{C} → Hom C A → Hom C B → Hom C Pr 27 | law1 : ∀{C}{f : Hom C A}{g} → comp π₁ [ f , g ] ≅ f 28 | law2 : ∀{C}{f : Hom C A}{g} → comp π₂ [ f , g ] ≅ g 29 | law3 : ∀{C}{f : Hom C A}{g : Hom C B}{h : Hom C Pr} → 30 | comp π₁ h ≅ f → comp π₂ h ≅ g → h ≅ [ f , g ] 31 | -------------------------------------------------------------------------------- /RMonads/RKleisli/Functors.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | open import RMonads 4 | 5 | module RMonads.RKleisli.Functors {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 6 | {J : Fun C D}(M : RMonad J) where 7 | 8 | open import Library 9 | open import RMonads.RKleisli M 10 | open import RAdjunctions 11 | 12 | open Cat 13 | open Fun 14 | open RMonad M 15 | 16 | RKlL : Fun C Kl 17 | RKlL = record{ 18 | OMap = id; 19 | HMap = λ f → comp D η (HMap J f); 20 | fid = 21 | proof 22 | comp D η (HMap J (iden C)) 23 | ≅⟨ cong (comp D η) (fid J) ⟩ 24 | comp D η (iden D) 25 | ≅⟨ idr D ⟩ 26 | η ∎; 27 | fcomp = λ{_ _ _ f g} → 28 | proof 29 | comp D η (HMap J (comp C f g)) 30 | ≅⟨ cong (comp D η) (fcomp J) ⟩ 31 | comp D η (comp D (HMap J f) (HMap J g)) 32 | ≅⟨ sym (ass D) ⟩ 33 | comp D (comp D η (HMap J f)) (HMap J g) 34 | ≅⟨ cong (λ f → comp D f (HMap J g)) (sym law2) ⟩ 35 | comp D (comp D (bind (comp D η (HMap J f))) η) (HMap J g) 36 | ≅⟨ ass D ⟩ 37 | comp D (bind (comp D η (HMap J f))) (comp D η (HMap J g)) 38 | ∎} 39 | 40 | RKlR : Fun Kl D 41 | RKlR = record{ 42 | OMap = T; 43 | HMap = bind; 44 | fid = law1; 45 | fcomp = law3} 46 | -------------------------------------------------------------------------------- /Monads.agda: -------------------------------------------------------------------------------- 1 | module Monads where 2 | 3 | open import Library 4 | open import Categories 5 | 6 | record Monad {a}{b}(C : Cat {a}{b}) : Set (a ⊔ b) where 7 | constructor monad 8 | open Cat C 9 | field T : Obj → Obj 10 | η : ∀ {X} → Hom X (T X) 11 | bind : ∀{X Y} → Hom X (T Y) → Hom (T X) (T Y) 12 | law1 : ∀{X} → bind (η {X}) ≅ iden {T X} 13 | law2 : ∀{X Y}{f : Hom X (T Y)} → comp (bind f) η ≅ f 14 | law3 : ∀{X Y Z}{f : Hom X (T Y)}{g : Hom Y (T Z)} → 15 | bind (comp (bind g) f) ≅ comp (bind g) (bind f) 16 | 17 | open import Functors 18 | 19 | TFun : ∀{a b}{C : Cat {a}{b}} → Monad C → Fun C C 20 | TFun {C = C} M = let open Monad M; open Cat C in record { 21 | OMap = T; 22 | HMap = bind ∘ comp η; 23 | fid = 24 | proof 25 | bind (comp η iden) 26 | ≅⟨ cong bind idr ⟩ 27 | bind η 28 | ≅⟨ law1 ⟩ 29 | iden ∎; 30 | fcomp = λ {_}{_}{_}{f}{g} → 31 | proof 32 | bind (comp η (comp f g)) 33 | ≅⟨ cong bind (sym ass) ⟩ 34 | bind (comp (comp η f) g) 35 | ≅⟨ cong (λ f → bind (comp f g)) (sym law2) ⟩ 36 | bind (comp (comp (bind (comp η f)) η) g) 37 | ≅⟨ cong bind ass ⟩ 38 | bind (comp (bind (comp η f)) (comp η g)) 39 | ≅⟨ law3 ⟩ 40 | comp (bind (comp η f)) (bind (comp η g)) 41 | ∎} 42 | -------------------------------------------------------------------------------- /Adjunctions.agda: -------------------------------------------------------------------------------- 1 | module Adjunctions where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | 7 | open Cat 8 | open Fun 9 | 10 | record Adj {a b c d}(C : Cat {a}{b})(D : Cat {c}{d}) : Set (a ⊔ b ⊔ c ⊔ d) 11 | where 12 | constructor adjunction 13 | field L : Fun C D 14 | R : Fun D C 15 | left : {X : Obj C}{Y : Obj D} → 16 | Hom D (OMap L X) Y → Hom C X (OMap R Y) 17 | right : {X : Obj C}{Y : Obj D} → 18 | Hom C X (OMap R Y) → Hom D (OMap L X) Y 19 | lawa : {X : Obj C}{Y : Obj D}(f : Hom D (OMap L X) Y) → 20 | right (left f) ≅ f 21 | lawb : {X : Obj C}{Y : Obj D}(f : Hom C X (OMap R Y)) → 22 | left (right f) ≅ f 23 | natleft : {X X' : Obj C}{Y Y' : Obj D} 24 | (f : Hom C X' X)(g : Hom D Y Y') 25 | (h : Hom D (OMap L X) Y) → 26 | comp C (HMap R g) (comp C (left h) f) 27 | ≅ 28 | left (comp D g (comp D h (HMap L f))) 29 | natright : {X X' : Obj C}{Y Y' : Obj D} 30 | (f : Hom C X' X)(g : Hom D Y Y') 31 | (h : Hom C X (OMap R Y)) → 32 | right (comp C (HMap R g) (comp C h f)) 33 | ≅ 34 | comp D g (comp D (right h) (HMap L f)) 35 | -------------------------------------------------------------------------------- /RMonads/Restriction.agda: -------------------------------------------------------------------------------- 1 | module RMonads.Restriction where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | open import Naturals 7 | open import Monads 8 | open import RMonads 9 | 10 | open Cat 11 | open Fun 12 | 13 | restrictM : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}(J : Fun C D) → 14 | Monad D → RMonad J 15 | restrictM J M = record { 16 | T = T ∘ OMap J; 17 | η = η; 18 | bind = bind; 19 | law1 = law1; 20 | law2 = law2; 21 | law3 = law3} 22 | where open Monad M 23 | 24 | open import Monads.MonadMorphs 25 | open import RMonads.RMonadMorphs 26 | 27 | restrictMM : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{M M' : Monad D} 28 | (J : Fun C D) → MonadMorph M M' → 29 | RMonadMorph (restrictM J M) (restrictM J M') 30 | restrictMM J MM = record { 31 | morph = λ{X} → morph {OMap J X}; 32 | lawη = lawη; 33 | lawbind = lawbind} 34 | where open MonadMorph MM 35 | open import Adjunctions 36 | open import RAdjunctions 37 | 38 | restrictA : ∀{a b c d e f}{C : Cat {a}{b}}{D : Cat {c}{d}}{E : Cat {e}{f}} 39 | (J : Fun C D) → Adj D E → RAdj J E 40 | restrictA J A = record{ 41 | L = L ○ J; 42 | R = R; 43 | left = left; 44 | right = right; 45 | lawa = lawa; 46 | lawb = lawb; 47 | natleft = natleft ∘ HMap J; 48 | natright = natright ∘ HMap J} 49 | where open Adj A 50 | -------------------------------------------------------------------------------- /YonedaLemma.agda: -------------------------------------------------------------------------------- 1 | module YonedaLemma where 2 | 3 | open import Library 4 | open import Categories 5 | open import Categories.Sets 6 | open import Functors 7 | open import Naturals 8 | open import FunctorCat 9 | 10 | HomF[-,_] : ∀{l m}{C : Cat {l}{m}}(X : Cat.Obj C) -> Fun (C Op) (Sets {m}) 11 | HomF[-,_] {C = C} B = functor 12 | (\ X -> Hom X B) 13 | (\ f g -> comp g f) 14 | (ext \ _ -> idr) 15 | (ext \ _ -> sym ass) 16 | where open Cat C 17 | 18 | HomN[-,_] : ∀{l m}{C : Cat {l}{m}}{X Y : Cat.Obj C}(f : Cat.Hom C X Y) -> 19 | NatT (HomF[-,_] {C = C} X) HomF[-, Y ] 20 | HomN[-,_] {C = C} f = natural 21 | (comp f) 22 | (ext \ _ -> ass) 23 | where open Cat C 24 | 25 | y : ∀{l m}(C : Cat {l}{m}) -> Fun C (FunctorCat (C Op) (Sets {m})) 26 | y C = functor 27 | HomF[-,_] 28 | HomN[-,_] 29 | (NatTEq (iext \ _ -> ext \ _ -> idl)) 30 | (NatTEq (iext \ _ -> ext \ _ -> ass)) 31 | where open Cat C 32 | 33 | yleminv : ∀{l m}(C : Cat {l}{m})(F : Fun (C Op) (Sets {m}))(X : Cat.Obj C) -> 34 | NatT (Fun.OMap (y C) X) F -> Fun.OMap F X 35 | yleminv C F X α = NatT.cmp α {X} (Cat.iden C) 36 | 37 | ylem : ∀{l m}(C : Cat {l}{m})(F : Fun (C Op) (Sets {m}))(X : Cat.Obj C) -> 38 | Fun.OMap F X -> NatT (Fun.OMap (y C) X) F 39 | ylem C F X FX = natural 40 | (\ {X'} f -> HMap f FX) 41 | (\{X'}{Y}{f} -> ext \ g -> sym (fcong FX (fcomp {f = f}{g = g})) ) 42 | where open Cat C; open Fun F 43 | -------------------------------------------------------------------------------- /Monads/EM/Functors.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.EM.Functors {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open import Library 7 | open import Functors 8 | open import Monads.EM M 9 | 10 | open Cat C 11 | open Fun 12 | open Monad M 13 | open Alg 14 | open AlgMorph 15 | 16 | EML : Fun C EM 17 | EML = record { 18 | OMap = λ X → record { 19 | acar = T X; 20 | astr = λ _ → bind; 21 | alaw1 = sym law2; 22 | alaw2 = law3}; 23 | HMap = λ f → record { 24 | amor = bind (comp η f); 25 | ahom = λ {Z} {g} → 26 | proof 27 | comp (bind (comp η f)) (bind g) 28 | ≅⟨ sym law3 ⟩ 29 | bind (comp (bind (comp η f)) g) 30 | ∎}; 31 | fid = AlgMorphEq ( 32 | proof 33 | bind (comp η iden) 34 | ≅⟨ cong bind idr ⟩ 35 | bind η 36 | ≅⟨ law1 ⟩ 37 | iden ∎); 38 | fcomp = λ {_}{_}{_}{f}{g} → 39 | AlgMorphEq ( 40 | proof 41 | bind (comp η (comp f g)) 42 | ≅⟨ cong bind (sym ass) ⟩ 43 | bind (comp (comp η f) g) 44 | ≅⟨ cong (λ f → bind (comp f g)) (sym law2) ⟩ 45 | bind (comp (comp (bind (comp η f)) η) g) 46 | ≅⟨ cong bind ass ⟩ 47 | bind (comp (bind (comp η f)) (comp η g)) 48 | ≅⟨ law3 ⟩ 49 | comp (bind (comp η f)) (bind (comp η g)) 50 | ∎)} 51 | 52 | EMR : Fun EM C 53 | EMR = record { 54 | OMap = acar; 55 | HMap = amor; 56 | fid = refl; 57 | fcomp = refl} 58 | -------------------------------------------------------------------------------- /DayConvolution.agda: -------------------------------------------------------------------------------- 1 | module DayConvolution where 2 | 3 | open import Library 4 | open import Functors 5 | open import Categories 6 | open import Categories.Sets 7 | open import MonoidalCat 8 | 9 | -- first draft 10 | 11 | ODay : ∀{l m}(M : Monoidal {l}{m}) 12 | (F G : Fun (Monoidal.C M) (Sets {l})) -> 13 | Cat.Obj (Monoidal.C M) -> Set _ 14 | ODay M F G X = 15 | Σ Obj \Y -> Σ Obj \Z -> OMap F Y × OMap G Z × Hom X (OMap ⊗ (Y , Z)) 16 | where open Monoidal M; open Cat C; open Fun 17 | 18 | HDay : ∀{l m}(M : Monoidal {l}{m}) 19 | (F G : Fun (Monoidal.C M) (Sets {l})) -> ∀{X X'} -> 20 | Cat.Hom (Monoidal.C M) X X' -> ODay M F G X' -> ODay M F G X 21 | HDay M F G f (y , z , fy , gz , g) = y , z , fy , gz , comp g f 22 | where open Monoidal M; open Cat C 23 | 24 | 25 | DayF : ∀{l m}(M : Monoidal {l}{m}) 26 | (F G : Fun (Monoidal.C M) (Sets {l})) -> Fun ((Monoidal.C M) Op) Sets 27 | DayF M F G = functor 28 | (ODay M F G) 29 | (HDay M F G) 30 | (\ {X} -> ext (\{ (y , z , fy , gz , g) -> 31 | cong (λ g → y , z , fy , gz , g) idr})) 32 | (\ {X Y Z h k} -> ext (\{ (y , z , fy , gz , g) -> 33 | cong (λ l → y , z , fy , gz , l) (sym ass)})) 34 | where open Monoidal M; open Cat C 35 | 36 | -- subject to some additional conditions 37 | -- forall h : y -> y'. 38 | -- (y , F h fy' , gz , g) ~ (y' , z , fy' , gz , [h,id] . g) 39 | -- 40 | -- forall k : z -> z' 41 | -- (y , z , fy , G k gz' , g) ~ (y , z' , fy , gz' , [id,k] . g) 42 | 43 | -------------------------------------------------------------------------------- /RAdjunctions.agda: -------------------------------------------------------------------------------- 1 | module RAdjunctions where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | 7 | open Cat 8 | open Fun 9 | 10 | record RAdj {a b c d e f}{C : Cat {a}{b}}{D : Cat {c}{d}} 11 | (J : Fun C D)(E : Cat {e}{f}) : Set (a ⊔ b ⊔ c ⊔ d ⊔ e ⊔ f) where 12 | constructor radjunction 13 | field L : Fun C E 14 | R : Fun E D 15 | left : {X : Obj C}{Y : Obj E} → 16 | Hom E (OMap L X) Y → Hom D (OMap J X) (OMap R Y) 17 | right : {X : Obj C}{Y : Obj E} → 18 | Hom D (OMap J X) (OMap R Y) → Hom E (OMap L X) Y 19 | lawa : {X : Obj C}{Y : Obj E}(f : Hom E (OMap L X) Y) → 20 | right (left f) ≅ f 21 | lawb : {X : Obj C}{Y : Obj E}(f : Hom D (OMap J X) (OMap R Y)) → 22 | left (right f) ≅ f 23 | natleft : {X X' : Obj C}{Y Y' : Obj E} 24 | (f : Hom C X' X)(g : Hom E Y Y') 25 | (h : Hom E (OMap L X) Y) → 26 | comp D (HMap R g) (comp D (left h) (HMap J f)) 27 | ≅ 28 | left (comp E g (comp E h (HMap L f))) 29 | natright : {X X' : Obj C}{Y Y' : Obj E} 30 | (f : Hom C X' X)(g : Hom E Y Y') 31 | (h : Hom D (OMap J X) (OMap R Y)) → 32 | right (comp D (HMap R g) (comp D h (HMap J f))) 33 | ≅ 34 | comp E g (comp E (right h) (HMap L f)) 35 | -------------------------------------------------------------------------------- /Vectors.agda: -------------------------------------------------------------------------------- 1 | record SemiRing : Set1 where 2 | field 3 | R : Set 4 | _+_ : R -> R -> R 5 | zero : R 6 | _*_ : R -> R -> R 7 | one : R 8 | 9 | module Vectors (S : SemiRing) where 10 | 11 | open SemiRing S 12 | open import Data.Nat renaming (ℕ to Nat; zero to z; _*_ to times; _+_ to plus) 13 | open import Data.Fin renaming (zero to z) hiding (_+_) 14 | open import RMonads 15 | open import Functors.Fin 16 | open import Data.Bool 17 | open import Function 18 | open import Relation.Binary.HeterogeneousEquality 19 | 20 | Vec : Nat -> Set 21 | Vec n = Fin n -> R 22 | 23 | Matrix : Nat -> Nat -> Set 24 | Matrix m n = Fin m -> Vec n 25 | 26 | -- unit 27 | delta : forall {n} -> Matrix n n 28 | delta i j = if feq i j then one else zero 29 | 30 | transpose : forall {m n} -> Matrix m n -> Matrix n m 31 | transpose A = λ j i -> A i j 32 | 33 | dot : forall {n} -> Vec n -> Vec n -> R 34 | dot {z} x y = zero 35 | dot {suc n} x y = (x z * y z) + dot (x ∘ suc) (y ∘ suc) 36 | 37 | mult : forall {m n} -> Matrix m n -> Vec m -> Vec n 38 | mult A x = λ j -> dot (transpose A j) x 39 | 40 | VecRMon : RMonad FinF 41 | VecRMon = rmonad 42 | Vec 43 | delta 44 | mult 45 | {!!} 46 | {!!} 47 | {!!} 48 | 49 | 50 | 51 | where 52 | {- 53 | lem : forall {n}(x : Vec n)(j : Fin n) -> dot (λ i → if feq i j then one else zero) x ≅ x j 54 | lem {z} x () 55 | lem {suc n} x z = {!lem {n} (x ∘ suc) !} 56 | lem {suc n} x (suc j) = {!!} 57 | 58 | -} 59 | -------------------------------------------------------------------------------- /Categories/Setoids.agda: -------------------------------------------------------------------------------- 1 | module Categories.Setoids where 2 | 3 | open import Library 4 | open import Categories 5 | 6 | 7 | record Setoid {a b} : Set (lsuc (a ⊔ b)) where 8 | field set : Set a 9 | eq : set → set → Set b 10 | ref : {s : set} → eq s s 11 | sym' : {s s' : set} → eq s s' → eq s' s 12 | trn : {s s' s'' : set} → 13 | eq s s' → eq s' s'' → eq s s'' 14 | open Setoid 15 | 16 | record SetoidFun (S S' : Setoid) : Set where 17 | constructor setoidfun 18 | field fun : set S → set S' 19 | feq : {s s' : set S} → 20 | eq S s s' → eq S' (fun s) (fun s') 21 | open SetoidFun 22 | 23 | SetoidFunEq : {S S' : Setoid}{f g : SetoidFun S S'} → 24 | fun f ≅ fun g → 25 | (λ{s}{s'}(p : eq S s s') → feq f p) 26 | ≅ 27 | (λ{s}{s'}(p : eq S s s') → feq g p) → 28 | f ≅ g 29 | SetoidFunEq {f = setoidfun fun feq} {setoidfun .fun .feq} refl refl = refl 30 | 31 | idFun : {S : Setoid} → SetoidFun S S 32 | idFun = record {fun = id; feq = id} 33 | 34 | compFun : {S S' S'' : Setoid} → 35 | SetoidFun S' S'' → SetoidFun S S' → SetoidFun S S'' 36 | compFun f g = record {fun = fun f ∘ fun g; feq = feq f ∘ feq g} 37 | 38 | idl : {X Y : Setoid} {f : SetoidFun X Y} → compFun idFun f ≅ f 39 | idl {f = setoidfun _ _} = refl 40 | 41 | idr : {X Y : Setoid} {f : SetoidFun X Y} → compFun f idFun ≅ f 42 | idr {f = setoidfun _ _} = refl 43 | 44 | Setoids : Cat 45 | Setoids = record{ 46 | Obj = Setoid; 47 | Hom = SetoidFun; 48 | iden = idFun; 49 | comp = compFun; 50 | idl = idl; 51 | idr = idr; 52 | ass = refl} 53 | -- -} 54 | -------------------------------------------------------------------------------- /RMonads/REM/Functors.agda: -------------------------------------------------------------------------------- 1 | open import Functors 2 | open import Categories 3 | open import RMonads 4 | 5 | module RMonads.REM.Functors {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 6 | (J : Fun C D)(M : RMonad J) where 7 | 8 | open import Library 9 | open import RMonads.REM M 10 | 11 | open Cat 12 | open Fun 13 | open RAlg 14 | open RAlgMorph 15 | open RMonad M 16 | 17 | REML : Fun C EM 18 | REML = record { 19 | OMap = λ X → record { 20 | acar = T X; 21 | astr = bind; 22 | alaw1 = sym law2; 23 | alaw2 = law3}; 24 | HMap = λ f → record { 25 | amor = bind (comp D η (HMap J f)); 26 | ahom = sym law3}; 27 | fid = RAlgMorphEq ( 28 | proof 29 | bind (comp D η (HMap J (iden C))) 30 | ≅⟨ cong (bind ∘ comp D η) (fid J) ⟩ 31 | bind (comp D η (iden D)) 32 | ≅⟨ cong bind (idr D) ⟩ 33 | bind η 34 | ≅⟨ law1 ⟩ 35 | iden D ∎); 36 | fcomp = λ{_ _ _ f g} → RAlgMorphEq ( 37 | proof 38 | bind (comp D η (HMap J (comp C f g))) 39 | ≅⟨ cong (bind ∘ comp D η) (fcomp J) ⟩ 40 | bind (comp D η (comp D (HMap J f) (HMap J g))) 41 | ≅⟨ cong bind (sym (ass D)) ⟩ 42 | bind (comp D (comp D η (HMap J f)) (HMap J g)) 43 | ≅⟨ cong (λ f₁ → bind (comp D f₁ (HMap J g))) (sym law2) ⟩ 44 | bind (comp D (comp D (bind (comp D η (HMap J f))) η) (HMap J g)) 45 | ≅⟨ cong bind (ass D) ⟩ 46 | bind (comp D (bind (comp D η (HMap J f))) (comp D η (HMap J g))) 47 | ≅⟨ law3 ⟩ 48 | comp D (bind (comp D η (HMap J f))) (bind (comp D η (HMap J g))) 49 | ∎)} 50 | 51 | REMR : Fun EM D 52 | REMR = record { 53 | OMap = acar; 54 | HMap = amor; 55 | fid = refl; 56 | fcomp = refl} 57 | -------------------------------------------------------------------------------- /Functors.agda: -------------------------------------------------------------------------------- 1 | module Functors where 2 | 3 | open import Library 4 | open import Categories 5 | open Cat 6 | 7 | record Fun {a b c d} (C : Cat {a}{b})(D : Cat {c}{d}) : Set (a ⊔ b ⊔ c ⊔ d) 8 | where 9 | constructor functor 10 | field OMap : Obj C → Obj D 11 | HMap : ∀{X Y} → Hom C X Y → Hom D (OMap X) (OMap Y) 12 | fid : ∀{X} → HMap (iden C {X}) ≅ iden D {OMap X} 13 | fcomp : ∀{X Y Z}{f : Hom C Y Z}{g : Hom C X Y} → 14 | HMap (comp C f g) ≅ comp D (HMap f) (HMap g) 15 | open Fun 16 | 17 | IdF : ∀{a b}(C : Cat {a}{b}) → Fun C C 18 | IdF C = record{OMap = id;HMap = id;fid = refl;fcomp = refl} 19 | 20 | _○_ : ∀{a b c d e f}{C : Cat {a}{b}}{D : Cat {c}{d}}{E : Cat {e}{f}} → 21 | Fun D E → Fun C D → Fun C E 22 | _○_ {C = C}{D}{E} F G = record{ 23 | OMap = OMap F ∘ OMap G; 24 | HMap = HMap F ∘ HMap G; 25 | fid = 26 | proof 27 | HMap F (HMap G (iden C)) 28 | ≅⟨ cong (HMap F) (fid G) ⟩ 29 | HMap F (iden D) 30 | ≅⟨ fid F ⟩ 31 | iden E 32 | ∎; 33 | fcomp = λ {_}{_}{_}{f}{g} → 34 | proof 35 | HMap F (HMap G (comp C f g)) 36 | ≅⟨ cong (HMap F) (fcomp G) ⟩ 37 | HMap F (comp D (HMap G f) (HMap G g)) 38 | ≅⟨ fcomp F ⟩ 39 | comp E (HMap F (HMap G f)) (HMap F (HMap G g)) 40 | ∎} 41 | infix 10 _○_ 42 | 43 | FunctorEq : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}(F G : Fun C D) → 44 | OMap F ≅ OMap G → 45 | (λ {X Y} → HMap F {X}{Y}) ≅ (λ {X}{Y} → HMap G {X}{Y}) → F ≅ G 46 | FunctorEq (functor fo fh _ _) (functor .fo .fh _ _) refl refl = 47 | cong₂ (functor fo fh) 48 | (iext λ _ → ir _ _) 49 | (iext λ _ → iext λ _ → iext λ _ → iext λ _ → iext λ _ → ir _ _) 50 | -------------------------------------------------------------------------------- /README.md: -------------------------------------------------------------------------------- 1 | This repository contains a formalisation of relative monads. Monads 2 | are relative monads on the identity functor so it is not technically 3 | necessary to repeat the same constructions for ordinary 4 | monads. Nevertheless I include separate implementations of ordinary 5 | monads and related constructions as a warmup. 6 | 7 | Most of the code is quite polished. The stuff related to categories of 8 | adjunctions is still pretty gory. 9 | 10 | Basic category theory 11 | --------------------- 12 | * Categories 13 | * Initial object 14 | * Terminal object 15 | * CoProducts 16 | * Functors 17 | * Full and Faithfulness 18 | * Natural Transformations 19 | * Day Convolution 20 | * Yoneda Lemma 21 | 22 | Relative monad theory 23 | --------------------- 24 | * (Rel.) Monads 25 | * (Rel.) Adjunctions 26 | * (Rel.) Kleisli Category 27 | * (Rel.) Kleisli Adjunction 28 | * (Rel.) EM Category 29 | * (Rel.) EM Adjunction 30 | * (Rel.) Modules 31 | * Category of (Rel.) Adjunctions for a (Relative) Monad 32 | * Proof that (Rel.) Kleisli is the Initial Object in the Cat. of (Rel.) Adjs. 33 | * Proof that (Rel.) EM is the Terminal Object in the Cat. of (Rel.) Adjs. 34 | 35 | Examples 36 | -------- 37 | * Proof that Well-Scoped Lambda Terms are a Rel. Monad. 38 | * Proof that a Lambda-Model is a Rel. EM Algebra. 39 | * Proof that Well-Typed Lambda Terms are a Rel. Monad. 40 | * Proof that an evaluator for Well-Typed Lambda Terms is a Rel. EM Algebra. 41 | * Proof that weak arrows are relative monads on Yoneda 42 | * Proof that Rel. Monads on Fin are Lawvere theories 43 | * Proof that Algebras for Rel. Monads on Fin are models of Lawvere theories. 44 | * Proof that vector spaces are Rel. Monads 45 | * Proof that left-modules are algebras of Rel. Monads -------------------------------------------------------------------------------- /Adjunctions/Adj2Mon.agda: -------------------------------------------------------------------------------- 1 | module Adjunctions.Adj2Mon where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | open import Monads 7 | open import Adjunctions 8 | 9 | open Cat 10 | open Fun 11 | 12 | Adj2Mon : ∀{a b}{C D : Cat {a}{b}} → Adj C D → Monad C 13 | Adj2Mon {C = C}{D} A = record{ 14 | T = OMap R ∘ OMap L; 15 | η = left (iden D); 16 | bind = HMap R ∘ right; 17 | law1 = 18 | proof 19 | HMap R (right (left (iden D))) 20 | ≅⟨ cong (HMap R) (lawa (iden D)) ⟩ 21 | HMap R (iden D) 22 | ≅⟨ fid R ⟩ 23 | iden C ∎; 24 | law2 = λ{_}{_}{f} → 25 | proof 26 | comp C (HMap R (right f)) (left (iden D)) 27 | ≅⟨ cong (comp C (HMap R (right f))) (sym (idr C)) ⟩ 28 | comp C (HMap R (right f)) (comp C (left (iden D)) (iden C)) 29 | ≅⟨ natleft (iden C) (right f) (iden D) ⟩ 30 | left (comp D (right f) (comp D (iden D) (HMap L (iden C)))) 31 | ≅⟨ cong (left ∘ comp D (right f)) (idl D) ⟩ 32 | left (comp D (right f) (HMap L (iden C))) 33 | ≅⟨ cong (left ∘ comp D (right f)) (fid L) ⟩ 34 | left (comp D (right f) (iden D)) 35 | ≅⟨ cong (left) (idr D) ⟩ 36 | left (right f) 37 | ≅⟨ lawb f ⟩ 38 | f ∎; 39 | law3 = λ{_}{_}{_}{f}{g} → 40 | proof 41 | HMap R (right (comp C (HMap R (right g)) f)) 42 | ≅⟨ cong (HMap R ∘ right ∘ comp C (HMap R (right g))) 43 | (sym (idr C)) ⟩ 44 | HMap R (right (comp C (HMap R (right g)) (comp C f (iden C)))) 45 | ≅⟨ cong (HMap R) (natright (iden C) (right g) f) ⟩ 46 | HMap R (comp D (right g) (comp D (right f) (HMap L (iden C)))) 47 | ≅⟨ cong (HMap R ∘ comp D (right g) ∘ comp D (right f)) (fid L)⟩ 48 | HMap R (comp D (right g) (comp D (right f) (iden D))) 49 | ≅⟨ cong (HMap R ∘ comp D (right g)) (idr D) ⟩ 50 | HMap R (comp D (right g) (right f)) 51 | ≅⟨ fcomp R ⟩ 52 | comp C (HMap R (right g)) (HMap R (right f)) 53 | ∎} 54 | where open Adj A 55 | -------------------------------------------------------------------------------- /RAdjunctions/RAdj2RMon.agda: -------------------------------------------------------------------------------- 1 | module RAdjunctions.RAdj2RMon where 2 | 3 | open import Function 4 | open import Relation.Binary.HeterogeneousEquality 5 | open import Categories 6 | open import Functors 7 | open import RMonads 8 | open import RAdjunctions 9 | 10 | open Cat 11 | open Fun 12 | open RAdj 13 | 14 | Adj2Mon : ∀{a b c d e f}{C : Cat {a}{b}}{D : Cat {c}{d}}{E : Cat {e}{f}} 15 | {J : Fun C D} → RAdj J E → RMonad J 16 | Adj2Mon {C = C}{D}{E}{J} A = record{ 17 | T = OMap (R A) ∘ OMap (L A); 18 | η = left A (iden E); 19 | bind = HMap (R A) ∘ right A; 20 | law1 = trans (cong (HMap (R A)) (lawa A (iden E))) (fid (R A)); 21 | law2 = λ {_ _ f} → 22 | trans (cong (comp D (HMap (R A) (right A f))) 23 | (trans (sym (idr D)) 24 | (cong (comp D (left A (iden E))) (sym (fid J))))) 25 | (trans (natleft A (iden C) (right A f) (iden E)) 26 | (trans (cong (left A) 27 | (trans (cong (comp E (right A f)) 28 | (trans (idl E) (fid (L A)))) 29 | (idr E))) 30 | (lawb A f))); 31 | law3 = λ{_ _ _ f g} → 32 | trans (cong (HMap (R A)) 33 | (trans (trans (cong (right A) 34 | (cong (comp D (HMap (R A) (right A g))) 35 | (trans (sym (idr D)) 36 | (cong (comp D f) 37 | (sym (fid J)))))) 38 | (trans (natright A (iden C) (right A g) f) 39 | (trans (sym (ass E)) 40 | (cong (comp E (comp E (right A g) 41 | (right A f))) 42 | (fid (L A)))))) 43 | (idr E))) 44 | (fcomp (R A))} 45 | -------------------------------------------------------------------------------- /RMonads.agda: -------------------------------------------------------------------------------- 1 | module RMonads where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | 7 | 8 | record RMonad {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}(J : Fun C D) : 9 | Set (a ⊔ b ⊔ c ⊔ d) where 10 | constructor rmonad 11 | open Cat 12 | open Fun 13 | field T : Obj C → Obj D 14 | η : ∀{X} → Hom D (OMap J X) (T X) 15 | bind : ∀{X Y} → Hom D (OMap J X) (T Y) → Hom D (T X) (T Y) 16 | law1 : ∀{X} → bind (η {X}) ≅ iden D {T X} 17 | law2 : ∀{X Y}{f : Hom D (OMap J X) (T Y)} → comp D (bind f) η ≅ f 18 | law3 : ∀{X Y Z} 19 | {f : Hom D (OMap J X) (T Y)}{g : Hom D (OMap J Y) (T Z)} → 20 | bind (comp D (bind g) f) ≅ comp D (bind g) (bind f) 21 | 22 | open import Functors 23 | 24 | TFun : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D} → 25 | RMonad J → Fun C D 26 | TFun {C = C}{D}{J} M = let open RMonad M; open Cat; open Fun in record { 27 | OMap = T; 28 | HMap = bind ∘ comp D η ∘ HMap J; 29 | fid = 30 | proof 31 | bind (comp D η (HMap J (iden C))) 32 | ≅⟨ cong (bind ∘ comp D η) (fid J) ⟩ 33 | bind (comp D η (iden D)) 34 | ≅⟨ cong bind (idr D) ⟩ 35 | bind η 36 | ≅⟨ law1 ⟩ 37 | iden D 38 | ∎; 39 | fcomp = λ{_ _ _ f g} → 40 | proof 41 | bind (comp D η (HMap J (comp C f g))) 42 | ≅⟨ cong (bind ∘ comp D η) (fcomp J) ⟩ 43 | bind (comp D η (comp D (HMap J f) (HMap J g))) 44 | ≅⟨ cong bind (sym (ass D)) ⟩ 45 | bind (comp D (comp D η (HMap J f)) (HMap J g)) 46 | ≅⟨ cong (λ f → bind (comp D f (HMap J g))) (sym law2) ⟩ 47 | bind (comp D (comp D (bind (comp D η (HMap J f))) η) (HMap J g)) 48 | ≅⟨ cong bind (ass D) ⟩ 49 | bind (comp D (bind (comp D η (HMap J f))) (comp D η (HMap J g))) 50 | ≅⟨ law3 ⟩ 51 | comp D (bind (comp D η (HMap J f))) (bind (comp D η (HMap J g))) 52 | ∎} 53 | 54 | -- any functor is a rel monad over itself 55 | trivRM : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}(J : Fun C D) → RMonad J 56 | trivRM {D = D} J = rmonad OMap (iden D) id refl (idr D) refl 57 | where open Fun J; open Cat 58 | -------------------------------------------------------------------------------- /RMonads/REM/Adjunction.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | open import RMonads 4 | 5 | module RMonads.REM.Adjunction {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 6 | {J : Fun C D}(M : RMonad J) where 7 | 8 | open import Library 9 | open import RAdjunctions 10 | open import RMonads.REM M 11 | open import RMonads.REM.Functors J M 12 | 13 | open Fun 14 | open RAdj 15 | open RAlg 16 | open RAlgMorph 17 | open RMonad M 18 | open Cat D 19 | 20 | REMAdj : RAdj J EM 21 | REMAdj = record { 22 | L = REML; 23 | R = REMR; 24 | left = λ f → comp (amor f) η; 25 | right = λ{_ B} f → record{ 26 | amor = astr B f; 27 | ahom = sym (alaw2 B)}; 28 | lawa = λ {_ Y} f → 29 | RAlgMorphEq ( 30 | proof 31 | astr Y (comp (amor f) η) 32 | ≅⟨ sym (ahom f) ⟩ 33 | comp (amor f) (bind η) 34 | ≅⟨ cong (comp (amor f)) law1 ⟩ 35 | comp (amor f) iden 36 | ≅⟨ idr ⟩ 37 | amor f ∎); 38 | lawb = λ {_ B} f → sym (alaw1 B); 39 | natleft = λ f g h → 40 | proof 41 | comp (amor g) (comp (comp (amor h) η) (HMap J f)) 42 | ≅⟨ cong (comp (amor g)) ass ⟩ 43 | comp (amor g) (comp (amor h) (comp η (HMap J f))) 44 | ≅⟨ cong (comp (amor g) ∘ comp (amor h)) (sym law2) ⟩ 45 | comp (amor g) (comp (amor h) (comp (bind (comp η (HMap J f))) η)) 46 | ≅⟨ cong (comp (amor g)) (sym ass) ⟩ 47 | comp (amor g) (comp (comp (amor h) (bind (comp η (HMap J f)))) η) 48 | ≅⟨ sym ass ⟩ 49 | comp (comp (amor g) (comp (amor h) (bind (comp η (HMap J f))))) η 50 | ∎; 51 | natright = λ{W}{X}{Y}{Z} f g h → 52 | RAlgMorphEq ( 53 | proof 54 | astr Z (comp (amor g) (comp h (HMap J f))) 55 | ≅⟨ sym (ahom g) ⟩ 56 | comp (amor g) (astr Y (comp h (HMap J f))) 57 | ≅⟨ cong (λ h₁ → comp (amor g) (astr Y (comp h₁ (HMap J f)))) (alaw1 Y) ⟩ 58 | comp (amor g) (astr Y (comp (comp (astr Y h) η) (HMap J f))) 59 | ≅⟨ cong (comp (amor g) ∘ astr Y) ass ⟩ 60 | comp (amor g) (astr Y (comp (astr Y h) (comp η (HMap J f)))) 61 | ≅⟨ cong (comp (amor g)) (alaw2 Y) ⟩ 62 | comp (amor g) (comp (astr Y h) (bind (comp η (HMap J f)))) 63 | ∎)} 64 | -------------------------------------------------------------------------------- /Everything.agda: -------------------------------------------------------------------------------- 1 | module Everything where 2 | 3 | -- basic utilities 4 | open import Library 5 | open import Isomorphism 6 | 7 | -- basic category theory 8 | open import Categories 9 | open import Categories.Sets 10 | open import Categories.Families 11 | open import Categories.Initial 12 | open import Categories.Terminal 13 | open import Categories.CoProducts 14 | open import Categories.PushOuts 15 | open import Categories.Setoids -- should be replaced by standard libary def 16 | 17 | open import Functors 18 | open import Functors.Fin 19 | open import Functors.FullyFaithful 20 | 21 | open import Naturals 22 | 23 | -- basic examples 24 | open import Monoids 25 | open import FunctorCat 26 | 27 | -- ordinary monads 28 | open import Monads 29 | open import Monads.MonadMorphs 30 | open import Adjunctions 31 | open import Adjunctions.Adj2Mon 32 | open import Monads.Kleisli 33 | open import Monads.Kleisli.Functors 34 | open import Monads.Kleisli.Adjunction 35 | open import Monads.EM 36 | open import Monads.EM.Functors 37 | open import Monads.EM.Adjunction 38 | open import Monads.CatofAdj 39 | open import Monads.CatofAdj.InitAdj 40 | open import Monads.CatofAdj.TermAdjObj 41 | open import Monads.CatofAdj.TermAdjHom 42 | open import Monads.CatofAdj.TermAdjUniq 43 | open import Monads.CatofAdj.TermAdj 44 | 45 | -- relative monads 46 | open import RMonads 47 | open import RMonads.RMonadMorphs 48 | open import RAdjunctions 49 | open import RAdjunctions.RAdj2RMon 50 | open import RMonads.REM 51 | open import RMonads.REM.Functors 52 | open import RMonads.REM.Adjunction 53 | open import RMonads.RKleisli 54 | open import RMonads.RKleisli.Functors 55 | open import RMonads.RKleisli.Adjunction 56 | open import RMonads.Restriction 57 | open import RMonads.SpecialCase 58 | open import RMonads.CatofRAdj 59 | open import RMonads.CatofRAdj.InitRAdj 60 | open import RMonads.CatofRAdj.TermRAdjObj 61 | open import RMonads.CatofRAdj.TermRAdjHom 62 | open import RMonads.CatofRAdj.TermRAdj 63 | open import RMonads.Modules 64 | 65 | -- rmonad examples 66 | open import WellScopedTerms 67 | open import WellScopedTermsModel 68 | open import WellTypedTerms 69 | open import WellTypedTermsModel 70 | 71 | open import Lawvere 72 | -------------------------------------------------------------------------------- /Monads/EM/Adjunction.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.EM.Adjunction {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open import Library 7 | 8 | open import Functors 9 | open import Adjunctions 10 | open import Monads.EM M 11 | open import Monads.EM.Functors M 12 | 13 | open Cat C 14 | open Fun 15 | open Monad M 16 | open Adj 17 | open Alg 18 | open AlgMorph 19 | 20 | EMAdj : Adj C EM 21 | EMAdj = record { 22 | L = EML; 23 | R = EMR; 24 | left = λ f → comp (amor f) η; 25 | right = λ {X}{Y} f → 26 | record{amor = astr Y X f; 27 | ahom = λ {Z}{g} → 28 | proof 29 | comp (astr Y X f) (astr (OMap EML X) Z g) 30 | ≅⟨ sym (alaw2 Y) ⟩ 31 | astr Y Z (comp (astr Y X f) g) 32 | ∎}; 33 | lawa = λ {X}{Y}(f : AlgMorph (OMap EML X) Y) → AlgMorphEq ( 34 | proof 35 | astr Y X (comp (amor f) η) 36 | ≅⟨ sym (ahom f) ⟩ 37 | comp (amor f) (astr (OMap EML X) X η) 38 | ≡⟨⟩ 39 | comp (amor f) (bind η) 40 | ≅⟨ cong (comp (amor f)) law1 ⟩ 41 | comp (amor f) iden 42 | ≅⟨ idr ⟩ 43 | amor f 44 | ∎); 45 | lawb = λ {X}{Y} f → 46 | proof 47 | comp (astr Y X f) η 48 | ≅⟨ sym (alaw1 Y) ⟩ 49 | f 50 | ∎; 51 | natleft = λ{X}{X'}{Y}{Y'} f g h → 52 | proof 53 | comp (amor g) (comp (comp (amor h) η) f) 54 | ≅⟨ cong (comp (amor g)) ass ⟩ 55 | comp (amor g) (comp (amor h) (comp η f)) 56 | ≅⟨ cong (comp (amor g) ∘ comp (amor h)) (sym law2) ⟩ 57 | comp (amor g) (comp (amor h) (comp (bind (comp η f)) η)) 58 | ≅⟨ cong (comp (amor g)) (sym ass) ⟩ 59 | comp (amor g) (comp (comp (amor h) (bind (comp η f))) η) 60 | ≅⟨ sym ass ⟩ 61 | comp (comp (amor g) (comp (amor h) (bind (comp η f)))) η 62 | ∎; 63 | natright = λ{X}{X'}{Y}{Y'} f g h → AlgMorphEq ( 64 | proof 65 | astr Y' X' (comp (amor g) (comp h f)) 66 | ≅⟨ sym (ahom g) ⟩ 67 | comp (amor g) (astr Y X' (comp h f)) 68 | ≅⟨ cong (λ h → comp (amor g) (astr Y X' (comp h f))) (alaw1 Y) ⟩ 69 | comp (amor g) (astr Y X' (comp (comp (astr Y X h) η) f)) 70 | ≅⟨ cong (comp (amor g) ∘ astr Y X') ass ⟩ 71 | comp (amor g) (astr Y X' (comp (astr Y X h) (comp η f))) 72 | ≅⟨ cong (comp (amor g)) (alaw2 Y) ⟩ 73 | comp (amor g) (comp (astr Y X h) (bind (comp η f))) 74 | ∎)} 75 | -------------------------------------------------------------------------------- /RMonads/Modules.agda: -------------------------------------------------------------------------------- 1 | module RMonads.Modules where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | open import RMonads 7 | 8 | 9 | 10 | record Mod {a}{b}{c}{d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D} 11 | (RM : RMonad J) : Set (a ⊔ c ⊔ d) where 12 | constructor mod 13 | open Cat 14 | open Fun 15 | open RMonad RM 16 | field M : Obj C → Obj D 17 | mbind : ∀ {X Y} → Hom D (OMap J X) (T Y) → Hom D (M X) (M Y) 18 | laweta : ∀{X} → mbind (η {X}) ≅ iden D {M X} 19 | lawbind : ∀{X Y Z} 20 | {f : Hom D (OMap J X) (T Y)}{g : Hom D (OMap J Y) (T Z)} → 21 | mbind (comp D (bind g) f) ≅ comp D (mbind g) (mbind f) 22 | 23 | -- any rel. monad is trivially a module over its J. 24 | ModJ : forall {a}{b}{c}{d}{C : Cat {a}{b}}{D : Cat {c}{d}} 25 | {J : Fun C D} (RM : RMonad J) -> Mod (trivRM J) 26 | ModJ {D = D}{J = J} RM = 27 | mod T 28 | (\ f -> bind (comp η f)) 29 | (trans (cong bind idr) law1) 30 | (trans (cong bind (trans (trans (sym ass) 31 | (cong (\ g -> comp g _) (sym law2))) 32 | ass)) 33 | law3) 34 | where open Fun J; open RMonad RM; open Cat D 35 | 36 | 37 | open import Functors.FullyFaithful 38 | open import Isomorphism 39 | -- if J is fully faithful then any suitably typed functor is a module for trivial rel monad given by J 40 | ModF : forall {a}{b}{c}{d}{C : Cat {a}{b}}{D : Cat {c}{d}} 41 | {J : Fun C D} → Full J → Faithful J → (F : Fun C D) -> Mod (trivRM J) 42 | ModF {C = C}{D = D}{J = J} P Q F = mod 43 | (OMap F) 44 | (λ f → HMap F (fst (P f))) 45 | (trans (cong (HMap F) (Q (trans (snd (P (Cat.iden D))) (sym $ fid J)))) (fid F)) 46 | (trans (cong (HMap F) 47 | (Q (trans (snd (P _)) 48 | (trans (cong₂ (Cat.comp D) 49 | (sym $ snd (P _)) 50 | (sym $ snd (P _))) 51 | (sym $ fcomp J))))) 52 | (fcomp F)) where open Fun 53 | 54 | 55 | -- any rel. monad is trivially a module over itself, 'tautalogical module' 56 | ModRM : forall {a}{b}{c}{d}{C : Cat {a}{b}}{D : Cat {c}{d}} 57 | {J : Fun C D} (RM : RMonad J) -> Mod RM 58 | ModRM RM = mod T bind law1 law3 where open RMonad RM 59 | -------------------------------------------------------------------------------- /RMonads/SpecialCase.agda: -------------------------------------------------------------------------------- 1 | module RMonads.SpecialCase where 2 | 3 | open import Categories 4 | open import Functors 5 | open import Naturals hiding (Iso) 6 | open import Monads 7 | open import RMonads 8 | 9 | 10 | leftM : ∀{a}{b}{C : Cat {a}{b}} → Monad C → RMonad (IdF C) 11 | leftM M = record { 12 | T = T; 13 | η = η; 14 | bind = bind; 15 | law1 = law1; 16 | law2 = law2; 17 | law3 = law3} 18 | where open Monad M 19 | 20 | rightM : ∀{a b}{C : Cat {a}{b}} → RMonad (IdF C) → Monad C 21 | rightM {C = C} M = record { 22 | T = T; 23 | η = η; 24 | bind = bind; 25 | law1 = law1; 26 | law2 = law2; 27 | law3 = law3} where open RMonad M 28 | 29 | open import Relation.Binary.HeterogeneousEquality 30 | open import Isomorphism 31 | 32 | isoM : ∀{a b}{C : Cat {a}{b}} → Iso (RMonad (IdF C)) (Monad C) 33 | isoM = record {fun = rightM; inv = leftM; law1 = λ {(monad _ _ _ _ _ _) → refl}; law2 = λ {(rmonad _ _ _ _ _ _) → refl}} 34 | 35 | open import Monads.MonadMorphs 36 | open import RMonads.RMonadMorphs 37 | 38 | leftMM : ∀{a b}{C : Cat {a}{b}}{M M' : Monad C} → MonadMorph M M' → 39 | RMonadMorph (leftM M) (leftM M') 40 | leftMM MM = record { 41 | morph = morph; 42 | lawη = lawη; 43 | lawbind = lawbind} where open MonadMorph MM 44 | 45 | rightMM : ∀{a b}{C : Cat {a}{b}}{M M' : RMonad (IdF C)} → RMonadMorph M M' → 46 | MonadMorph (rightM M) (rightM M') 47 | rightMM MM = record { 48 | morph = morph; 49 | lawη = lawη; 50 | lawbind = lawbind} where open RMonadMorph MM 51 | 52 | 53 | isoMM : ∀{a b}{C : Cat {a}{b}}{M M' : Monad C} → 54 | Iso (RMonadMorph (leftM M) (leftM M')) (MonadMorph M M') 55 | isoMM {M = monad _ _ _ _ _ _}{M' = monad _ _ _ _ _ _} = record { 56 | fun = rightMM; 57 | inv = leftMM; 58 | law1 = λ {(monadmorph _ _ _) → refl}; 59 | law2 = λ {(rmonadmorph _ _ _) → refl}} 60 | 61 | open import Adjunctions 62 | open import RAdjunctions 63 | 64 | leftA : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} → Adj C D → RAdj (IdF C) D 65 | leftA {C}{D} A = record{ 66 | L = L; 67 | R = R; 68 | left = left; 69 | right = right; 70 | lawa = lawa; 71 | lawb = lawb; 72 | natleft = natleft; 73 | natright = natright} where open Adj A 74 | 75 | rightA : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} → RAdj (IdF C) D → Adj C D 76 | rightA {C = C}{D = D} A = record{ 77 | L = L; 78 | R = R; 79 | left = left; 80 | right = right; 81 | lawa = lawa; 82 | lawb = lawb; 83 | natleft = natleft; 84 | natright = natright} where open RAdj A 85 | 86 | isoA : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} → 87 | Iso (RAdj (IdF C) D) (Adj C D) 88 | isoA = record { 89 | fun = rightA; 90 | inv = leftA; 91 | law1 = λ {(adjunction _ _ _ _ _ _ _ _) → refl}; 92 | law2 = λ {(radjunction _ _ _ _ _ _ _ _) → refl}} 93 | 94 | -------------------------------------------------------------------------------- /WellTypedTermsModel.agda: -------------------------------------------------------------------------------- 1 | module WellTypedTermsModel where 2 | 3 | open import Library 4 | open import WellTypedTerms 5 | open import RMonads.REM 6 | open import FunctorCat 7 | open import Categories.Sets 8 | 9 | 10 | -- interpretation of types 11 | 12 | Val : Ty → Set 13 | Val ι = ⊤ 14 | Val (σ ⇒ τ) = Val σ → Val τ 15 | 16 | -- interpretation of contexts 17 | 18 | Env : Con → Set 19 | Env Γ = ∀{σ} → Var Γ σ → Val σ 20 | 21 | _<<_ : ∀{Γ σ} → Env Γ → Val σ → Env (Γ < σ) 22 | (γ << v) vz = v 23 | (γ << v) (vs x) = γ x 24 | 25 | -- intepretation of terms 26 | 27 | eval : ∀{Γ σ} → Env Γ → Tm Γ σ → Val σ 28 | eval γ (var x) = γ x 29 | eval γ (app t u) = eval γ t (eval γ u) 30 | eval γ (lam t) = λ v → eval (γ << v) t 31 | 32 | substeval : ∀{σ τ}(p : σ ≅ τ){Γ : Con}{γ : Env Γ}(t : Tm Γ σ) → 33 | (subst Val p ∘ eval γ) t ≅ (eval γ ∘ subst (Tm Γ) p) t 34 | substeval refl t = refl 35 | 36 | wk<< : ∀{Γ Δ}(α : Ren Γ Δ)(β : Env Δ){σ}(v : Val σ) → 37 | ∀{ρ}(y : Var (Γ < σ) ρ) → 38 | ((β ∘ α) << v) y ≅ (β << v) (wk α y) 39 | wk<< α β v vz = refl 40 | wk<< α β v (vs y) = refl 41 | 42 | reneval : ∀{Γ Δ σ}(α : Ren Γ Δ)(β : Env Δ)(t : Tm Γ σ) → 43 | eval (β ∘ α) t ≅ (eval β ∘ ren α) t 44 | reneval α β (var x) = refl 45 | reneval α β (app t u) = 46 | proof 47 | eval (β ∘ α) t (eval (β ∘ α) u) 48 | ≅⟨ cong₂ (λ f x → f x) (reneval α β t) (reneval α β u) ⟩ 49 | eval β (ren α t) (eval β (ren α u)) 50 | ∎ 51 | reneval α β (lam t) = ext λ v → 52 | proof 53 | eval ((β ∘ α) << v) t 54 | ≅⟨ cong (λ (γ : Env _) → eval γ t) (iext λ _ → ext (wk<< α β v)) ⟩ 55 | eval ((β << v) ∘ wk α) t 56 | ≅⟨ reneval (wk α) (β << v) t ⟩ 57 | eval (β << v) (ren (wk α) t) 58 | ∎ 59 | 60 | lifteval : ∀{Γ Δ σ τ}(α : Sub Γ Δ)(β : Env Δ)(v : Val σ)(y : Var (Γ < σ) τ) → 61 | ((eval β ∘ α) << v) y ≅ (eval (β << v) ∘ lift α) y 62 | lifteval α β v vz = refl 63 | lifteval α β v (vs x) = 64 | proof 65 | eval β (α x) 66 | ≅⟨ reneval vs (β << v) (α x) ⟩ 67 | eval (β << v) (ren vs (α x)) 68 | ∎ 69 | 70 | subeval : ∀{Γ Δ σ}(α : Sub Γ Δ)(β : Env Δ)(t : Tm Γ σ) → 71 | eval (eval β ∘ α) t ≅ (eval β ∘ sub α) t 72 | subeval α β (var x) = refl 73 | subeval α β (app t u) = 74 | proof 75 | eval (eval β ∘ α) t (eval (eval β ∘ α) u) 76 | ≅⟨ cong₂ (λ f x → f x) (subeval α β t) (subeval α β u) ⟩ 77 | eval β (sub α t) (eval β (sub α u)) 78 | ∎ 79 | subeval α β (lam t) = ext λ v → 80 | proof 81 | eval ((eval β ∘ α) << v) t 82 | ≅⟨ cong (λ (γ : Env _) → eval γ t) (iext λ _ → ext (lifteval α β v)) ⟩ 83 | eval (eval (β << v) ∘ lift α) t 84 | ≅⟨ subeval (lift α) (β << v) t ⟩ 85 | eval (β << v) (sub (lift α) t) 86 | ∎ 87 | 88 | modelRAlg : RAlg TmRMonad 89 | modelRAlg = record { 90 | acar = Val; 91 | astr = λ {Γ} → λ γ → eval γ; 92 | alaw1 = refl; 93 | alaw2 = λ {Γ Δ α γ} → iext λ σ → ext (subeval α γ)} 94 | -------------------------------------------------------------------------------- /RMonads/RMonadMorphs.agda: -------------------------------------------------------------------------------- 1 | module RMonads.RMonadMorphs where 2 | 3 | open import Library 4 | open import Functors 5 | open import Categories 6 | open import RMonads 7 | 8 | open Fun 9 | open RMonad 10 | 11 | record RMonadMorph {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D} 12 | (M M' : RMonad J) : Set (a ⊔ b ⊔ c ⊔ d) where 13 | constructor rmonadmorph 14 | open Cat D 15 | field morph : ∀ {X} → Hom (T M X) (T M' X) 16 | lawη : ∀ {X} → comp morph (η M {X}) ≅ η M' {X} 17 | lawbind : ∀ {X Y}{k : Hom (OMap J X) (T M Y)} → 18 | comp (morph {Y}) (bind M k) 19 | ≅ 20 | comp (bind M' (comp (morph {Y}) k)) (morph {X}) 21 | 22 | 23 | RMonadMorphEq : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D} 24 | {M M' : RMonad J}(α α' : RMonadMorph M M') -> 25 | (λ {X} -> RMonadMorph.morph α {X}) ≅ (λ {X} -> RMonadMorph.morph α' {X}) → 26 | α ≅ α' 27 | RMonadMorphEq (rmonadmorph m lη lbind) (rmonadmorph .m lη' lbind') refl = 28 | cong₂ (rmonadmorph m) 29 | (iext λ _ → hir refl) 30 | (iext λ _ → iext λ _ → iext λ _ → hir refl) 31 | 32 | IdRMonadMorph : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D} 33 | (M : RMonad J) → RMonadMorph M M 34 | IdRMonadMorph {D = D} M = rmonadmorph 35 | iden 36 | idl 37 | (trans idl (trans (cong (bind M) (sym idl)) (sym idr))) 38 | where open Cat D 39 | 40 | CompRMonadMorph : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D} 41 | {M M' M'' : RMonad J} → 42 | RMonadMorph M' M'' → RMonadMorph M M' → RMonadMorph M M'' 43 | CompRMonadMorph {D = D}{M'' = M''} 44 | (rmonadmorph f lawηf lawbindf) 45 | (rmonadmorph g lawηg lawbindg) = 46 | rmonadmorph 47 | (comp f g) 48 | (trans ass (trans (cong (comp f) lawηg) lawηf)) 49 | \ {_ _ k} -> trans 50 | ass 51 | (trans (cong (comp f) lawbindg) 52 | (trans (trans (sym ass) 53 | (cong (λ f → comp f g) 54 | (trans (lawbindf {k = comp g k}) 55 | (cong (λ g → comp (bind M'' g) f) 56 | (sym ass))))) 57 | ass)) 58 | where open Cat D 59 | 60 | idr : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D}{M M' : RMonad J} 61 | (f : RMonadMorph M M') → CompRMonadMorph f (IdRMonadMorph _) ≅ f 62 | idr {D = D} f = RMonadMorphEq _ _ (iext λ _ → Cat.idr D) 63 | 64 | idl : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D}{M M' : RMonad J} 65 | (f : RMonadMorph M M') → CompRMonadMorph (IdRMonadMorph _) f ≅ f 66 | idl {D = D} f = RMonadMorphEq _ _ (iext λ _ → Cat.idl D) 67 | 68 | ass : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 69 | {J : Fun C D}{M M' M'' M''' : RMonad J} 70 | (f : RMonadMorph M'' M''')(g : RMonadMorph M' M'')(h : RMonadMorph M M') → 71 | CompRMonadMorph (CompRMonadMorph f g) h 72 | ≅ 73 | CompRMonadMorph f (CompRMonadMorph g h) 74 | ass {D = D} f g h = RMonadMorphEq _ _ (iext λ _ → Cat.ass D) 75 | 76 | -------------------------------------------------------------------------------- /Monads/EM.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.EM {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open import Library 7 | 8 | open import Functors 9 | open Cat C 10 | open Monad M 11 | 12 | record Alg : Set (a ⊔ b) where 13 | constructor alg 14 | field acar : Obj 15 | astr : ∀ Z → Hom Z acar → Hom (T Z) acar 16 | alaw1 : ∀ {Z}{f : Hom Z acar} → f ≅ comp (astr Z f) η 17 | alaw2 : ∀{Z}{W}{k : Hom Z (T W)}{f : Hom W acar} → 18 | astr Z (comp (astr W f) k) ≅ comp (astr W f) (bind k) 19 | open Alg 20 | 21 | AlgEq : {X Y : Alg} → acar X ≅ acar Y → (astr X ≅ astr Y) → X ≅ Y 22 | AlgEq {alg acar astr _ _} {alg ._ ._ _ _} refl refl = 23 | cong₂ (alg acar astr) 24 | (iext λ _ → iext λ _ → ir _ _) 25 | (iext λ _ → iext λ _ → iext λ _ → iext λ _ → ir _ _) 26 | 27 | record AlgMorph (A B : Alg) : Set (a ⊔ b) where 28 | constructor algmorph 29 | field amor : Hom (acar A) (acar B) 30 | ahom : ∀{Z}{f : Hom Z (acar A)} → 31 | comp amor (astr A Z f) ≅ astr B Z (comp amor f) 32 | open AlgMorph 33 | 34 | AlgMorphEq : {X Y : Alg}{f g : AlgMorph X Y} → amor f ≅ amor g → f ≅ g 35 | AlgMorphEq {f = algmorph amor _}{algmorph .amor _} refl = 36 | cong (algmorph amor) 37 | (iext λ _ → iext λ _ → ir _ _) 38 | 39 | AlgMorphEq' : {X X' Y Y' : Alg} 40 | {f : AlgMorph X Y}{g : AlgMorph X' Y'} → X ≅ X' → Y ≅ Y' → 41 | amor f ≅ amor g → f ≅ g 42 | AlgMorphEq' refl refl = AlgMorphEq 43 | 44 | IdMorph : {A : Alg} → AlgMorph A A 45 | IdMorph {A} = record { 46 | amor = iden; 47 | ahom = λ {Z} {f} → 48 | proof 49 | comp iden (astr A Z f) 50 | ≅⟨ idl ⟩ 51 | astr A Z f 52 | ≅⟨ cong (astr A Z) (sym idl) ⟩ 53 | astr A Z (comp iden f) 54 | ∎} 55 | 56 | CompMorph : {X Y Z : Alg} → AlgMorph Y Z → AlgMorph X Y → AlgMorph X Z 57 | CompMorph {X}{Y}{Z} f g = record { 58 | amor = comp (amor f) (amor g); 59 | ahom = λ{W}{h} → 60 | proof 61 | comp (comp (amor f) (amor g)) (astr X W h) 62 | ≅⟨ ass ⟩ 63 | comp (amor f) (comp (amor g) (astr X W h)) 64 | ≅⟨ cong (comp (amor f)) (ahom g) ⟩ 65 | comp (amor f) (astr Y W (comp (amor g) h)) 66 | ≅⟨ ahom f ⟩ 67 | astr Z W (comp (amor f) (comp (amor g) h)) 68 | ≅⟨ cong (astr Z W) (sym ass) ⟩ 69 | astr Z W (comp (comp (amor f) (amor g)) h) 70 | ∎} 71 | 72 | idlMorph : {X Y : Alg}{f : AlgMorph X Y} → CompMorph IdMorph f ≅ f 73 | idlMorph = AlgMorphEq idl 74 | 75 | idrMorph : {X Y : Alg}{f : AlgMorph X Y} → CompMorph f IdMorph ≅ f 76 | idrMorph = AlgMorphEq idr 77 | 78 | assMorph : {W X Y Z : Alg} 79 | {f : AlgMorph Y Z}{g : AlgMorph X Y}{h : AlgMorph W X} → 80 | CompMorph (CompMorph f g) h ≅ CompMorph f (CompMorph g h) 81 | assMorph = AlgMorphEq ass 82 | 83 | EM : Cat {a ⊔ b}{a ⊔ b} 84 | EM = record{ 85 | Obj = Alg; 86 | Hom = AlgMorph; 87 | iden = IdMorph; 88 | comp = CompMorph; 89 | idl = idlMorph; 90 | idr = idrMorph; 91 | ass = λ{_}{_}{_}{_}{f}{g}{h} → assMorph {f = f}{g}{h}} 92 | -- -} 93 | -------------------------------------------------------------------------------- /Naturals.agda: -------------------------------------------------------------------------------- 1 | module Naturals where 2 | 3 | open import Library 4 | open import Categories 5 | open import Functors 6 | 7 | 8 | open Fun 9 | 10 | record NatT {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}(F G : Fun C D) : 11 | Set (a ⊔ b ⊔ c ⊔ d) where 12 | constructor natural 13 | open Cat 14 | field cmp : ∀ {X} → Hom D (OMap F X) (OMap G X) 15 | nat : ∀{X Y}{f : Hom C X Y} → 16 | comp D (HMap G f) cmp ≅ comp D cmp (HMap F f) 17 | 18 | NatTEq : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{F G : Fun C D} 19 | {α β : NatT F G} → 20 | (λ {X} → NatT.cmp α {X}) ≅ (λ {X} → NatT.cmp β {X}) → 21 | α ≅ β 22 | NatTEq {α = natural cmp _} {natural .cmp _} refl = 23 | cong (natural cmp) (iext λ _ → iext λ _ → iext λ _ → ir _ _) 24 | 25 | idNat : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{F : Fun C D} → NatT F F 26 | idNat {D = D}{F} = let open Cat D in record { 27 | cmp = iden; 28 | nat = λ{X}{Y}{f} → 29 | proof 30 | comp (HMap F f) iden 31 | ≅⟨ idr ⟩ 32 | HMap F f 33 | ≅⟨ sym idl ⟩ 34 | comp iden (HMap F f) ∎} 35 | 36 | compNat : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{F G H : Fun C D} → 37 | NatT G H → NatT F G → NatT F H 38 | compNat {D = D}{F}{G}{H} α β = let open Cat D; open NatT in record { 39 | cmp = comp (cmp α) (cmp β); 40 | nat = λ{X}{Y}{f} → 41 | proof 42 | comp (HMap H f) (comp (cmp α) (cmp β)) 43 | ≅⟨ sym ass ⟩ 44 | comp (comp (HMap H f) (cmp α)) (cmp β) 45 | ≅⟨ cong (λ f₁ → comp f₁ (cmp β)) (nat α) ⟩ 46 | comp (comp (cmp α) (HMap G f)) (cmp β) 47 | ≅⟨ ass ⟩ 48 | comp (cmp α) (comp (HMap G f) (cmp β)) 49 | ≅⟨ cong (comp (cmp α)) (nat β) ⟩ 50 | comp (cmp α) (comp (cmp β) (HMap F f)) 51 | ≅⟨ sym ass ⟩ 52 | comp (comp (cmp α) (cmp β)) (HMap F f) 53 | ∎} 54 | 55 | idlNat : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{F G : Fun C D} 56 | {α : NatT F G} → compNat idNat α ≅ α 57 | idlNat {D = D} = NatTEq (iext λ _ → Cat.idl D) 58 | 59 | idrNat : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{F G : Fun C D} 60 | {α : NatT F G} → compNat α idNat ≅ α 61 | idrNat {D = D} = NatTEq (iext λ _ → Cat.idr D) 62 | 63 | assNat : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{E F G H : Fun C D} 64 | {α : NatT G H}{β : NatT F G}{η : NatT E F} → 65 | compNat (compNat α β) η ≅ compNat α (compNat β η) 66 | assNat {D = D} = NatTEq (iext λ _ → Cat.ass D) 67 | 68 | -- Natural isomorphism 69 | 70 | record Iso {l m}(C : Cat {l}{m}){A B}(f : Cat.Hom C A B) : Set (l ⊔ m) 71 | where 72 | constructor iso 73 | open Cat C 74 | field inv : Hom B A 75 | rinv : comp f inv ≅ iden {B} 76 | linv : comp inv f ≅ iden {A} 77 | 78 | 79 | record NatI {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}(F G : Fun C D) : 80 | Set (a ⊔ b ⊔ c ⊔ d) where 81 | constructor natural 82 | open Cat 83 | field cmp : ∀ {X} → Hom D (OMap F X) (OMap G X) 84 | cmpI : ∀{X} -> Iso D (cmp {X}) 85 | nat : ∀{X Y}{f : Hom C X Y} → 86 | comp D (HMap G f) cmp ≅ comp D cmp (HMap F f) 87 | -------------------------------------------------------------------------------- /Functors/Fin.agda: -------------------------------------------------------------------------------- 1 | module Functors.Fin where 2 | 3 | open import Library 4 | open import Categories.Sets 5 | open import Categories.Setoids 6 | open import Categories 7 | open import Categories.Initial 8 | open import Categories.CoProducts 9 | open import Functors 10 | open import Isomorphism 11 | open import Functors.FullyFaithful 12 | 13 | open Cat 14 | open Fun 15 | open Iso 16 | 17 | Nats : Cat {lzero}{lzero} 18 | Nats = record{ 19 | Obj = ℕ; 20 | Hom = λ m n → Fin m → Fin n; 21 | iden = id; 22 | comp = λ f g → f ∘ g; 23 | idl = refl; 24 | idr = refl; 25 | ass = refl} 26 | 27 | -- initial object 28 | 29 | initN : Init Nats zero 30 | initN = record { 31 | i = λ (); 32 | law = ext λ ()} 33 | 34 | -- coproducts 35 | 36 | extend : ∀ {m n} -> Fin m -> Fin (m + n) 37 | extend zero = zero 38 | extend (suc i) = suc (extend i) 39 | 40 | lift : ∀ m {n} -> Fin n -> Fin (m + n) 41 | lift zero i = i 42 | lift (suc m) i = suc (lift m i) 43 | 44 | case : ∀ (m : ℕ){n : ℕ}{X : Set} → 45 | (Fin m → X) → (Fin n → X) → Fin (m + n) → X 46 | case zero f g i = g i 47 | case (suc m) f g zero = f zero 48 | case (suc m) f g (suc i) = case m (f ∘ suc) g i 49 | 50 | lem1 : ∀ A {B C}(f : Fin A → C) (g : Fin B → C)(i : Fin A) → 51 | case A f g (extend i) ≅ f i 52 | lem1 zero f g () 53 | lem1 (suc A) f g zero = refl 54 | lem1 (suc A) f g (suc i) = lem1 A (f ∘ suc) g i 55 | 56 | lem2 : ∀ A {B C} (f : Fin A → C) (g : Fin B → C)(i : Fin B) → 57 | case A f g (lift A i) ≅ g i 58 | lem2 zero f g zero = refl 59 | lem2 zero f g (suc i) = refl 60 | lem2 (suc A) f g i = lem2 A (f ∘ suc) g i 61 | 62 | lem3 : ∀ A {B C}(f : Fin A → C) (g : Fin B → C) 63 | (h : Fin (A + B) → C) → 64 | (λ x → h (extend {A} x)) ≅ f → 65 | (λ x → h (lift A x)) ≅ g → ∀ i → h i ≅ case A f g i 66 | lem3 zero f g h p q i = fcong i q 67 | lem3 (suc A) f g h p q zero = fcong zero p 68 | lem3 (suc A) f g h p q (suc i) = 69 | lem3 A (f ∘ suc) g (h ∘ suc) (ext (λ i → fcong (suc i) p)) q i 70 | 71 | coprod : CoProd Nats 72 | coprod = record 73 | { _+_ = _+_ 74 | ; inl = extend 75 | ; inr = λ{m} → lift m 76 | ; [_,_] = λ{m} → case m 77 | ; law1 = λ{m} f g → ext (lem1 m f g) 78 | ; law2 = λ{m} f g → ext (lem2 m f g) 79 | ; law3 = λ{m} f g h p q → ext (lem3 m f g h p q) 80 | } 81 | 82 | -- 83 | 84 | FinF : Fun Nats Sets 85 | FinF = record { 86 | OMap = Fin; 87 | HMap = id; 88 | fid = refl; 89 | fcomp = refl} 90 | 91 | FinFoid : Fun Nats Setoids 92 | FinFoid = record { 93 | OMap = λ n → record { 94 | set = Fin n ; 95 | eq = λ i j → i ≅ j; 96 | ref = refl; 97 | sym' = sym; 98 | trn = trans}; 99 | HMap = λ f → record { 100 | fun = f; feq = cong f}; 101 | fid = SetoidFunEq refl (iext λ _ → iext λ _ → ext congid); 102 | fcomp = λ{_ _ _ f g} → 103 | SetoidFunEq refl (iext λ _ → iext λ _ → ext (congcomp f g))} 104 | 105 | FinFF : FullyFaithful FinF 106 | FinFF X Y = record { 107 | fun = id; 108 | inv = id; 109 | law1 = λ _ → refl; 110 | law2 = λ _ → refl} 111 | 112 | open import Data.Bool 113 | 114 | feq : forall {n} -> Fin n -> Fin n -> Bool 115 | feq zero zero = true 116 | feq zero (suc j) = false 117 | feq (suc i) zero = false 118 | feq (suc i) (suc j) = true 119 | -------------------------------------------------------------------------------- /Monads/CatofAdj/TermAdjUniq.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | import Monads.CatofAdj 4 | import Monads.CatofAdj.TermAdjObj 5 | 6 | 7 | module Monads.CatofAdj.TermAdjUniq 8 | {a b} 9 | {C : Cat {a}{b}} 10 | (M : Monad C) 11 | (A : Monads.CatofAdj.ObjAdj M) 12 | (V : Monads.CatofAdj.HomAdj M A (Monads.CatofAdj.TermAdjObj.EMObj M)) 13 | where 14 | 15 | open import Library 16 | open import Functors 17 | open import Adjunctions 18 | open Monads.CatofAdj M 19 | open import Monads.EM M 20 | open Monads.CatofAdj.TermAdjObj M 21 | open import Monads.CatofAdj.TermAdjHom M A 22 | open import Categories.Terminal 23 | open Fun 24 | open Adj 25 | open ObjAdj 26 | open Cat 27 | 28 | omaplem : OMap (HomAdj.K EMHom ) ≅ OMap (HomAdj.K V) 29 | omaplem = ext 30 | λ X → AlgEq (fcong X (cong OMap (HomAdj.Rlaw V))) 31 | ((ext λ Y → 32 | dext 33 | (λ {f} {f'} p → 34 | trans 35 | (trans 36 | (trans 37 | (stripsubst 38 | (λ Z → Hom C Z (OMap (R (adj A)) X)) 39 | (HMap (R (adj A)) (right (adj A) f)) 40 | (fcong Y (cong OMap (law A)))) 41 | (cong' refl (cong 42 | (λ (F : Obj (D A) → Obj C) → 43 | λ (_ : Hom (D A) (OMap (L (adj A)) Y) X) → 44 | Hom C (F (OMap (L (adj A)) Y)) (F X)) 45 | (cong OMap (HomAdj.Rlaw V))) 46 | (icong' refl 47 | (cong 48 | (λ (F : Obj (D A) → Obj C) → 49 | λ (z : Obj (D A)) → 50 | Hom (D A) (OMap (L (adj A)) Y) z → 51 | Hom C (F (OMap (L (adj A)) Y)) (F z)) 52 | (cong OMap (HomAdj.Rlaw V))) 53 | (icong' refl 54 | (cong 55 | (λ (F : Obj (D A) → Obj C) → 56 | λ (z : Obj (D A)) → 57 | {Y₁ : Obj (D A)} → Hom (D A) z Y₁ → Hom C (F z) (F Y₁)) 58 | (cong OMap (HomAdj.Rlaw V))) 59 | (cong HMap (HomAdj.Rlaw V)) (refl {x = OMap (L (adj A)) Y})) 60 | (refl {x = X})) 61 | (refl {x = right (adj A) f}))) 62 | (cong₃ (λ A₁ B → AlgMorph.amor {A₁} {B}) 63 | (fcong Y (cong OMap (HomAdj.Llaw V))) refl 64 | (HomAdj.rightlaw V {Y} {X} {f}))) 65 | (cong (Alg.astr (OMap (HomAdj.K V) X) Y) 66 | (trans 67 | (stripsubst (Hom C Y) f (fcong X (cong OMap (HomAdj.Rlaw V)))) 68 | p))))) 69 | 70 | hmaplem : {X Y : Obj (D A)} (f : Hom (D A) X Y) → 71 | HMap (HomAdj.K EMHom) f ≅ HMap (HomAdj.K V) f 72 | hmaplem {X}{Y} f = AlgMorphEq' 73 | (fcong X omaplem) 74 | (fcong Y omaplem) 75 | (cong' 76 | refl 77 | (cong 78 | (λ (F : Obj (D A) → Obj C) → λ (_ : Hom (D A) X Y) → Hom C (F X) (F Y)) 79 | (cong OMap (HomAdj.Rlaw V))) 80 | (icong' 81 | refl 82 | (cong 83 | (λ (F : Obj (D A) → Obj C) → λ (z : Obj (D A)) → 84 | Hom (D A) X z → Hom C (F X) (F z)) 85 | (cong OMap (HomAdj.Rlaw V))) 86 | (icong' 87 | refl 88 | (cong 89 | (λ (F : Obj (D A) → Obj C) → λ (z : Obj (D A)) → {Y₁ : Obj (D A)} → 90 | Hom (D A) z Y₁ → Hom C (F z) (F Y₁)) 91 | (cong OMap (HomAdj.Rlaw V))) 92 | (cong HMap (HomAdj.Rlaw V)) 93 | (refl {x = X})) 94 | (refl {x = Y})) 95 | (refl {x = f})) 96 | -------------------------------------------------------------------------------- /WeakArrows.agda: -------------------------------------------------------------------------------- 1 | module WeakArrows where 2 | 3 | open import Library 4 | open import Categories 5 | 6 | record Arrow {l m}(J : Cat {l}{m}) : Set (lsuc (l ⊔ m)) where 7 | constructor arrow 8 | open Cat J 9 | field R : Obj → Obj → Set m 10 | pure : ∀{X Y} -> Hom X Y -> R X Y 11 | _<<<_ : ∀{X Y Z} → R Y Z -> R X Y -> R X Z 12 | alaw1 : ∀{X Y Z}(g : Hom Y Z)(f : Hom X Y) -> 13 | pure (comp g f) ≅ pure g <<< pure f 14 | alaw2 : ∀{X Y}(s : R X Y) -> s <<< pure iden ≅ s 15 | alaw3 : ∀{X Y}(r : R X Y) -> pure iden <<< r ≅ r 16 | alaw4 : ∀{W X Y Z}(t : R Y Z)(s : R X Y)(r : R W X) -> 17 | t <<< (s <<< r) ≅ (t <<< s) <<< r 18 | 19 | open import Functors 20 | open import Naturals 21 | open import Categories.Sets 22 | open import YonedaLemma 23 | open import RMonads 24 | 25 | module Arrow2RMonad {l m}(C : Cat {l}{m})(A : Arrow C) where 26 | open Cat C 27 | open Arrow A 28 | 29 | T : Cat.Obj C -> Fun (C Op) (Sets {m}) 30 | T X = functor 31 | (λ Y -> R Y X) 32 | (λ f s -> s <<< pure f) 33 | (ext alaw2) 34 | (ext (λ s -> trans (cong (s <<<_) (alaw1 _ _)) (alaw4 _ _ _))) 35 | 36 | η : {X : Obj} -> NatT HomF[-, X ] (T X) 37 | η = natural pure (ext λ _ -> sym (alaw1 _ _)) 38 | 39 | bind : {X Y : Obj} -> 40 | NatT HomF[-, X ] (T Y) -> NatT (T X) (T Y) 41 | bind α = natural (λ s -> cmp iden <<< s) (ext λ _ -> sym (alaw4 _ _ _)) 42 | where open NatT α 43 | -- cmp iden is one direction of the yoneda lemma 44 | 45 | law1 : {X : Obj} → bind (η {X}) ≅ idNat {F = T X} 46 | law1 = NatTEq (iext (\ _ -> ext alaw3)) 47 | 48 | law2 : {X Y : Obj}{f : NatT HomF[-, X ] (T Y)} → 49 | compNat (bind f) η ≅ f 50 | law2 {f = f} = NatTEq 51 | (iext \ _ -> ext \ s -> trans (fcong iden (nat {f = s})) (cong cmp idl)) 52 | where open NatT f 53 | 54 | law3 : {X Y Z : Obj} 55 | {f : NatT HomF[-, X ] (T Y)} → 56 | {g : NatT HomF[-, Y ] (T Z)} → 57 | bind (compNat (bind g) f) ≅ compNat (bind g) (bind f) 58 | law3 = NatTEq (iext \ W -> ext (\ s -> sym (alaw4 _ _ s))) 59 | 60 | ArrowRMonad : RMonad (y C) 61 | ArrowRMonad = rmonad 62 | T 63 | η 64 | bind 65 | law1 66 | law2 67 | (λ {_ _ _ f g} -> law3 {f = f}{g = g}) 68 | 69 | module RMonad2Arrow {l m}(C : Cat {l}{m})(M : RMonad (y C)) where 70 | open Cat C 71 | open RMonad M 72 | open Fun 73 | open NatT 74 | 75 | R : Obj → Obj → Set m 76 | R X Y = OMap (T Y) X 77 | 78 | pure : {X Y : Obj} → Hom X Y → R X Y 79 | pure f = cmp η f 80 | 81 | _<<<_ : ∀{X Y Z} → R Y Z → R X Y → R X Z 82 | s <<< t = cmp (bind (ylem C (T _) _ s)) t 83 | 84 | alaw1 : ∀{X Y Z}(g : Hom Y Z)(f : Hom X Y) → 85 | pure (comp g f) ≅ (pure g <<< pure f) 86 | alaw1 g f = trans 87 | (sym (fcong g (nat η))) 88 | (sym (fcong f (ifcong _ (cong cmp law2)))) 89 | 90 | alaw2 : ∀{X Y} -> (s : R X Y) → (s <<< pure iden) ≅ s 91 | alaw2 s = trans 92 | (fcong iden (ifcong _ (cong cmp law2))) 93 | (fcong s (fid (T _))) 94 | 95 | alaw3 : ∀{X Y}(r : R X Y) → (pure iden <<< r) ≅ r 96 | alaw3 r = trans 97 | (cong (\ α -> cmp (bind α) r) 98 | (NatTEq (iext \ Z -> ext \ f -> trans 99 | (fcong iden (nat η)) 100 | (cong (cmp η) idl)))) 101 | (fcong r (ifcong _ (cong cmp law1))) 102 | 103 | alaw4 : ∀{W X Y Z}(t : R Y Z)(s : R X Y)(r : R W X) → 104 | (t <<< (s <<< r)) ≅ ((t <<< s) <<< r) 105 | alaw4 t s r = trans 106 | (sym (fcong r (ifcong _ (cong cmp law3)))) 107 | (cong (\ α -> cmp (bind α) r) 108 | (NatTEq (iext \ _ -> ext \ _ -> 109 | sym (fcong s (nat (bind (ylem C (T _) _ t))))))) 110 | 111 | RMonadArrow : Arrow C 112 | RMonadArrow = arrow 113 | R 114 | pure 115 | _<<<_ 116 | alaw1 117 | alaw2 118 | alaw3 119 | alaw4 120 | -------------------------------------------------------------------------------- /RMonads/CatofRAdj/TermRAdj.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | open import RMonads 4 | 5 | module RMonads.CatofRAdj.TermRAdj {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 6 | {J : Fun C D}(M : RMonad J) where 7 | 8 | open import Library 9 | open import RAdjunctions 10 | open import RMonads.CatofRAdj M 11 | open import Categories.Terminal 12 | open import RMonads.REM M 13 | open import RMonads.REM.Adjunction M 14 | open import RAdjunctions.RAdj2RMon 15 | open import RMonads.CatofRAdj.TermRAdjObj M 16 | open import RMonads.CatofRAdj.TermRAdjHom M 17 | 18 | open Cat 19 | open Fun 20 | open RAdj 21 | 22 | omaplem : {X : Obj CatofAdj} {f : Hom CatofAdj X EMObj} → 23 | OMap (HomAdj.K (EMHom {X})) ≅ OMap (HomAdj.K f) 24 | omaplem {A}{f} = ext (λ X → AlgEq 25 | (fcong X (cong OMap (HomAdj.Rlaw f))) 26 | (iext λ Y → 27 | dext 28 | (λ {g} {g'} p → 29 | trans 30 | (trans 31 | (stripsubst (λ Z → Hom D Z (OMap (R (ObjAdj.adj A)) X)) 32 | (HMap (R (ObjAdj.adj A)) (right (ObjAdj.adj A) g)) 33 | (fcong Y (cong OMap (ObjAdj.law A)))) 34 | (cong' refl (ext (λ g₁ → 35 | cong 36 | (λ (F : Obj (ObjAdj.E A) → Obj D) → 37 | Hom D (F (OMap (L (ObjAdj.adj A)) Y)) (F X)) 38 | (cong OMap (HomAdj.Rlaw f)))) 39 | (icong' refl (ext (λ Z → 40 | cong 41 | (λ (F : Obj (ObjAdj.E A) → Obj D) → 42 | Hom (ObjAdj.E A) (OMap (L (ObjAdj.adj A)) Y) Z → 43 | Hom D (F (OMap (L (ObjAdj.adj A)) Y)) (F Z)) 44 | (cong OMap (HomAdj.Rlaw f)))) 45 | (icong' 46 | refl 47 | (ext (λ Z → cong 48 | (λ (F : Obj (ObjAdj.E A) → Obj D) → 49 | {Y₁ : Obj (ObjAdj.E A)} → 50 | Hom (ObjAdj.E A) Z Y₁ → Hom D (F Z) (F Y₁)) 51 | (cong OMap (HomAdj.Rlaw f)))) 52 | (cong HMap (HomAdj.Rlaw f)) 53 | (refl {x = OMap (L (ObjAdj.adj A)) Y})) 54 | (refl {x = X})) 55 | (refl {x = right (ObjAdj.adj A) g}))) 56 | (trans 57 | (cong₃ (λ A1 A2 → RAlgMorph.amor {A1}{A2}) 58 | (fcong Y (cong OMap (HomAdj.Llaw f))) refl 59 | (HomAdj.rightlaw f {Y} {X} {g})) 60 | (cong (RAlg.astr (OMap (HomAdj.K f) X)) 61 | (trans 62 | (stripsubst (Hom D (OMap J Y)) g 63 | (fcong X (cong OMap (HomAdj.Rlaw f)))) 64 | p)))))) 65 | 66 | 67 | hmaplem : {X : Obj CatofAdj} {f : Hom CatofAdj X EMObj} → 68 | {X₁ Y : Obj (ObjAdj.E X)} (f₁ : Hom (ObjAdj.E X) X₁ Y) → 69 | HMap (HomAdj.K (EMHom {X})) f₁ ≅ HMap (HomAdj.K f) f₁ 70 | hmaplem {A}{V}{X}{Y} f = lemZ 71 | (fcong X (omaplem {A} {V})) 72 | (fcong Y (omaplem {A} {V})) 73 | (cong' 74 | refl 75 | (ext (λ Z → cong 76 | (λ (F : Obj (ObjAdj.E A) → Obj D) → Hom D (F X) (F Y)) 77 | (cong OMap (HomAdj.Rlaw V)))) 78 | (icong' 79 | refl 80 | (ext (λ Z → cong 81 | (λ (F : Obj (ObjAdj.E A) → Obj D) → 82 | Hom (ObjAdj.E A) X Z → Hom D (F X) (F Z)) 83 | (cong OMap (HomAdj.Rlaw V)))) 84 | (icong' 85 | refl 86 | (ext (λ Z → cong 87 | (λ (F : Obj (ObjAdj.E A) → Obj D) → 88 | {Y₁ : Obj (ObjAdj.E A)} → 89 | Hom (ObjAdj.E A) Z Y₁ → Hom D (F Z) (F Y₁)) 90 | (cong OMap (HomAdj.Rlaw V)))) 91 | (cong HMap (HomAdj.Rlaw V)) 92 | (refl {x = X})) 93 | (refl {x = Y})) 94 | (refl {x = f})) 95 | 96 | 97 | uniq : {X : Obj CatofAdj} {f : Hom CatofAdj X EMObj} → 98 | EMHom {X} ≅ f 99 | uniq {X} {f} = HomAdjEq _ _ (FunctorEq _ _ 100 | (omaplem {X} {f}) 101 | (iext λ _ → iext λ _ → ext (hmaplem {X}{f}))) 102 | 103 | EMIsTerm : Term CatofAdj EMObj 104 | EMIsTerm = record { 105 | t = EMHom; 106 | law = uniq} 107 | -------------------------------------------------------------------------------- /RMonads/REM.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | open import RMonads 4 | 5 | module RMonads.REM {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}}{J : Fun C D} 6 | (M : RMonad J) where 7 | open import Library 8 | open RMonad M 9 | open Fun 10 | 11 | record RAlg : Set (a ⊔ c ⊔ d) where 12 | constructor ralg 13 | open Cat D 14 | field acar : Obj 15 | astr : ∀ {Z} → Hom (OMap J Z) acar → Hom (T Z) acar 16 | alaw1 : ∀ {Z}{f : Hom (OMap J Z) acar} → 17 | f ≅ comp (astr f) η 18 | alaw2 : ∀{Z}{W}{k : Hom (OMap J Z) (T W)} 19 | {f : Hom (OMap J W) acar} → 20 | astr (comp (astr f) k) ≅ comp (astr f) (bind k) 21 | 22 | AlgEq : {X Y : RAlg} → RAlg.acar X ≅ RAlg.acar Y → 23 | ((λ {Z} → RAlg.astr X {Z}) ≅ (λ {Z} → RAlg.astr Y {Z})) → 24 | X ≅ Y 25 | AlgEq {ralg acar astr _ _}{ralg .acar .astr _ _} refl refl = let open Cat in 26 | cong₂ (ralg acar astr) 27 | (iext λ _ → iext λ _ → ir _ _) 28 | (iext λ _ → iext λ _ → iext λ _ → iext λ _ → ir _ _) 29 | 30 | astrnat : ∀(alg : RAlg){X Y} 31 | (f : Cat.Hom C X Y) → 32 | (g : Cat.Hom D (OMap J X) (RAlg.acar alg)) 33 | (g' : Cat.Hom D (OMap J Y) (RAlg.acar alg)) → 34 | Cat.comp D g' (HMap J f) ≅ g → 35 | Cat.comp D (RAlg.astr alg g') 36 | (RMonad.bind M (Cat.comp D (RMonad.η M) (HMap J f))) 37 | ≅ 38 | RAlg.astr alg g 39 | astrnat alg f g g' p = let 40 | open RAlg alg; open Cat D in 41 | proof 42 | comp (astr g') (bind (comp η (HMap J f))) 43 | ≅⟨ sym alaw2 ⟩ 44 | astr (comp (astr g') (comp η (HMap J f))) 45 | ≅⟨ cong astr (sym ass) ⟩ 46 | astr (comp (comp (astr g') η) (HMap J f)) 47 | ≅⟨ cong (λ g₁ → astr (comp g₁ (HMap J f))) (sym alaw1) ⟩ 48 | astr (comp g' (HMap J f)) 49 | ≅⟨ cong astr p ⟩ 50 | astr g ∎ 51 | 52 | 53 | record RAlgMorph (A B : RAlg) : Set (a ⊔ c ⊔ d) 54 | where 55 | constructor ralgmorph 56 | open Cat D 57 | open RAlg 58 | field amor : Hom (acar A) (acar B) 59 | ahom : ∀{Z}{f : Hom (OMap J Z) (acar A)} → 60 | comp amor (astr A f) ≅ astr B (comp amor f) 61 | open RAlgMorph 62 | 63 | RAlgMorphEq : ∀{X Y : RAlg}{f g : RAlgMorph X Y} → amor f ≅ amor g → f ≅ g 64 | RAlgMorphEq {X}{Y}{ralgmorph amor _}{ralgmorph .amor _} refl = 65 | cong (ralgmorph amor) (iext λ _ → iext λ _ → ir _ _) 66 | 67 | lemZ : ∀{X X' Y Y' : RAlg} 68 | {f : RAlgMorph X Y}{g : RAlgMorph X' Y'} → X ≅ X' → Y ≅ Y' → 69 | amor f ≅ amor g → f ≅ g 70 | lemZ refl refl = RAlgMorphEq 71 | 72 | IdMorph : ∀{A : RAlg} → RAlgMorph A A 73 | IdMorph {A} = let open Cat D; open RAlg A in record { 74 | amor = iden; 75 | ahom = λ {_ f} → 76 | proof 77 | comp iden (astr f) 78 | ≅⟨ idl ⟩ 79 | astr f 80 | ≅⟨ cong astr (sym idl) ⟩ 81 | astr (comp iden f) 82 | ∎} 83 | 84 | CompMorph : ∀{X Y Z : RAlg} → 85 | RAlgMorph Y Z → RAlgMorph X Y → RAlgMorph X Z 86 | CompMorph {X}{Y}{Z} f g = let open Cat D; open RAlg in record { 87 | amor = comp (amor f) (amor g); 88 | ahom = λ {_ f'} → 89 | proof 90 | comp (comp (amor f) (amor g)) (astr X f') 91 | ≅⟨ ass ⟩ 92 | comp (amor f) (comp (amor g) (astr X f')) 93 | ≅⟨ cong (comp (amor f)) (ahom g) ⟩ 94 | comp (amor f) (astr Y (comp (amor g) f')) 95 | ≅⟨ ahom f ⟩ 96 | astr Z (comp (amor f) (comp (amor g) f')) 97 | ≅⟨ cong (astr Z) (sym ass) ⟩ 98 | astr Z (comp (comp (amor f) (amor g)) f') 99 | ∎} 100 | 101 | idlMorph : {X Y : RAlg}{f : RAlgMorph X Y} → 102 | CompMorph IdMorph f ≅ f 103 | idlMorph = RAlgMorphEq (Cat.idl D) 104 | 105 | idrMorph : ∀{X Y : RAlg}{f : RAlgMorph X Y} → 106 | CompMorph f IdMorph ≅ f 107 | idrMorph = RAlgMorphEq (Cat.idr D) 108 | 109 | assMorph : ∀{W X Y Z : RAlg} 110 | {f : RAlgMorph Y Z}{g : RAlgMorph X Y}{h : RAlgMorph W X} → 111 | CompMorph (CompMorph f g) h ≅ CompMorph f (CompMorph g h) 112 | assMorph = RAlgMorphEq (Cat.ass D) 113 | 114 | EM : Cat 115 | EM = record{ 116 | Obj = RAlg; 117 | Hom = RAlgMorph; 118 | iden = IdMorph; 119 | comp = CompMorph; 120 | idl = idlMorph; 121 | idr = idrMorph; 122 | ass = λ{_ _ _ _ f g h} → assMorph {f = f}{g}{h}} 123 | 124 | -------------------------------------------------------------------------------- /MonoidalCat.agda: -------------------------------------------------------------------------------- 1 | module MonoidalCat where 2 | 3 | open import Library hiding (_×_) 4 | open import Categories 5 | open import Categories.Products 6 | open import Functors 7 | open import Naturals 8 | 9 | record Monoidal {l}{m} : Set (lsuc (l ⊔ m)) where 10 | field C : Cat {l}{m} 11 | open Cat C 12 | open Fun 13 | open NatI 14 | field ⊗ : Fun (C × C) C 15 | I : Obj 16 | 17 | I⊗- : Fun C C 18 | I⊗- = functor 19 | (\ A -> OMap ⊗ (I , A)) 20 | (\ f -> HMap ⊗ (iden , f)) 21 | (fid ⊗) 22 | (trans (cong (\f -> HMap ⊗ (f , _)) (sym idl)) (fcomp ⊗)) 23 | 24 | field λ' : NatI I⊗- (IdF C) 25 | 26 | -⊗I : Fun C C 27 | -⊗I = functor 28 | (\ A -> OMap ⊗ (A , I)) 29 | (\ f -> HMap ⊗ (f , iden)) 30 | (fid ⊗) 31 | (trans (cong (\f -> HMap ⊗ (_ , f)) (sym idl)) (fcomp ⊗)) 32 | 33 | field ρ : NatI -⊗I (IdF C) 34 | 35 | [-⊗-]⊗- : Fun (C × C × C) C 36 | [-⊗-]⊗- = functor 37 | (λ {(X , Y , Z) → OMap ⊗ (OMap ⊗ (X , Y) , Z)}) 38 | (λ {(f , g , h) → HMap ⊗ (HMap ⊗ (f , g) , h)}) 39 | (trans (cong (\f -> HMap ⊗ (f , _)) (fid ⊗) ) (fid ⊗)) 40 | (trans (cong (\f -> HMap ⊗ (f , _)) (fcomp ⊗)) (fcomp ⊗)) 41 | 42 | -⊗[-⊗-] : Fun (C × C × C) C 43 | -⊗[-⊗-] = functor 44 | (λ {(X , Y , Z) → OMap ⊗ (X , OMap ⊗ (Y , Z))}) 45 | (λ {(f , g , h) → HMap ⊗ (f , HMap ⊗ (g , h))}) 46 | (trans (cong (\f -> HMap ⊗ (_ , f)) (fid ⊗) ) (fid ⊗)) 47 | (trans (cong (\f -> HMap ⊗ (_ , f)) (fcomp ⊗)) (fcomp ⊗)) 48 | 49 | field α : NatI [-⊗-]⊗- -⊗[-⊗-] 50 | 51 | triangle : ∀{A B} -> 52 | comp (HMap ⊗ (iden {A} , (cmp λ' {B}))) (cmp α {A , I , B}) 53 | ≅ HMap ⊗ (cmp ρ {A} , iden {B}) 54 | 55 | square : ∀{A B C D} -> 56 | comp (HMap ⊗ (iden {A} , cmp α {B , C , D})) 57 | (comp (cmp α {A , OMap ⊗ (B , C) , D}) 58 | (HMap ⊗ (cmp α {A , B , C} , iden {D}))) 59 | ≅ 60 | comp (cmp α {A , B , OMap ⊗ (C , D)}) 61 | (cmp α {OMap ⊗ (A , B) , C , D}) 62 | 63 | record MonoidalFun {a b c d}(M : Monoidal {a}{b})(M' : Monoidal {c}{d}) 64 | : Set (a ⊔ b ⊔ c ⊔ d) where 65 | open Monoidal 66 | open Cat 67 | 68 | field F : Fun (C M) (C M') 69 | 70 | open Fun 71 | 72 | field e : Hom (C M') (I M') (OMap F (I M)) 73 | 74 | F-⊗'F- : Fun (C M × C M) (C M') 75 | F-⊗'F- = functor 76 | (\ {(A , B) -> OMap (⊗ M') (OMap F A , OMap F B)}) 77 | (\ {(f , g) -> HMap (⊗ M') (HMap F f , HMap F g)}) 78 | (trans (cong₂ (\f g -> HMap (⊗ M') (f , g)) (fid F) (fid F)) (fid (⊗ M'))) 79 | (trans (cong₂ (\f g -> HMap (⊗ M') (f , g)) 80 | (fcomp F) (fcomp F)) (fcomp (⊗ M'))) 81 | 82 | F[-⊗-] : Fun (C M × C M) (C M') 83 | F[-⊗-] = functor 84 | (λ { (A , B) → OMap F (OMap (⊗ M) (A , B)) }) 85 | ((λ { (f , g) → HMap F (HMap (⊗ M) (f , g)) })) 86 | (trans (cong (HMap F) (fid (⊗ M))) (fid F)) 87 | (trans (cong (HMap F) (fcomp (⊗ M))) (fcomp F)) 88 | 89 | 90 | field m : NatT F-⊗'F- F[-⊗-] 91 | 92 | field square1 : ∀{A} -> 93 | NatI.cmp (ρ M') {OMap F A} 94 | ≅ 95 | comp (C M') (HMap F (NatI.cmp (ρ M) {A})) 96 | (comp (C M') (NatT.cmp m {A , I M}) 97 | (HMap (⊗ M') (iden (C M') {OMap F A} , e))) 98 | 99 | square2 : ∀{B} -> 100 | NatI.cmp (λ' M') {OMap F B} 101 | ≅ 102 | comp (C M') (HMap F (NatI.cmp (λ' M) {B})) 103 | (comp (C M') (NatT.cmp m {I M , B}) 104 | (HMap (⊗ M') (e , iden (C M') {OMap F B}))) 105 | 106 | 107 | hexagon : ∀{A B B'} -> 108 | comp (C M') (HMap F (NatI.cmp (α M) {A , B , B'})) 109 | (comp (C M') (NatT.cmp m {OMap (⊗ M) (A , B) , B'}) 110 | (HMap (⊗ M') (NatT.cmp m {A , B} , 111 | iden (C M') {OMap F B'}))) 112 | ≅ 113 | comp (C M') (NatT.cmp m {A , OMap (⊗ M) (B , B')}) 114 | (comp (C M') (HMap (⊗ M') (iden (C M') {OMap F A} , 115 | NatT.cmp m {B , B'})) 116 | (NatI.cmp (α M') {OMap F A , 117 | OMap F B , 118 | OMap F B'})) 119 | -------------------------------------------------------------------------------- /Library.agda: -------------------------------------------------------------------------------- 1 | module Library where 2 | 3 | open import Function using (id; _∘_; _$_) public 4 | open import Relation.Binary.HeterogeneousEquality public 5 | open ≅-Reasoning renaming (begin_ to proof_) public 6 | open import Data.Product renaming (proj₁ to fst; proj₂ to snd) public 7 | open import Data.Empty public using (⊥; ⊥-elim) 8 | open import Data.Unit public using (⊤) 9 | open import Data.Nat public using (ℕ; zero; suc; _+_; module ℕ) 10 | open import Data.Nat.Properties.Simple public 11 | open import Data.Fin public using (Fin; zero; suc; raise) 12 | renaming (_+_ to _F+_; fromℕ to fromNat) 13 | open import Level public renaming (suc to lsuc; zero to lzero) hiding (lift) 14 | 15 | -- needed for setoids 16 | congid : ∀{a}{A : Set a}{a a' : A}(p : a ≅ a') → cong id p ≅ p 17 | congid refl = refl 18 | 19 | congcomp : ∀{a b c}{A : Set a}{B : Set b}{C : Set c} 20 | {a a' : A}(f : B → C)(g : A → B)(p : a ≅ a') → 21 | cong (f ∘ g) p ≅ cong f (cong g p) 22 | congcomp f g refl = refl 23 | 24 | -- should be replaced by dcong 25 | cong' : ∀{a b}{A A' : Set a} → A ≅ A' → 26 | {B : A → Set b}{B' : A' → Set b} → B ≅ B' → 27 | {f : ∀ a → B a}{f' : ∀ a → B' a} → f ≅ f' → 28 | {a : A}{a' : A'} → a ≅ a' → f a ≅ f' a' 29 | cong' refl refl refl refl = refl 30 | 31 | -- should be replaced by dicong 32 | icong' : ∀{a b}{A A' : Set a} → A ≅ A' → 33 | {B : A → Set b}{B' : A' → Set b} → B ≅ B' → 34 | {f : ∀ {a} → B a}{f' : ∀ {a} → B' a} → 35 | (λ {a} → f {a}) ≅ (λ {a} → f' {a}) → 36 | {a : A}{a' : A'} → a ≅ a' → f {a} ≅ f' {a'} 37 | icong' refl refl refl refl = refl 38 | 39 | fcong : ∀{a b}{A : Set a}{B : A → Set b}{f f' : (x : A) → B x} 40 | (a : A) → f ≅ f' → f a ≅ f' a 41 | fcong a refl = refl 42 | 43 | dcong : ∀{a b}{A A' : Set a}{B : A → Set b}{B' : A' → Set b} 44 | {f : (a : A) → B a}{f' : (a : A') → B' a}{a : A}{a' : A'} → 45 | a ≅ a' → B ≅ B' → f ≅ f' → f a ≅ f' a' 46 | dcong refl refl refl = refl 47 | 48 | dicong : ∀{a b}{A A' : Set a}{B : A → Set b}{B' : A' → Set b} 49 | {f : ∀ {a} → B a}{f' : ∀ {a} → B' a} → {a : A}{a' : A'} → 50 | a ≅ a' → B ≅ B' → 51 | (λ {a} → f {a}) ≅ (λ {a} → f' {a}) → 52 | f {a} ≅ f' {a'} 53 | dicong refl refl refl = refl 54 | 55 | ifcong : ∀{a b}{A : Set a}{B : A → Set b}{f f' : {x : A} → B x}(a : A) → 56 | _≅_ {_}{ {x : A} → B x} f { {x : A} → B x} f' → f {a} ≅ f' {a} 57 | ifcong a refl = refl 58 | 59 | cong₃ : ∀{a b c d}{A : Set a}{B : A → Set b} 60 | {C : ∀ x → B x → Set c }{D : ∀ x y → C x y → Set d} 61 | (f : ∀ x y z → D x y z) 62 | {a a' : A} → a ≅ a' → 63 | {b : B a}{b' : B a'} → b ≅ b' → 64 | {c : C a b}{c' : C a' b'} → c ≅ c' → 65 | f a b c ≅ f a' b' c' 66 | cong₃ f refl refl refl = refl 67 | 68 | ir : ∀{a}{A B : Set a}{x : A}{y : B}(p q : x ≅ y) → p ≅ q 69 | ir refl refl = refl 70 | 71 | stripsubst : ∀{a c}{A : Set a}(C : A → Set c) → 72 | {a : A} → (c : C a) → 73 | {b : A} → (p : a ≅ b) → 74 | subst C p c ≅ c 75 | stripsubst C c refl = refl 76 | 77 | postulate ext : ∀{a b}{A : Set a}{B B' : A → Set b} 78 | {f : ∀ a → B a}{g : ∀ a → B' a} → 79 | (∀ a → f a ≅ g a) → f ≅ g 80 | 81 | postulate dext : ∀{a b}{A A' : Set a}{B : A → Set b}{B' : A' → Set b} 82 | {f : ∀ a → B a}{g : ∀ a → B' a} → 83 | (∀ {a a'} → a ≅ a' → f a ≅ g a') → f ≅ g 84 | 85 | -- this could just be derived from ext 86 | 87 | postulate iext : ∀{a b}{A : Set a}{B B' : A → Set b} 88 | {f : ∀ {a} → B a}{g : ∀{a} → B' a} → 89 | (∀ a → f {a} ≅ g {a}) → 90 | _≅_ {_}{ {a : A} → B a} f { {a : A} → B' a} g 91 | 92 | postulate diext : ∀{a b}{A A' : Set a}{B : A → Set b}{B' : A' → Set b} 93 | {f : ∀ {a} → B a}{f' : ∀{a'} → B' a'} → 94 | (∀{a a'} → a ≅ a' → f {a} ≅ f' {a'}) → 95 | _≅_ {_}{ {a : A} → B a} f { {a' : A'} → B' a'} f' 96 | 97 | hir : ∀{a}{A A' A'' A''' : Set a}{a : A}{a' : A'}{a'' : A''}{a''' : A'''} 98 | {p : a ≅ a'}{q : a'' ≅ a'''} → a ≅ a'' → p ≅ q 99 | hir {p = refl} {q = refl} refl = refl 100 | 101 | hir' : ∀{a}{A A' A'' A''' : Set a}{a : A}{a' : A'}{a'' : A''}{a''' : A'''} 102 | {p : a ≅ a'}{q : a'' ≅ a'''} → a' ≅ a''' → p ≅ q 103 | hir' {p = refl} {q = refl} refl = refl 104 | -------------------------------------------------------------------------------- /Monads/CatofAdj/TermAdjHom.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | open import Level 4 | import Monads.CatofAdj 5 | 6 | module Monads.CatofAdj.TermAdjHom 7 | {c d} 8 | {C : Cat {c}{d}} 9 | (M : Monad C) 10 | (A : Cat.Obj (Monads.CatofAdj.CatofAdj M {c ⊔ d}{c ⊔ d})) where 11 | 12 | open import Library 13 | 14 | open import Functors 15 | open import Adjunctions 16 | open import Monads.EM M 17 | open Monads.CatofAdj M 18 | open import Monads.CatofAdj.TermAdjObj M 19 | 20 | 21 | open Cat 22 | open Fun 23 | open Adj 24 | open Monad 25 | 26 | open ObjAdj 27 | 28 | K' : Fun (D A) (D EMObj) 29 | K' = record { 30 | OMap = λ X → record { 31 | acar = OMap (R (adj A)) X; 32 | astr = λ Z f → subst (λ Z → Hom C Z (OMap (R (adj A)) X)) 33 | (fcong Z (cong OMap (law A))) 34 | (HMap (R (adj A)) (right (adj A) f)); 35 | alaw1 = λ {Z} {f} → alaw1lem 36 | (TFun M) 37 | (L (adj A)) 38 | (R (adj A)) 39 | (law A) 40 | (η M) 41 | (right (adj A)) 42 | (left (adj A)) 43 | (ηlaw A) 44 | (natleft (adj A) (iden C) (right (adj A) f) (iden (D A))) 45 | (lawb (adj A) f); 46 | alaw2 = λ {Z} {W} {k} {f} → alaw2lem 47 | (TFun M) 48 | (L (adj A)) 49 | (R (adj A)) 50 | (law A) 51 | (right (adj A)) 52 | (bind M) 53 | (natright (adj A)) 54 | (bindlaw A {_}{_}{k})}; 55 | HMap = λ {X} {Y} f → record { 56 | amor = HMap (R (adj A)) f; 57 | ahom = λ {Z} {g} → ahomlem 58 | (TFun M) 59 | (L (adj A)) 60 | (R (adj A)) 61 | (law A) 62 | (right (adj A)) 63 | (natright (adj A))}; 64 | fid = AlgMorphEq (fid (R (adj A))); 65 | fcomp = AlgMorphEq (fcomp (R (adj A)))} 66 | 67 | Llaw' : K' ○ L (adj A) ≅ L (adj EMObj) 68 | Llaw' = FunctorEq 69 | _ 70 | _ 71 | (ext (λ X → AlgEq 72 | (fcong X (cong OMap (law A))) 73 | ((ext λ Z → dext (λ {f} {f'} p → Llawlem 74 | (TFun M) 75 | (L (adj A)) 76 | (R (adj A)) 77 | (law A) 78 | (right (adj A)) 79 | (bind M) 80 | (bindlaw A) 81 | p))))) 82 | (iext λ X → iext λ Y → ext λ f → AlgMorphEq' 83 | (AlgEq 84 | (fcong X (cong OMap (law A))) 85 | ((ext λ Z → dext (λ {f₁} {f'} p → Llawlem 86 | (TFun M) 87 | (L (adj A)) 88 | (R (adj A)) 89 | (law A) 90 | (right (adj A)) 91 | (bind M) 92 | (bindlaw A) 93 | p)))) 94 | (AlgEq 95 | (fcong Y (cong OMap (law A))) 96 | ((ext λ Z → dext (λ {f₁} {f'} p → Llawlem 97 | (TFun M) 98 | (L (adj A)) 99 | (R (adj A)) 100 | (law A) 101 | (right (adj A)) 102 | (bind M) 103 | (bindlaw A) 104 | p)))) 105 | (dcong 106 | (refl {x = f}) 107 | (ext (λ _ → cong₂ 108 | (Hom C) 109 | (fcong X (cong OMap (law A))) 110 | (fcong Y (cong OMap (law A))))) 111 | (dicong 112 | (refl {x = Y}) 113 | (ext (λ z → cong 114 | (λ F → Hom C X z → Hom C (F X) (F z)) 115 | (cong OMap (law A)))) 116 | (dicong 117 | (refl {x = X}) 118 | (ext (λ z → cong 119 | (λ F → ∀ {y} → Hom C z y → Hom C (F z) (F y)) 120 | (cong OMap (law A)))) 121 | (cong HMap (law A)))))) 122 | 123 | Rlaw' : R (adj A) ≅ R (adj EMObj) ○ K' 124 | Rlaw' = FunctorEq _ _ refl refl 125 | 126 | rightlaw' : {X : Obj C} {Y : Obj (D A)} {f : Hom C X (OMap (R (adj A)) Y)} → 127 | HMap K' (right (adj A) f) 128 | ≅ 129 | right (adj EMObj) 130 | {X} 131 | {OMap K' Y} 132 | (subst (Hom C X) (fcong Y (cong OMap Rlaw')) f) 133 | rightlaw' {X = X}{Y = Y}{f = f} = AlgMorphEq' 134 | (AlgEq (fcong X (cong OMap (law A))) 135 | (ext λ Z → dext (λ p → Llawlem (TFun M) 136 | (L (adj A)) 137 | (R (adj A)) 138 | (law A) 139 | (right (adj A)) 140 | (bind M) 141 | (bindlaw A) p ))) 142 | refl 143 | (trans (cong (λ (f₁ : Hom C X (OMap (R (adj A)) Y)) → 144 | HMap (R (adj A)) (right (adj A) f₁)) (sym (stripsubst (Hom C X) f (fcong Y (cong OMap Rlaw')))) ) (sym (stripsubst (λ Z → Hom C Z (OMap (R (adj A)) Y)) ( ((HMap (R (adj A)) 145 | (right (adj A) 146 | (subst (Hom C X) (fcong Y (cong (λ r → OMap r) Rlaw')) f))))) (fcong X (cong OMap (law A)))))) 147 | 148 | EMHom : Hom CatofAdj A EMObj 149 | EMHom = record { 150 | K = K'; 151 | Llaw = Llaw'; 152 | Rlaw = Rlaw'; 153 | rightlaw = rightlaw'} 154 | 155 | -------------------------------------------------------------------------------- /Lawvere.agda: -------------------------------------------------------------------------------- 1 | module Lawvere where 2 | 3 | open import Library 4 | open import Data.Sum 5 | open import Categories 6 | open import Categories.Sets 7 | open import Categories.Initial 8 | open import Categories.PushOuts 9 | open import Categories.Products hiding (_×_) 10 | open import Categories.CoProducts 11 | open import Categories.Terminal 12 | 13 | open import Functors 14 | open import Functors.Fin 15 | 16 | record Lawvere {a}{b} : Set (lsuc (a ⊔ b)) where 17 | constructor lawvere 18 | field T : Cat {a}{b} 19 | L : Fun (Nats Op) T 20 | L0 : Term T (Fun.OMap L zero) 21 | LP : ∀ m n → Prod T (Fun.OMap L m) (Fun.OMap L n) 22 | 23 | -- it's not the identity, it switches some implicit in fid and fcomp I think 24 | FunOp : ∀{a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} → 25 | Fun C D → Fun (C Op) (D Op) 26 | FunOp (functor OMap HMap fid fcomp) = functor OMap HMap fid fcomp 27 | 28 | LSet : Lawvere 29 | LSet = lawvere 30 | (Sets Op) 31 | (FunOp FinF) 32 | (term (λ ()) (ext λ ())) 33 | λ m n → prod 34 | (Fin m ⊎ Fin n) 35 | inj₁ 36 | inj₂ 37 | [_,_]′ 38 | refl 39 | refl 40 | λ p q → ext (λ{ (inj₁ a) -> fcong a p; (inj₂ b) -> fcong b q}) 41 | 42 | open import RMonads 43 | open import RMonads.RKleisli 44 | open import RMonads.RKleisli.Functors 45 | 46 | lem : RMonad FinF → Lawvere {lzero}{lzero} 47 | lem M = lawvere 48 | (Kl M Op) 49 | (FunOp (RKlL M) ) 50 | (term (λ ()) (ext λ ())) 51 | λ m n → prod 52 | (m + n) 53 | (η ∘ extend) 54 | (η ∘ lift m) 55 | (case m) 56 | (ext λ i → trans (fcong (extend i) law2) (lem1 m _ _ i)) 57 | (ext λ i → trans (fcong (lift m i) law2) (lem2 m _ _ i)) 58 | λ {o} {f} {g} {h} p q → ext (lem3 59 | m 60 | f 61 | g 62 | h 63 | (trans (ext λ i → sym (fcong (extend i) law2)) p) 64 | (trans (ext λ i → sym (fcong (lift m i) law2)) q)) 65 | where open RMonad M 66 | 67 | lem-1 : Lawvere {lzero}{lzero} → RMonad FinF 68 | lem-1 LT = rmonad 69 | (λ n → Cat.Hom (Lawvere.T LT) (Fun.OMap (Lawvere.L LT) 1) (Fun.OMap (Lawvere.L LT) n)) 70 | {!!} 71 | {!!} 72 | {!!} 73 | {!!} 74 | {!!} 75 | 76 | 77 | open import Categories.Products 78 | 79 | record Model {a}{b}{c}{d} (Law : Lawvere {a}{b}) : Set (lsuc (a ⊔ b ⊔ c ⊔ d)) 80 | where 81 | open Lawvere Law 82 | field C : Cat {c}{d} 83 | F : Fun T C 84 | F0 : Term C (Fun.OMap F (Fun.OMap L zero)) 85 | FP : ∀ m n → Prod C (Fun.OMap F (Fun.OMap L m)) 86 | (Fun.OMap F (Fun.OMap L n)) 87 | open import RMonads.REM 88 | open import RMonads.CatofRAdj.InitRAdj 89 | open import RMonads.CatofRAdj.TermRAdjObj 90 | open import RMonads.REM.Functors 91 | 92 | model : (T : RMonad FinF) → Model (lem T) 93 | model T = record { 94 | C = EM T Op ; 95 | F = FunOp (K' T (EMObj T)); 96 | F0 = term (λ{alg} → ralgmorph (RAlg.astr alg {0} (λ ())) 97 | (λ {n}{f} → 98 | sym $ RAlg.alaw2 alg {n}{zero}{f}{λ ()} )) 99 | (λ{alg}{f} → RAlgMorphEq T (ext (λ t → trans 100 | (trans (cong (λ f₁ → RAlg.astr alg f₁ t) (ext (λ ()))) 101 | (sym (fcong t (RAlgMorph.ahom f {0}{RMonad.η T})))) 102 | (cong (RAlgMorph.amor f) (fcong t (RMonad.law1 T)))))); 103 | FP = λ m n → prod 104 | (Fun.OMap (REML FinF T) (m + n) ) 105 | (Fun.HMap (REML FinF T) extend) 106 | (Fun.HMap (REML FinF T) (lift m)) 107 | (λ{alg} f g → ralgmorph 108 | (RAlg.astr alg 109 | (case m (RAlgMorph.amor f ∘ RMonad.η T) 110 | (RAlgMorph.amor g ∘ RMonad.η T))) 111 | (sym (RAlg.alaw2 alg))) 112 | (λ {alg}{f}{g} → RAlgMorphEq T (trans 113 | (sym (RAlg.alaw2 alg)) 114 | (trans (cong (RAlg.astr alg) 115 | (ext λ i → trans (fcong (extend i) (sym (RAlg.alaw1 alg))) 116 | (lem1 m _ _ i))) 117 | (trans (sym (RAlgMorph.ahom f)) 118 | (ext λ i → cong (RAlgMorph.amor f) 119 | (fcong i (RMonad.law1 T))))))) 120 | (λ {alg}{f}{g} → RAlgMorphEq T (trans 121 | (sym (RAlg.alaw2 alg)) 122 | (trans (cong (RAlg.astr alg) 123 | (ext λ i → trans (fcong (lift m i) (sym (RAlg.alaw1 alg))) 124 | (lem2 m _ _ i))) 125 | (trans (sym (RAlgMorph.ahom g)) 126 | (ext λ i → cong (RAlgMorph.amor g) 127 | (fcong i (RMonad.law1 T))))))) 128 | λ{alg}{f}{g}{h} p q → RAlgMorphEq T (trans 129 | (trans (ext λ i → cong (RAlgMorph.amor h) 130 | (sym (fcong i (RMonad.law1 T)))) 131 | (RAlgMorph.ahom h)) 132 | (cong (RAlg.astr alg) (ext (lem3 133 | m 134 | (RAlgMorph.amor f ∘ RMonad.η T) 135 | (RAlgMorph.amor g ∘ RMonad.η T) 136 | (RAlgMorph.amor h ∘ RMonad.η T) 137 | (ext λ i → trans 138 | (cong (RAlgMorph.amor h) (sym (fcong i (RMonad.law2 T)))) 139 | (fcong (RMonad.η T i) (cong RAlgMorph.amor p))) 140 | (ext λ i → trans 141 | (cong (RAlgMorph.amor h) (sym (fcong i (RMonad.law2 T)))) 142 | (fcong (RMonad.η T i) (cong RAlgMorph.amor q)))))))} 143 | -------------------------------------------------------------------------------- /Monads/CatofAdj/TermAdjObj.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Monads 3 | 4 | module Monads.CatofAdj.TermAdjObj {a b}{C : Cat {a}{b}}(M : Monad C) where 5 | 6 | open import Library 7 | open import Functors 8 | open import Naturals 9 | open import Adjunctions 10 | open import Monads.CatofAdj M 11 | open import Categories.Terminal 12 | open import Monads.EM M 13 | open import Monads.EM.Adjunction M 14 | open import Adjunctions.Adj2Mon 15 | 16 | open Cat 17 | open Fun 18 | open Monad M 19 | open NatT 20 | open Adj 21 | 22 | lemX : R EMAdj ○ L EMAdj ≅ TFun M 23 | lemX = FunctorEq _ _ refl refl 24 | 25 | EMObj : Obj CatofAdj 26 | EMObj = record { 27 | D = EM; 28 | adj = EMAdj; 29 | law = lemX; 30 | ηlaw = idl C; 31 | bindlaw = λ{X Y f} → 32 | cong bind 33 | (stripsubst (Hom C X) f (fcong Y (cong OMap (sym lemX))))} 34 | 35 | open ObjAdj 36 | open Adj 37 | 38 | alaw1lem : ∀{c d}{D : Cat {c}{d}} 39 | (T : Fun C C)(L : Fun C D)(R : Fun D C) 40 | (p : R ○ L ≅ T) 41 | (η : ∀ {X} → Hom C X (OMap T X)) → 42 | (right : ∀ {X Y} → Hom C X (OMap R Y) → Hom D (OMap L X) Y) → 43 | (left : ∀ {X Y} → Hom D (OMap L X) Y → Hom C X (OMap R Y)) → 44 | (ηlaw : ∀ {X} → left (iden D {OMap L X}) ≅ η {X}) → 45 | ∀ {X}{Z}{f : Hom C Z (OMap R X)} → 46 | (nat : comp C (HMap R (right f)) (comp C (left (iden D)) (iden C)) 47 | ≅ 48 | left (comp D (right f) (comp D (iden D) (HMap L (iden C))))) → 49 | (lawb : left (right f) ≅ f) → 50 | f 51 | ≅ 52 | comp C (subst (λ Z₁ → Hom C Z₁ (OMap R X)) 53 | (fcong Z (cong OMap p)) 54 | (HMap R (right f))) 55 | η 56 | alaw1lem {D = D} .(R ○ L) L R refl η right left ηlaw {X}{Z}{f} nat lawb = 57 | trans (trans (trans (sym lawb) 58 | (cong left 59 | (trans (sym (idr D)) 60 | (cong (comp D (right f)) 61 | (trans (sym (fid L)) 62 | (sym (idl D))))))) 63 | (trans (sym nat) 64 | (cong (comp C (HMap R (right f))) 65 | (idr C)))) 66 | (cong (comp C (HMap R (right f))) ηlaw) 67 | 68 | alaw2lem : ∀{c d}{D : Cat {c}{d}} 69 | (T : Fun C C)(L : Fun C D)(R : Fun D C) 70 | (p : R ○ L ≅ T) → 71 | (right : ∀ {X Y} → Hom C X (OMap R Y) → Hom D (OMap L X) Y) → 72 | (bind : ∀ {X Y} → Hom C X (OMap T Y) → Hom C (OMap T X) (OMap T Y)) → 73 | (natright : {X₁ X' : Obj C} {Y Y' : Obj D} (f₁ : Hom C X' X₁) 74 | (g : Hom D Y Y') (h : Hom C X₁ (OMap R Y)) → 75 | right (comp C (HMap R g) (comp C h f₁)) 76 | ≅ 77 | comp D g (comp D (right h) (HMap L f₁))) → 78 | ∀ {X}{Z} {W} {k : Hom C Z (OMap T W)}{f : Hom C W (OMap R X)} → 79 | (bindlaw : HMap R (right (subst (Hom C Z) (fcong W (cong OMap (sym p))) k)) 80 | ≅ bind k) → 81 | subst (λ Z → Hom C Z (OMap R X)) 82 | (fcong Z (cong OMap p)) 83 | (HMap R 84 | (right 85 | (comp C 86 | (subst (λ Z → Hom C Z (OMap R X)) 87 | (fcong W (cong OMap p)) (HMap R (right f))) 88 | k))) 89 | ≅ 90 | comp C 91 | (subst (λ Z → Hom C Z (OMap R X)) 92 | (fcong W (cong OMap p)) (HMap R (right f))) 93 | (bind k) 94 | alaw2lem {D = D} .(R ○ L) L R refl right bind natright {X}{Z}{W}{k}{f} bindlaw = 95 | trans (trans (cong (HMap R) 96 | (trans (cong (λ k₁ → right (comp C (HMap R (right f)) k₁)) 97 | (sym (idr C))) 98 | (trans (natright (iden C) (right f) k) 99 | (trans (cong (λ h → comp D 100 | (right f) 101 | (comp D 102 | (right k) 103 | h)) 104 | (fid L)) 105 | (trans (sym (ass D)) 106 | (idr D)))))) 107 | (fcomp R)) 108 | (cong (comp C (HMap R (right f))) bindlaw) 109 | 110 | 111 | ahomlem : ∀{c d}{D : Cat {c}{d}} 112 | (T : Fun C C)(L : Fun C D)(R : Fun D C)(p : R ○ L ≅ T) → 113 | (right : ∀ {X Y} → Hom C X (OMap R Y) → Hom D (OMap L X) Y) → 114 | (natright : {X₁ X' : Obj C} {Y Y' : Obj D} (f₁ : Hom C X' X₁) 115 | (g : Hom D Y Y') (h : Hom C X₁ (OMap R Y)) → 116 | right (comp C (HMap R g) (comp C h f₁)) 117 | ≅ 118 | comp D g (comp D (right h) (HMap L f₁))) → 119 | {X : Obj D}{Y : Obj D}{f : Hom D X Y} → 120 | {Z : Obj C} {f₁ : Hom C Z (OMap R X)} → 121 | comp C (HMap R f) 122 | (subst (λ Z₁ → Hom C Z₁ (OMap R X)) 123 | (fcong Z (cong OMap p)) 124 | (HMap R (right f₁))) 125 | ≅ 126 | subst (λ Z₁ → Hom C Z₁ (OMap R Y)) 127 | (fcong Z (cong OMap p)) 128 | (HMap R (right (comp C (HMap R f) f₁))) 129 | ahomlem {D = D} .(R ○ L) L R refl right natright {X}{Y}{f}{Z}{g} = 130 | trans (sym (fcomp R)) 131 | (cong (HMap R) 132 | (sym 133 | (trans (cong (λ g₁ → right (comp C (HMap R f) g₁)) (sym (idr C))) 134 | (trans (natright (iden C) f g) 135 | (trans (cong (λ h → comp D f (comp D (right g) h)) (fid L)) 136 | (trans (sym (ass D)) (idr D))))))) 137 | 138 | 139 | Llawlem : ∀{c d}{D : Cat {c}{d}} 140 | (T : Fun C C)(L : Fun C D)(R : Fun D C)(p : R ○ L ≅ T) → 141 | (right : ∀ {X Y} → Hom C X (OMap R Y) → Hom D (OMap L X) Y) → 142 | (bind : ∀ {X Y} → Hom C X (OMap T Y) → Hom C (OMap T X) (OMap T Y)) → 143 | (bindlaw : {X Y : Obj C} {f : Hom C X (OMap T Y)} → 144 | HMap R 145 | (right 146 | (subst (Hom C X) (fcong Y (cong OMap (sym p))) f)) 147 | ≅ bind f) → 148 | ∀{X Z} → 149 | {f : Hom C Z (OMap R (OMap L X))} 150 | {f' : Hom C Z (OMap T X)} → (q : f ≅ f') → 151 | subst (λ Z₁ → Hom C Z₁ (OMap R (OMap L X))) 152 | (fcong Z (cong OMap p)) (HMap R (right f)) 153 | ≅ bind f' 154 | Llawlem .(R ○ L) L R refl right bind bindlaw {X}{Z}{f}{.f} refl = bindlaw 155 | 156 | -------------------------------------------------------------------------------- /WellScopedTermsModel.agda: -------------------------------------------------------------------------------- 1 | module WellScopedTermsModel where 2 | 3 | open import Library 4 | open import WellScopedTerms 5 | open import RMonads 6 | open import RMonads.REM 7 | open import Categories.Sets 8 | 9 | _<<_ : ∀{n}{X : Set} → (Fin n → X) → X → Fin (suc n) → X 10 | (f << x) zero = x 11 | (f << x) (suc i) = f i 12 | 13 | record LambdaModel : Set₁ where 14 | field S : Set 15 | Env = λ n → Fin n → S 16 | field eval : ∀{n} → Env n → Tm n → S 17 | ap : S → S → S 18 | lawvar : ∀{n}{i : Fin n}{γ : Env n} → 19 | eval γ (var i) ≅ γ i 20 | lawapp : ∀{n}{t u : Tm n}{γ : Env n} → 21 | eval γ (app t u) ≅ ap (eval γ t) (eval γ u) 22 | lawlam : ∀{n}{t : Tm (suc n)}{γ : Env n}{s : S} → 23 | ap (eval γ (lam t)) s ≅ eval (γ << s) t 24 | lawext : ∀{f g : S} → ((a : S) → ap f a ≅ ap g a) → f ≅ g 25 | 26 | module Model (l : LambdaModel) where 27 | open LambdaModel l 28 | 29 | wk<< : ∀{m n}(α : Ren m n)(β : Env n) 30 | (v : S) → (y : Fin (suc m)) → 31 | ((β ∘ α) << v) y ≅ (β << v) (wk α y) 32 | wk<< α β v zero = refl 33 | wk<< α β v (suc i) = refl 34 | 35 | reneval : ∀{m n}(α : Ren m n)(β : Env n) 36 | (t : Tm m) → eval (β ∘ α) t ≅ (eval β ∘ ren α) t 37 | reneval α β (var i) = 38 | proof 39 | eval (β ∘ α) (var i) 40 | ≅⟨ lawvar ⟩ 41 | β (α i) 42 | ≅⟨ sym lawvar ⟩ 43 | eval β (var (α i)) ∎ 44 | reneval α β (lam t) = lawext λ a → 45 | proof 46 | ap (eval (β ∘ α) (lam t)) a 47 | ≅⟨ lawlam ⟩ 48 | eval ((β ∘ α) << a) t 49 | ≅⟨ cong (λ (f : Env _) → eval f t) (ext (wk<< α β a)) ⟩ 50 | eval ((β << a) ∘ wk α) t 51 | ≅⟨ reneval (wk α) (β << a) t ⟩ 52 | eval (β << a) (ren (wk α) t) 53 | ≅⟨ sym lawlam ⟩ 54 | ap (eval β (lam (ren (wk α) t))) a 55 | ∎ 56 | reneval α β (app t u) = 57 | proof 58 | eval (β ∘ α) (app t u) 59 | ≅⟨ lawapp ⟩ 60 | ap (eval (β ∘ α) t) (eval (β ∘ α) u) 61 | ≅⟨ cong₂ ap (reneval α β t) (reneval α β u) ⟩ 62 | ap (eval β (ren α t)) (eval β (ren α u)) 63 | ≅⟨ sym lawapp ⟩ 64 | eval β (app (ren α t) (ren α u)) 65 | ∎ 66 | 67 | lift<< : ∀{m n}(γ : Sub m n)(α : Env n) 68 | (a : S)(i : Fin (suc m)) → 69 | ((eval α ∘ γ ) << a) i ≅ (eval (α << a) ∘ lift γ) i 70 | lift<< γ α a zero = 71 | proof 72 | a 73 | ≅⟨ sym lawvar ⟩ 74 | eval (α << a) (var zero) 75 | ∎ 76 | lift<< γ α a (suc i) = 77 | proof 78 | eval α (γ i) 79 | ≡⟨⟩ 80 | eval ((α << a) ∘ suc) (γ i) 81 | ≅⟨ reneval suc (α << a) (γ i) ⟩ 82 | eval (α << a) (ren suc (γ i)) 83 | ∎ 84 | 85 | subeval : ∀{m n}(t : Tm m)(γ : Sub m n)(α : Env n) → 86 | eval (eval α ∘ γ) t ≅ (eval α ∘ sub γ) t 87 | subeval (var i) γ α = 88 | proof 89 | eval (eval α ∘ γ) (var i) 90 | ≅⟨ lawvar ⟩ 91 | eval α (γ i) 92 | ∎ 93 | subeval (lam t) γ α = lawext λ a → 94 | proof 95 | ap (eval (eval α ∘ γ) (lam t)) a 96 | ≅⟨ lawlam ⟩ 97 | eval ((eval α ∘ γ) << a) t 98 | ≅⟨ cong (λ (f : Env _) → eval f t) (ext (lift<< γ α a)) ⟩ 99 | eval (eval (α << a) ∘ lift γ) t 100 | ≅⟨ subeval t (lift γ) (α << a) ⟩ 101 | eval (α << a) (sub (lift γ) t) 102 | ≅⟨ sym lawlam ⟩ 103 | ap (eval α (lam (sub (lift γ) t))) a 104 | ∎ 105 | subeval (app t u) γ α = 106 | proof 107 | eval (eval α ∘ γ) (app t u) 108 | ≅⟨ lawapp ⟩ 109 | ap (eval (eval α ∘ γ) t) (eval (eval α ∘ γ) u) 110 | ≅⟨ cong₂ ap (subeval t γ α) (subeval u γ α) ⟩ 111 | ap (eval α (sub γ t)) (eval α (sub γ u)) 112 | ≅⟨ sym lawapp ⟩ 113 | eval α (app (sub γ t) (sub γ u)) 114 | ∎ 115 | 116 | TmRAlg : RAlg TmRMonad 117 | TmRAlg = record{ 118 | acar = S; 119 | astr = eval; 120 | alaw1 = ext λ _ → sym lawvar; 121 | alaw2 = ext λ t → subeval t _ _} 122 | 123 | 124 | module VEnv where 125 | open import Delay 126 | open import Size 127 | 128 | mutual 129 | Env : ℕ → Set 130 | Env n = Fin n → Val 131 | 132 | data Val : Set where 133 | clo : ∀{n} → Env n → Tm (suc n) → Val 134 | 135 | 136 | -- the RAlg is expecting a env containing 'values' here the values 137 | -- evaluator takes an env of undelayed values and makes a delayed 138 | -- values. 139 | mutual 140 | ev : ∀{i n} → Env n → Tm n → Delay Val i 141 | ev γ (var x) = now (γ x) 142 | ev γ (lam t) = now (clo γ t) 143 | ev γ (app t u) = ev γ t >>= λ f → ev γ u >>= λ v → f $$ v 144 | 145 | _∞$$_ : ∀{i} → Val → Val → ∞Delay Val i 146 | force (clo γ t ∞$$ v) = ev (γ << v) t 147 | 148 | _$$_ : ∀{i} → Val → Val → Delay Val i 149 | f $$ v = later (f ∞$$ v) 150 | 151 | module FusedVals where 152 | open import Size 153 | mutual 154 | data Env (i : Size) : (n : ℕ) → Set where 155 | ε : Env i zero 156 | _,_ : ∀ {n} (ρ : Env i n) (v : Val i) → Env i (suc n) 157 | 158 | data Val (i : Size) : Set where 159 | lam : forall {n} (t : Tm (suc n)) (ρ : Env i n) → Val i 160 | later : (v∞ : ∞Val i) → Val i 161 | 162 | record ∞Val (i : Size) : Set where 163 | coinductive 164 | constructor ∞val 165 | field 166 | vforce : {j : Size< i} → Val j 167 | 168 | open ∞Val 169 | 170 | lookup : ∀{i n} -> Env i n -> Fin n -> Val i 171 | lookup (ρ , v) zero = v 172 | lookup (ρ , v) (suc i) = lookup ρ i 173 | 174 | tabulate : ∀{i n} -> (Fin n -> Val i) -> Env i n 175 | tabulate {n = zero} f = ε 176 | tabulate {n = suc n} f = (tabulate {n = n} (f ∘ suc)) , f zero 177 | 178 | mutual 179 | ev : ∀{i n} -> Env i n -> Tm n -> Val i 180 | ev ρ (var x) = lookup ρ x 181 | ev ρ (lam t) = lam t ρ 182 | ev ρ (app t u) = ev ρ t $$ ev ρ u 183 | 184 | _$$_ : ∀{i} → Val i → Val i → Val i 185 | f $$ v = later (f ∞$$ v) 186 | 187 | _∞$$_ : ∀{i} → Val i → Val i → ∞Val i 188 | vforce (lam t ρ ∞$$ v) = ev (ρ , v) t 189 | vforce (later p ∞$$ v) = later (vforce p ∞$$ v) 190 | 191 | {- 192 | FRAlg : ∀ {i} -> RAlg TmRMonad 193 | FRAlg {i} = ralg 194 | (Val i) 195 | (ev ∘ tabulate) 196 | {!!} 197 | {!!} 198 | -} 199 | -------------------------------------------------------------------------------- /WellTypedTermsNBEModel.agda: -------------------------------------------------------------------------------- 1 | module WellTypedTermsNBEModel where 2 | 3 | open import Library 4 | open import Naturals 5 | open import WellTypedTerms 6 | open import RMonads.REM 7 | open import FunctorCat 8 | open import Categories.Sets 9 | open NatT 10 | open Σ 11 | 12 | -- normal forms 13 | 14 | mutual 15 | data Nf (Γ : Con) : Ty → Set where 16 | lam : ∀{σ τ} → Nf (Γ < σ) τ → Nf Γ (σ ⇒ τ) 17 | ne : Ne Γ ι → Nf Γ ι 18 | 19 | data Ne (Γ : Con) : Ty → Set where 20 | var : ∀{σ} → Var Γ σ → Ne Γ σ 21 | app : ∀{σ τ} → Ne Γ (σ ⇒ τ) → Nf Γ σ → Ne Γ τ 22 | 23 | mutual 24 | renNf : ∀{Γ Δ} → Ren Δ Γ → ∀{σ} → Nf Δ σ → Nf Γ σ 25 | renNf ρ (lam n) = lam (renNf (wk ρ) n) 26 | renNf ρ (ne n) = ne (renNe ρ n) 27 | 28 | renNe : ∀{Γ Δ} → Ren Δ Γ → ∀{σ} → Ne Δ σ → Ne Γ σ 29 | renNe ρ (var x) = var (ρ x) 30 | renNe ρ (app n n') = app (renNe ρ n) (renNf ρ n') 31 | 32 | mutual 33 | renNfid : ∀{Γ} σ (v : Nf Γ σ) → renNf id v ≅ v 34 | renNfid ι (ne x) = cong ne (renNeid ι x) 35 | renNfid (σ ⇒ τ) (lam v) = cong lam $ 36 | proof 37 | renNf (wk id) v 38 | ≅⟨ cong (λ (ρ : Ren _ _) → renNf ρ v) 39 | (iext λ _ → ext wkid) ⟩ 40 | renNf id v 41 | ≅⟨ renNfid τ v ⟩ 42 | v 43 | ∎ 44 | 45 | renNeid : ∀{Γ} σ (v : Ne Γ σ) → renNe id v ≅ v 46 | renNeid σ (var x) = refl 47 | renNeid τ (app {σ} v x) = cong₂ app (renNeid (σ ⇒ τ) v) (renNfid σ x) 48 | 49 | mutual 50 | renNecomp : ∀{Δ Γ B}(ρ : Ren Δ Γ)(ρ' : Ren Γ B) σ (v : Ne Δ σ) → 51 | renNe (ρ' ∘ ρ) v ≅ renNe ρ' (renNe ρ v) 52 | renNecomp ρ ρ' σ (var x) = refl 53 | renNecomp ρ ρ' τ (app {σ} v x) = cong₂ 54 | Ne.app 55 | (renNecomp ρ ρ' (σ ⇒ τ) v) 56 | (renNfcomp ρ ρ' σ x) 57 | 58 | renNfcomp : ∀{Δ Γ B}(ρ : Ren Δ Γ)(ρ' : Ren Γ B) σ (v : Nf Δ σ) → 59 | renNf (ρ' ∘ ρ) v ≅ renNf ρ' (renNf ρ v) 60 | renNfcomp ρ ρ' (σ ⇒ τ) (lam v) = cong Nf.lam $ 61 | proof 62 | renNf (wk (ρ' ∘ ρ)) v 63 | ≅⟨ cong (λ (ρ : Ren _ _) → renNf ρ v) 64 | (iext λ _ → ext (wkcomp ρ' ρ)) ⟩ 65 | renNf (wk ρ' ∘ wk ρ) v 66 | ≅⟨ renNfcomp (wk ρ) (wk ρ') τ v ⟩ 67 | renNf (wk ρ') (renNf (wk ρ) v) 68 | ∎ 69 | renNfcomp ρ ρ' ι (ne x) = cong ne (renNecomp ρ ρ' ι x) 70 | 71 | mutual 72 | Val : Con → Ty → Set 73 | Val Γ ι = Ne Γ ι 74 | Val Γ (σ ⇒ τ) = ∀{B} → Ren Γ B → Val B σ → Val B τ 75 | 76 | -- (λ f → ∀{B B'}(ρ : Ren Γ B)(ρ' : Ren B B')(a : Val B σ) → renV ρ' (f ρ a) ≅ f (ρ' ∘ ρ) (renV ρ' a)) 77 | 78 | renV : ∀{Γ Δ} → Ren Δ Γ → ∀{σ} → Val Δ σ → Val Γ σ 79 | renV ρ {ι} n = renNe ρ n 80 | renV {Γ}{Δ} ρ {σ ⇒ τ} f = λ ρ' → f (ρ' ∘ ρ) 81 | 82 | 83 | {- 84 | Ren Γ B → Val B σ → Val B τ 85 | | | 86 | \/ \/ 87 | Ren Γ B' → Val B' σ → Val B' τ 88 | 89 | 90 | -} 91 | -- interpretation of contexts 92 | 93 | Env : Con → Con → Set 94 | Env Γ Δ = ∀{σ} → Var Γ σ → Val Δ σ 95 | 96 | _<<_ : ∀{Γ Δ σ} → Env Γ Δ → Val Δ σ → Env (Γ < σ) Δ 97 | (γ << v) vz = v 98 | (γ << v) (vs x) = γ x 99 | 100 | eval : ∀{Γ Δ σ} → Env Δ Γ → Tm Δ σ → Val Γ σ 101 | eval γ (var i) = γ i 102 | eval γ (lam t) = λ ρ v → eval ((renV ρ ∘ γ) << v) t 103 | eval γ (app t u) = eval γ t id (eval γ u) 104 | 105 | 106 | 107 | lem : ∀{B Γ Δ σ}(ρ : Ren Γ B)(γ : Env Δ Γ)(t : Tm Δ σ) → 108 | renV ρ (eval γ t) ≅ eval (renV ρ ∘ γ) t 109 | lem = {!!} 110 | 111 | 112 | renVid : ∀{Γ} σ (v : Val Γ σ) → renV id v ≅ v 113 | renVid ι v = renNeid ι v 114 | renVid (σ ⇒ τ) v = refl 115 | 116 | renVcomp : ∀{Δ Γ B}(ρ : Ren Δ Γ)(ρ' : Ren Γ B) σ (v : Val Δ σ) → 117 | renV (ρ' ∘ ρ) v ≅ renV ρ' (renV ρ v) 118 | renVcomp ρ ρ' ι v = renNecomp ρ ρ' ι v 119 | renVcomp ρ ρ' (σ ⇒ τ) v = refl 120 | 121 | 122 | renV<< : ∀{B' B Γ}(α : Ren B B')(β : Env Γ B){σ}(v : Val B σ) → 123 | ∀{ρ}(y : Var (Γ < σ) ρ) → 124 | ((renV α ∘ β) << renV α v) y ≅ (renV α ∘ (β << v)) y 125 | renV<< α β v vz = refl 126 | renV<< α β v (vs y) = refl 127 | 128 | substeval : ∀{σ τ}(p : σ ≅ τ){Γ B : Con}{γ : Env Γ B}(t : Tm Γ σ) → 129 | (subst (Val B) p ∘ eval γ) t ≅ (eval γ ∘ subst (Tm Γ) p) t 130 | substeval refl t = refl 131 | 132 | wk<< : ∀{B Γ Δ}(α : Ren Γ Δ)(β : Env Δ B){σ}(v : Val B σ) → 133 | ∀{ρ}(y : Var (Γ < σ) ρ) → 134 | ((β ∘ α) << v) y ≅ (β << v) (wk α y) 135 | wk<< α β v vz = refl 136 | wk<< α β v (vs y) = refl 137 | 138 | <>=_ : ∀ {i A B} → Delay A i → (A → Delay B i) → Delay B i 40 | now x >>= f = f x 41 | later x >>= f = later (x ∞>>= f) 42 | 43 | _∞>>=_ : ∀ {i A B} → ∞Delay A i → (A → Delay B i) → ∞Delay B i 44 | force (x ∞>>= f) = force x >>= f 45 | 46 | delayMonad : ∀ {i} → RawMonad {f = lzero} (λ A → Delay A i) 47 | delayMonad {i} = record 48 | { return = now 49 | ; _>>=_ = _>>=_ {i} 50 | } where open Bind 51 | 52 | module _ {i : Size} where 53 | open module DelayMonad = RawMonad (delayMonad {i = i}) 54 | public renaming (_⊛_ to _<*>_) 55 | open Bind public using (_∞>>=_) 56 | 57 | -- Map for ∞Delay 58 | 59 | _∞<$>_ : ∀ {i A B} (f : A → B) (∞a : ∞Delay A i) → ∞Delay B i 60 | f ∞<$> ∞a = ∞a ∞>>= λ a → return (f a) 61 | -- force (f ∞<$> ∞a) = f <$> force ∞a 62 | 63 | -- Double bind 64 | 65 | _=<<2_,_ : ∀ {i A B C} → (A → B → Delay C i) → Delay A i → Delay B i → Delay C i 66 | f =<<2 x , y = x >>= λ a → y >>= λ b → f a b 67 | 68 | -- Strong bisimilarity 69 | 70 | mutual 71 | data _~_ {i : Size} {A : Set} : (a? b? : Delay A ∞) → Set where 72 | ~now : ∀ a → now a ~ now a 73 | ~later : ∀ {a∞ b∞} (eq : _∞~_ {i} a∞ b∞) → later a∞ ~ later b∞ 74 | 75 | record _∞~_ {i : Size} {A : Set} (a∞ b∞ : ∞Delay A ∞) : Set where 76 | coinductive 77 | constructor ~delay 78 | field 79 | ~force : {j : Size< i} → _~_ {j} (force a∞) (force b∞) 80 | 81 | open _∞~_ public 82 | 83 | -- Reflexivity 84 | 85 | mutual 86 | ~refl : ∀ {i A} (a? : Delay A ∞) → _~_ {i} a? a? 87 | ~refl (now a) = ~now a 88 | ~refl (later a∞) = ~later (∞~refl a∞) 89 | 90 | ∞~refl : ∀ {i A} (a∞ : ∞Delay A ∞) → _∞~_ {i} a∞ a∞ 91 | ~force (∞~refl a∞) = ~refl (force a∞) 92 | 93 | -- Symmetry 94 | 95 | mutual 96 | ~sym : ∀ {i A} {a? b? : Delay A ∞} → _~_ {i} a? b? → _~_ {i} b? a? 97 | ~sym (~now a) = ~now a 98 | ~sym (~later eq) = ~later (∞~sym eq) 99 | 100 | ∞~sym : ∀ {i A} {a? b? : ∞Delay A ∞} → _∞~_ {i} a? b? → _∞~_ {i} b? a? 101 | ~force (∞~sym eq) = ~sym (~force eq) 102 | 103 | -- Transitivity 104 | 105 | mutual 106 | ~trans : ∀ {i A} {a? b? c? : Delay A ∞} 107 | (eq : _~_ {i} a? b?) (eq' : _~_ {i} b? c?) → _~_ {i} a? c? 108 | ~trans (~now a) (~now .a) = ~now a 109 | ~trans (~later eq) (~later eq') = ~later (∞~trans eq eq') 110 | 111 | ∞~trans : ∀ {i A} {a∞ b∞ c∞ : ∞Delay A ∞} 112 | (eq : _∞~_ {i} a∞ b∞) (eq' : _∞~_ {i} b∞ c∞) → _∞~_ {i} a∞ c∞ 113 | ~force (∞~trans eq eq') = ~trans (~force eq) (~force eq') 114 | 115 | -- Equality reasoning 116 | 117 | ~setoid : (i : Size) (A : Set) → Setoid lzero lzero 118 | ~setoid i A = record 119 | { Carrier = Delay A ∞ 120 | ; _≈_ = _~_ {i} 121 | ; isEquivalence = record 122 | { refl = λ {a?} → ~refl a? 123 | ; sym = ~sym 124 | ; trans = ~trans 125 | } 126 | } 127 | 128 | module ~-Reasoning {i : Size} {A : Set} where 129 | open Pre (Setoid.preorder (~setoid i A)) public 130 | -- using (begin_; _∎) (_≈⟨⟩_ to _~⟨⟩_; _≈⟨_⟩_ to _~⟨_⟩_) 131 | renaming (_≈⟨⟩_ to _≡⟨⟩_; _≈⟨_⟩_ to _≡⟨_⟩_; _∼⟨_⟩_ to _~⟨_⟩_; begin_ to proof_) 132 | 133 | ∞~setoid : (i : Size) (A : Set) → Setoid lzero lzero 134 | ∞~setoid i A = record 135 | { Carrier = ∞Delay A ∞ 136 | ; _≈_ = _∞~_ {i} 137 | ; isEquivalence = record 138 | { refl = λ {a?} → ∞~refl a? 139 | ; sym = ∞~sym 140 | ; trans = ∞~trans 141 | } 142 | } 143 | 144 | module ∞~-Reasoning {i : Size} {A : Set} where 145 | open Pre (Setoid.preorder (∞~setoid i A)) public 146 | -- using (begin_; _∎) (_≈⟨⟩_ to _~⟨⟩_; _≈⟨_⟩_ to _~⟨_⟩_) 147 | renaming (_≈⟨⟩_ to _≡⟨⟩_; _≈⟨_⟩_ to _≡⟨_⟩_; _∼⟨_⟩_ to _∞~⟨_⟩_; begin_ to proof_) 148 | 149 | 150 | -- Congruence laws. 151 | 152 | mutual 153 | bind-cong-l : ∀ {i A B} {a? b? : Delay A ∞} (eq : _~_ {i} a? b?) 154 | (k : A → Delay B ∞) → _~_ {i} (a? >>= k) (b? >>= k) 155 | bind-cong-l (~now a) k = ~refl _ 156 | bind-cong-l (~later eq) k = ~later (∞bind-cong-l eq k) 157 | 158 | ∞bind-cong-l : ∀ {i A B} {a∞ b∞ : ∞Delay A ∞} (eq : _∞~_ {i} a∞ b∞) → 159 | (k : A → Delay B ∞) → 160 | _∞~_ {i} (a∞ ∞>>= k) (b∞ ∞>>= k) 161 | ~force (∞bind-cong-l eq k) = bind-cong-l (~force eq) k 162 | 163 | _>>=l_ = bind-cong-l 164 | 165 | mutual 166 | bind-cong-r : ∀ {i A B} (a? : Delay A ∞) {k l : A → Delay B ∞} → 167 | (h : ∀ a → _~_ {i} (k a) (l a)) → _~_ {i} (a? >>= k) (a? >>= l) 168 | bind-cong-r (now a) h = h a 169 | bind-cong-r (later a∞) h = ~later (∞bind-cong-r a∞ h) 170 | 171 | ∞bind-cong-r : ∀ {i A B} (a∞ : ∞Delay A ∞) {k l : A → Delay B ∞} → 172 | (h : ∀ a → _~_ {i} (k a) (l a)) → _∞~_ {i} (a∞ ∞>>= k) (a∞ ∞>>= l) 173 | ~force (∞bind-cong-r a∞ h) = bind-cong-r (force a∞) h 174 | 175 | _>>=r_ = bind-cong-r 176 | 177 | mutual 178 | bind-cong : ∀ {i A B} {a? b? : Delay A ∞} (eq : _~_ {i} a? b?) 179 | {k l : A → Delay B ∞} (h : ∀ a → _~_ {i} (k a) (l a)) → 180 | _~_ {i} (a? >>= k) (b? >>= l) 181 | bind-cong (~now a) h = h a 182 | bind-cong (~later eq) h = ~later (∞bind-cong eq h) 183 | 184 | ∞bind-cong : ∀ {i A B} {a∞ b∞ : ∞Delay A ∞} (eq : _∞~_ {i} a∞ b∞) 185 | {k l : A → Delay B ∞} (h : ∀ a → _~_ {i} (k a) (l a)) → 186 | _∞~_ {i} (a∞ ∞>>= k) (b∞ ∞>>= l) 187 | ~force (∞bind-cong eq h) = bind-cong (~force eq) h 188 | 189 | _~>>=_ = bind-cong 190 | 191 | -- Monad laws. 192 | 193 | mutual 194 | bind-assoc : ∀{i A B C}(m : Delay A ∞) 195 | {k : A → Delay B ∞}{l : B → Delay C ∞} → 196 | _~_ {i} ((m >>= k) >>= l) (m >>= λ a → k a >>= l) 197 | bind-assoc (now a) = ~refl _ 198 | bind-assoc (later a∞) = ~later (∞bind-assoc a∞) 199 | 200 | ∞bind-assoc : ∀{i A B C}(a∞ : ∞Delay A ∞) 201 | {k : A → Delay B ∞}{l : B → Delay C ∞} → 202 | _∞~_ {i} ((a∞ ∞>>= λ a → k a) ∞>>= l) (a∞ ∞>>= λ a → k a >>= l) 203 | ~force (∞bind-assoc a∞) = bind-assoc (force a∞) 204 | 205 | -- Termination/Convergence. Makes sense only for Delay A ∞. 206 | 207 | data _⇓_ {A : Set} : (a? : Delay A ∞) (a : A) → Set where 208 | now⇓ : ∀ {a} → now a ⇓ a 209 | later⇓ : ∀ {a} {a∞ : ∞Delay A ∞} → force a∞ ⇓ a → later a∞ ⇓ a 210 | 211 | _⇓ : {A : Set} (x : Delay A ∞) → Set 212 | x ⇓ = ∃ λ a → x ⇓ a 213 | 214 | -- Monotonicity. 215 | 216 | map⇓ : ∀ {A B} {a : A} {a? : Delay A ∞} 217 | (f : A → B) (a⇓ : a? ⇓ a) → (f <$> a?) ⇓ f a 218 | map⇓ f now⇓ = now⇓ 219 | map⇓ f (later⇓ a⇓) = later⇓ (map⇓ f a⇓) 220 | 221 | -- some lemmas about convergence 222 | subst~⇓ : ∀{A}{t t' : Delay A ∞}{n : A} → t ⇓ n → t ~ t' → t' ⇓ n 223 | subst~⇓ now⇓ (~now a) = now⇓ 224 | subst~⇓ (later⇓ p) (~later eq) = later⇓ (subst~⇓ p (~force eq)) 225 | 226 | -- this should also hold for weak bisimularity right? 227 | {- 228 | subst≈⇓ : ∀{A}{t t' : Delay A ∞}{n : A} → t ⇓ n → t ≈ t' → t' ⇓ n 229 | subst≈⇓ = ? 230 | -} 231 | 232 | 233 | ⇓>>= : ∀{A B}(f : A → Delay B ∞) 234 | {?a : Delay A ∞}{a : A} → ?a ⇓ a → 235 | {b : B} → (?a >>= f) ⇓ b → f a ⇓ b 236 | ⇓>>= f now⇓ q = q 237 | ⇓>>= f (later⇓ p) (later⇓ q) = ⇓>>= f p q 238 | 239 | >>=⇓ : ∀{A B}(f : A → Delay B ∞) 240 | {?a : Delay A ∞}{a : A} → ?a ⇓ a → 241 | {b : B} → f a ⇓ b → (?a >>= f) ⇓ b 242 | >>=⇓ f now⇓ q = q 243 | >>=⇓ f (later⇓ p) q = later⇓ (>>=⇓ f p q) 244 | 245 | -- handy when you can't pattern match like in a let definition 246 | unlater : ∀{A}{∞a : ∞Delay A ∞}{a : A} → later ∞a ⇓ a → force ∞a ⇓ a 247 | unlater (later⇓ p) = p 248 | 249 | 250 | 251 | 252 | 253 | 254 | 255 | 256 | {- 257 | mutual 258 | _⇓_ : {A : Set} {i : Size} (x : Delay A i) (a : A) → Set 259 | x ⇓ a = Terminates a x 260 | 261 | data Terminates {A : Set} {i : Size} (a : A) : Delay A i → Set where 262 | now : now a ⇓ a 263 | later : ∀ {x : ∞Delay A i} → (force x {j = ?}) ⇓ a → later x ⇓ a 264 | 265 | mutual 266 | 267 | cast : ∀ {A i} → Delay A i → (j : Size< ↑ i) → Delay A j 268 | cast (now x) j = now x 269 | cast (later x) j = {!later (∞cast x j)!} 270 | 271 | ∞cast : ∀ {A i} → ∞Delay A i → (j : Size< ↑ i) → ∞Delay A j 272 | ♭ (∞cast x j) {k} = cast {i = j} (♭ x) k 273 | -} 274 | -------------------------------------------------------------------------------- /RMonads/CatofRAdj/TermRAdjHom.agda: -------------------------------------------------------------------------------- 1 | open import Categories 2 | open import Functors 3 | open import RMonads 4 | 5 | module RMonads.CatofRAdj.TermRAdjHom {a b c d}{C : Cat {a}{b}}{D : Cat {c}{d}} 6 | {J : Fun C D}(M : RMonad J) where 7 | 8 | open import Library 9 | open import RAdjunctions 10 | open import RMonads.CatofRAdj M 11 | open import Categories.Terminal 12 | open import RMonads.REM M 13 | open import RMonads.REM.Adjunction M 14 | open import RAdjunctions.RAdj2RMon 15 | open import RMonads.CatofRAdj.TermRAdjObj M 16 | 17 | open Cat 18 | open Fun 19 | open RAdj 20 | 21 | alaw1lem : ∀{e f}{E : Cat {e}{f}}(T : Fun C D)(L : Fun C E)(R : Fun E D) 22 | (p : R ○ L ≅ T) 23 | (η : ∀ {X} → Hom D (OMap J X) (OMap T X)) → 24 | (right : ∀ {X Y} → Hom D (OMap J X) (OMap R Y) → Hom E (OMap L X) Y) → 25 | (left : ∀ {X Y} → Hom E (OMap L X) Y → Hom D (OMap J X) (OMap R Y)) → 26 | (ηlaw : ∀ {X} → left (iden E {OMap L X}) ≅ η {X}) → 27 | ∀ {X}{Z}{f : Hom D (OMap J Z) (OMap R X)} → 28 | (nat : comp D (HMap R (right f)) (comp D (left (iden E)) (HMap J (iden C))) 29 | ≅ 30 | left (comp E (right f) (comp E (iden E) (HMap L (iden C))))) → 31 | (lawb : left (right f) ≅ f) → 32 | f 33 | ≅ 34 | comp D (subst (λ Z → Hom D Z (OMap R X)) 35 | (fcong Z (cong OMap p)) 36 | (HMap R (right f))) 37 | η 38 | alaw1lem {E = E} .(R ○ L) L R refl η right left ηlaw {X}{Z}{f} nat lawb = 39 | trans (trans (trans (sym lawb) 40 | (cong left 41 | (trans (sym (idr E)) 42 | (cong (comp E (right f)) 43 | (trans (sym (fid L)) 44 | (sym (idl E))))))) 45 | (trans (sym nat) 46 | (cong (comp D (HMap R (right f))) 47 | (trans (cong (comp D (left (iden E))) (fid J)) 48 | (idr D))))) 49 | (cong (comp D (HMap R (right f))) ηlaw) 50 | 51 | alaw2lem : ∀{e f}{E : Cat {e}{f}}(T : Fun C D)(L : Fun C E)(R : Fun E D) 52 | (p : R ○ L ≅ T) → 53 | (right : ∀ {X Y} → Hom D (OMap J X) (OMap R Y) → Hom E (OMap L X) Y) → 54 | (bind : ∀{X Y} → 55 | Hom D (OMap J X) (OMap T Y) → Hom D (OMap T X) (OMap T Y)) → 56 | (natright : {X X' : Obj C} {Y Y' : Obj E} (f : Hom C X' X) 57 | (g : Hom E Y Y') (h : Hom D (OMap J X) (OMap R Y)) → 58 | right (comp D (HMap R g) (comp D h (HMap J f))) 59 | ≅ 60 | comp E g (comp E (right h) (HMap L f))) → 61 | ∀{X}{Z}{W} 62 | {k : Hom D (OMap J Z) (OMap T W)} 63 | {f : Hom D (OMap J W) (OMap R X)} → 64 | (bindlaw : HMap R (right (subst (Hom D (OMap J Z)) 65 | (fcong W (cong OMap (sym p))) 66 | k)) 67 | ≅ 68 | bind k) → 69 | subst (λ Z → Hom D Z (OMap R X)) 70 | (fcong Z (cong OMap p)) 71 | (HMap R 72 | (right 73 | (comp D 74 | (subst (λ Z → Hom D Z (OMap R X)) 75 | (fcong W (cong OMap p)) (HMap R (right f))) 76 | k))) 77 | ≅ 78 | comp D 79 | (subst (λ Z → Hom D Z (OMap R X)) 80 | (fcong W (cong OMap p)) (HMap R (right f))) 81 | (bind k) 82 | alaw2lem {E = E} .(R ○ L) L R refl right bind natright {k = k}{f = f} bl = 83 | trans (trans (cong (HMap R) 84 | (trans (cong (λ k → right (comp D (HMap R (right f)) k)) 85 | (trans (sym (idr D)) (cong (comp D k) 86 | (sym (fid J))))) 87 | (trans (trans (natright (iden C) (right f) k) 88 | (trans (sym (ass E)) 89 | (cong (comp E 90 | (comp E 91 | (right f) 92 | (right k))) 93 | (fid L)))) 94 | (idr E)))) 95 | (fcomp R)) 96 | (cong (comp D (HMap R (right f))) bl) 97 | 98 | 99 | ahomlem : ∀{e f}{E : Cat {e}{f}}(T : Fun C D)(L : Fun C E)(R : Fun E D) 100 | (p : R ○ L ≅ T) → 101 | (right : ∀ {X Y} → Hom D (OMap J X) (OMap R Y) → Hom E (OMap L X) Y) → 102 | (natright : {X X' : Obj C} {Y Y' : Obj E} (f : Hom C X' X) 103 | (g : Hom E Y Y') (h : Hom D (OMap J X) (OMap R Y)) → 104 | right (comp D (HMap R g) (comp D h (HMap J f))) 105 | ≅ 106 | comp E g (comp E (right h) (HMap L f))) → 107 | {X : Obj E}{Y : Obj E}{f : Hom E X Y} → 108 | {Z : Obj C} {g : Hom D (OMap J Z) (OMap R X)} → 109 | comp D (HMap R f) 110 | (subst (λ Z → Hom D Z (OMap R X)) 111 | (fcong Z (cong OMap p)) 112 | (HMap R (right g))) 113 | ≅ 114 | subst (λ Z → Hom D Z (OMap R Y)) 115 | (fcong Z (cong OMap p)) 116 | (HMap R (right (comp D (HMap R f) g))) 117 | ahomlem {E = E} .(R ○ L) L R refl right natright {X}{Y}{f}{Z}{g} = 118 | trans (sym (fcomp R)) 119 | (cong (HMap R) 120 | (trans (trans (sym (idr E)) 121 | (trans (cong (comp E (comp E f (right g))) 122 | (sym (fid L))) 123 | (trans (ass E) 124 | (sym (natright (iden C) f g))))) 125 | (cong (λ g → right (comp D (HMap R f) g)) 126 | (trans (cong (comp D g) (fid J)) (idr D))))) 127 | open ObjAdj 128 | 129 | 130 | Llawlem : ∀{e f}{E : Cat {e}{f}}(T : Fun C D)(L : Fun C E)(R : Fun E D) 131 | (p : R ○ L ≅ T) → 132 | (right : ∀{X Y} → Hom D (OMap J X) (OMap R Y) → Hom E (OMap L X) Y) → 133 | (bind : ∀{X Y} → 134 | Hom D (OMap J X) (OMap T Y) → Hom D (OMap T X) (OMap T Y)) → 135 | (bindlaw : {X Y : Obj C} {f : Hom D (OMap J X) (OMap T Y)} → 136 | HMap R 137 | (right 138 | (subst (Hom D (OMap J X)) (fcong Y (cong OMap (sym p))) f)) 139 | ≅ bind f) → 140 | ∀{X Z} → 141 | {f : Hom D (OMap J Z) (OMap R (OMap L X))} 142 | {f' : Hom D (OMap J Z) (OMap T X)} → (q : f ≅ f') → 143 | subst (λ Z → Hom D Z (OMap R (OMap L X))) 144 | (fcong Z (cong OMap p)) (HMap R (right f)) 145 | ≅ bind f' 146 | Llawlem .(R ○ L) L R refl right bind bindlaw {X}{Z}{f}{.f} refl = bindlaw 147 | 148 | K : ∀{e f}(A : Obj (CatofAdj {e}{f})) → Fun (E A) (E EMObj) 149 | K A = record { 150 | OMap = λ X → record { 151 | acar = OMap (R (adj A)) X; 152 | astr = λ {Z} f → 153 | subst (λ Z₁ → Hom D Z₁ (OMap (R (adj A)) X)) 154 | (fcong Z (cong OMap (law A))) 155 | (HMap (R (adj A)) (right (adj A) f)); 156 | alaw1 = λ {Z} {f} → 157 | alaw1lem 158 | (TFun M) 159 | (L (adj A)) 160 | (R (adj A)) 161 | (law A) 162 | η 163 | (right (adj A)) 164 | (left (adj A)) 165 | (ηlaw A) 166 | (natleft (adj A) (iden C) (right (adj A) f) (iden (E A))) 167 | (lawb (adj A) f); 168 | alaw2 = λ {Z} {W} {k} {f} → 169 | alaw2lem 170 | (TFun M) 171 | (L (adj A)) 172 | (R (adj A)) 173 | (law A) 174 | (right (adj A)) 175 | bind 176 | (natright (adj A)) 177 | (bindlaw A {_} {_} {k})}; 178 | HMap = λ {X} {Y} f → record { 179 | amor = HMap (R (adj A)) f; 180 | ahom = λ {Z} {g} → 181 | ahomlem 182 | (TFun M) 183 | (L (adj A)) 184 | (R (adj A)) 185 | (law A) 186 | (right (adj A)) 187 | (natright (adj A)) }; 188 | fid = RAlgMorphEq (fid (R (adj A))); 189 | fcomp = RAlgMorphEq (fcomp (R (adj A)))} 190 | where open RMonad M 191 | 192 | 193 | Llaw' : ∀{e f}(A : Obj (CatofAdj {e}{f})) → 194 | K A ○ L (adj A) ≅ L (adj EMObj) 195 | Llaw' A = FunctorEq _ _ 196 | (ext λ X → AlgEq (fcong X (cong OMap (law A))) 197 | (iext λ Z → dext λ {f} {f'} → 198 | Llawlem (TFun M) 199 | (L (adj A)) 200 | (R (adj A)) 201 | (law A) 202 | (right (adj A)) 203 | bind 204 | (bindlaw A))) 205 | (iext λ X → iext λ Y → ext λ f → 206 | lemZ 207 | (AlgEq (fcong X (cong OMap (law A))) 208 | (iext λ Z → 209 | dext 210 | (Llawlem (TFun M) (L (adj A)) (R (adj A)) (law A) (right (adj A)) 211 | bind (bindlaw A)))) 212 | (AlgEq (fcong Y (cong OMap (law A))) 213 | (iext λ Z → 214 | dext 215 | (Llawlem (TFun M) (L (adj A)) (R (adj A)) (law A) (right (adj A)) 216 | bind (bindlaw A)))) 217 | (cong' refl 218 | (ext 219 | (λ _ → 220 | cong₂ (Hom D) (fcong X (cong OMap (law A))) 221 | (fcong Y (cong OMap (law A))))) 222 | (icong' refl 223 | (ext 224 | (λ Z → 225 | cong₂ (λ x y → Hom C X Z → Hom D x y) (fcong X (cong OMap (law A))) 226 | (fcong Z (cong OMap (law A))))) 227 | (icong' refl 228 | (ext 229 | (λ Z → 230 | cong 231 | (λ (F : Obj C → Obj D) → 232 | {Y₁ : Obj C} → Hom C Z Y₁ → Hom D (F Z) (F Y₁)) 233 | (cong OMap (law A)))) 234 | (cong HMap (law A)) (refl {x = X})) 235 | (refl {x = Y})) 236 | (refl {x = f}))) 237 | where open RMonad M 238 | 239 | Rlaw' : ∀{e f}(A : Obj (CatofAdj {e}{f})) → 240 | R (adj A) ≅ R (adj EMObj) ○ K A 241 | Rlaw' A = FunctorEq _ _ refl refl 242 | 243 | 244 | rightlaw' : ∀{e f}(A : Obj (CatofAdj {e}{f})) → 245 | {X : Obj C}{Y : Obj (E A)} → 246 | {f : Hom D (OMap J X) (OMap (R (adj A)) Y)} → 247 | HMap (K A) (right (adj A) f) 248 | ≅ 249 | right (adj EMObj) {X} {OMap (K A) Y} 250 | (subst (Hom D (OMap J X)) (fcong Y (cong OMap (Rlaw' A))) f) 251 | rightlaw' A {X}{Y}{f} = lemZ 252 | (AlgEq (fcong X (cong OMap (law A))) 253 | (iext λ Z → 254 | dext 255 | (λ {g} {g'} p → 256 | Llawlem (TFun M) (L (adj A)) (R (adj A)) (law A) (right (adj A)) 257 | bind (bindlaw A) p))) 258 | refl 259 | (trans 260 | (cong 261 | (λ (f₁ : Hom D (OMap J X) (OMap (R (adj A)) Y)) → 262 | HMap (R (adj A)) (right (adj A) f₁)) 263 | (sym 264 | (stripsubst (Hom D (OMap J X)) f 265 | (fcong Y (cong OMap (Rlaw' A)))))) 266 | (sym 267 | (stripsubst (λ (Z : Obj D) → Hom D Z (OMap (R (adj A)) Y)) _ 268 | (fcong X (cong OMap (law A)))))) 269 | where open RMonad M 270 | 271 | EMHom : {A : Obj CatofAdj} → Hom CatofAdj A EMObj 272 | EMHom {A} = record { 273 | K = K A; 274 | Llaw = Llaw' A; 275 | Rlaw = Rlaw' A; 276 | rightlaw = rightlaw' A } 277 | 278 | -- -} 279 | --------------------------------------------------------------------------------