bsc-leon-vatthauer/src/Monad/Instance/K.lagda.md

56 lines
1.8 KiB
Markdown
Raw Normal View History

2023-08-19 12:15:34 +02:00
<!--
```agda
open import Level
open import Categories.FreeObjects.Free
open import Categories.Functor.Core
open import Categories.Adjoint
open import Categories.Adjoint.Properties
open import Categories.Monad
open import Category.Instance.AmbientCategory using (Ambient)
2023-08-19 12:15:34 +02:00
```
-->
## Summary
In this file I explore the monad ***K*** and its properties:
- [X] *Lemma 16* Definition of the monad
- [ ] *Lemma 16* EilenbergMoore⇒UniformIterationAlgebras (use [crude monadicity theorem](https://agda.github.io/agda-categories/Categories.Adjoint.Monadic.Crude.html))
- [ ] *Proposition 19* ***K*** is strong
- [ ] *Theorem 22* ***K*** is an equational lifting monad
- [ ] *Proposition 23* The Kleisli category of ***K*** is enriched over pointed partial orders and strict monotone maps
- [ ] *Proposition 25* ***K*** is copyable and weakly discardable
- [ ] *Theorem 29* ***K*** is an initial pre-Elgot monad and an initial strong pre-Elgot monad
## Code
```agda
module Monad.Instance.K {o e} (ambient : Ambient o e) where
open Ambient ambient
open import Category.Construction.UniformIterationAlgebras ambient
2023-10-05 16:22:05 +02:00
open import Algebra.Properties ambient using (FreeUniformIterationAlgebra; uniformForgetfulF)
2023-08-19 12:15:34 +02:00
2023-10-05 16:22:05 +02:00
open Equiv
2023-08-19 12:15:34 +02:00
```
### *Lemma 16*: definition of monad ***K***
```agda
record MonadK : Set (suc o ⊔ suc ⊔ suc e) where
field
algebras : ∀ X → FreeUniformIterationAlgebra X
freeF : Functor C Uniform-Iteration-Algebras
2023-10-05 16:22:05 +02:00
freeF = FO⇒Functor uniformForgetfulF algebras
2023-08-19 12:15:34 +02:00
2023-10-05 16:22:05 +02:00
adjoint : freeF ⊣ uniformForgetfulF
adjoint = FO⇒LAdj uniformForgetfulF algebras
2023-08-19 12:15:34 +02:00
K : Monad C
K = adjoint⇒monad adjoint
-- EilenbergMoore⇒UniformIterationAlgebras : StrongEquivalence (EilenbergMoore K) (Uniform-Iteration-Algebras D)
-- EilenbergMoore⇒UniformIterationAlgebras = {! !}
```