🔨 refactor folder structure, define ambient category and use throughout

This commit is contained in:
Leon Vatthauer 2023-09-12 16:03:20 +02:00
parent 230b34da49
commit 481e1011e5
Signed by: leonv
SSH key fingerprint: SHA256:G4+ddwoZmhLPRB1agvXzZMXIzkVJ36dUYZXf5NxT+u8
12 changed files with 151 additions and 121 deletions

View file

@ -10,21 +10,25 @@ pandoc: public/*.md
agda: Everything.agda agda: Everything.agda
agda --html --html-dir=public Everything.agda --html-highlight=auto -i. agda --html --html-dir=public Everything.agda --html-highlight=auto -i.
mv public/Everything.html public/index.html
clean: clean:
rm Everything.agda rm -f Everything.agda
rm -rf public/* rm -rf public/*
find . -name '*.agdai' -exec rm \{\} \;
open: open:
firefox public/Everything.html firefox public/Everything.html
push: all # push compiled html to my cip directory
mv public/Everything.html public/index.html push: all push'
# just push without building
push':
chmod +w public/Agda.css chmod +w public/Agda.css
mv public bsc-thesis mv public bsc-thesis
scp -r bsc-thesis hy84coky@cip2a7.cip.cs.fau.de:.www/ scp -r bsc-thesis hy84coky@cip2a7.cip.cs.fau.de:.www/
mv bsc-thesis public mv bsc-thesis public
Everything.agda: Everything.agda:
git ls-tree --full-tree -r --name-only HEAD | egrep '^src/[^\.]*.l?agda(\.md)?' | sed -e 's|^src/[/]*|import |' -e 's|/|.|g' -e 's/.agda//' -e '/import Everything/d' -e 's/..md//' | LC_COLLATE='C' sort > Everything.agda git ls-tree --full-tree -r --name-only HEAD | egrep '^src/[^\.]*.l?agda(\.md)?' | sed -e 's|^src/[/]*|import |' -e 's|/|.|g' -e 's/.agda//' -e '/import Everything/d' -e 's/..md//' | LC_COLLATE='C' sort > Everything.agda

View file

@ -3,13 +3,7 @@
open import Level open import Level
open import Categories.Functor renaming (id to idF) open import Categories.Functor renaming (id to idF)
open import Categories.Functor.Algebra open import Categories.Functor.Algebra
open import Categories.Category open import Category.Instance.AmbientCategory using (Ambient)
open import Categories.Category.Cartesian
open import Categories.Category.BinaryProducts
open import Categories.Category.Cocartesian
open import Categories.Category.Extensive.Bundle
open import Categories.Category.Extensive
import Categories.Morphism.Reasoning as MR
``` ```
--> -->
@ -23,17 +17,8 @@ This file introduces (guarded) elgot algebras
## Code ## Code
```agda ```agda
module ElgotAlgebra where module Algebra.ElgotAlgebra {o e} (ambient : Ambient o e) where
open Ambient ambient
private
variable
o e : Level
module _ (D : ExtensiveDistributiveCategory o e) where
open ExtensiveDistributiveCategory D renaming (U to C; id to idC)
open Cocartesian (Extensive.cocartesian extensive)
open Cartesian (ExtensiveDistributiveCategory.cartesian D)
open MR C open MR C
``` ```

View file

@ -1,10 +1,7 @@
<!-- <!--
```agda ```agda
open import Level open import Level
open import Categories.Category.Core open import Category.Instance.AmbientCategory using (Ambient)
open import Categories.Category.Extensive.Bundle
open import Categories.Category.Extensive
open import Categories.Category.Cocartesian
``` ```
--> -->
## Summary ## Summary
@ -15,9 +12,8 @@ This file introduces *Uniform-Iteration Algebras*
## Code ## Code
```agda ```agda
module UniformIterationAlgebra {o e} (D : ExtensiveDistributiveCategory o e) where module Algebra.UniformIterationAlgebra {o e} (ambient : Ambient o e) where
open ExtensiveDistributiveCategory D renaming (U to C; id to idC) open Ambient ambient
open Cocartesian (Extensive.cocartesian extensive)
``` ```
### *Definition 12*: Uniform-Iteration Algebras ### *Definition 12*: Uniform-Iteration Algebras

View file

@ -11,12 +11,10 @@ open import Categories.Object.Product using (Product)
open import Categories.Object.Coproduct using (Coproduct) open import Categories.Object.Coproduct using (Coproduct)
open import Categories.Object.Exponential using (Exponential) open import Categories.Object.Exponential using (Exponential)
open import Categories.Category open import Categories.Category
open import ElgotAlgebra
open import Categories.Category.Distributive open import Categories.Category.Distributive
open import Categories.Category.Extensive.Bundle open import Categories.Category.Extensive.Bundle
open import Categories.Category.Extensive open import Categories.Category.Extensive
import Categories.Morphism as M open import Category.Instance.AmbientCategory using (Ambient)
import Categories.Morphism.Reasoning as MR
``` ```
--> -->
@ -30,17 +28,9 @@ This file introduces the category of *unguarded* elgot algebras
## Code ## Code
```agda ```agda
module ElgotAlgebras where module Category.Construction.ElgotAlgebras {o e} (ambient : Ambient o e) where
open Ambient ambient
private open import Algebra.ElgotAlgebra ambient
variable
o e : Level
module _ (D : ExtensiveDistributiveCategory o e) where
open ExtensiveDistributiveCategory D renaming (U to C; id to idC)
open Cocartesian (Extensive.cocartesian extensive)
open Cartesian (ExtensiveDistributiveCategory.cartesian D)
open BinaryProducts products
open M C open M C
open MR C open MR C
open HomReasoning open HomReasoning
@ -52,7 +42,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
```agda ```agda
-- iteration preversing morphism between two elgot-algebras -- iteration preversing morphism between two elgot-algebras
module _ (E₁ E₂ : Elgot-Algebra D) where module _ (E₁ E₂ : Elgot-Algebra) where
open Elgot-Algebra E₁ renaming (_# to _#₁) open Elgot-Algebra E₁ renaming (_# to _#₁)
open Elgot-Algebra E₂ renaming (_# to _#₂; A to B) open Elgot-Algebra E₂ renaming (_# to _#₂; A to B)
record Elgot-Algebra-Morphism : Set (o ⊔ ⊔ e) where record Elgot-Algebra-Morphism : Set (o ⊔ ⊔ e) where
@ -63,7 +53,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
-- the category of elgot algebras for a given category -- the category of elgot algebras for a given category
Elgot-Algebras : Category (o ⊔ ⊔ e) (o ⊔ ⊔ e) e Elgot-Algebras : Category (o ⊔ ⊔ e) (o ⊔ ⊔ e) e
Elgot-Algebras = record Elgot-Algebras = record
{ Obj = Elgot-Algebra D { Obj = Elgot-Algebra
; _⇒_ = Elgot-Algebra-Morphism ; _⇒_ = Elgot-Algebra-Morphism
; _≈_ = λ f g → Elgot-Algebra-Morphism.h f ≈ Elgot-Algebra-Morphism.h g ; _≈_ = λ f g → Elgot-Algebra-Morphism.h f ≈ Elgot-Algebra-Morphism.h g
; id = λ {EB} → let open Elgot-Algebra EB in ; id = λ {EB} → let open Elgot-Algebra EB in
@ -122,7 +112,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
open Terminal T open Terminal T
-- if the carriers of the algebra form a product, so do the algebras -- if the carriers of the algebra form a product, so do the algebras
A×B-Helper : ∀ {EA EB : Elgot-Algebra D} → Elgot-Algebra D A×B-Helper : ∀ {EA EB : Elgot-Algebra} → Elgot-Algebra
A×B-Helper {EA} {EB} = record A×B-Helper {EA} {EB} = record
{ A = A × B { A = A × B
; algebra = record ; algebra = record
@ -131,7 +121,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ≈⟨ ⟨⟩-cong₂ #ᵃ-Fixpoint #ᵇ-Fixpoint ⟩ ⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ≈⟨ ⟨⟩-cong₂ #ᵃ-Fixpoint #ᵇ-Fixpoint ⟩
⟨ [ idC , ((π₁ +₁ idC) ∘ f)#ᵃ ] ∘ ((π₁ +₁ idC) ∘ f) , [ idC , ((π₂ +₁ idC) ∘ f)#ᵇ ] ∘ ((π₂ +₁ idC) ∘ f) ⟩ ≈⟨ ⟨⟩-cong₂ (pullˡ []∘+₁) (pullˡ []∘+₁) ⟩ ⟨ [ idC , ((π₁ +₁ idC) ∘ f)#ᵃ ] ∘ ((π₁ +₁ idC) ∘ f) , [ idC , ((π₂ +₁ idC) ∘ f)#ᵇ ] ∘ ((π₂ +₁ idC) ∘ f) ⟩ ≈⟨ ⟨⟩-cong₂ (pullˡ []∘+₁) (pullˡ []∘+₁) ⟩
⟨ [ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] ∘ f , [ idC ∘ π₂ , ((π₂ +₁ idC) ∘ f)#ᵇ ∘ idC ] ∘ f ⟩ ≈˘⟨ ⟨⟩∘ ⟩ ⟨ [ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] ∘ f , [ idC ∘ π₂ , ((π₂ +₁ idC) ∘ f)#ᵇ ∘ idC ] ∘ f ⟩ ≈˘⟨ ⟨⟩∘ ⟩
(⟨ [ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] , [ idC ∘ π₂ , ((π₂ +₁ idC) ∘ f)#ᵇ ∘ idC ] ⟩ ∘ f) ≈⟨ ∘-resp-≈ˡ (unique (⟨ [ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] , [ idC ∘ π₂ , ((π₂ +₁ idC) ∘ f)#ᵇ ∘ idC ] ⟩ ∘ f) ≈⟨ ∘-resp-≈ˡ (⟨⟩-unique
(begin (begin
π₁ ∘ ⟨ [ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] , [ idC ∘ π₂ , ((π₂ +₁ idC) ∘ f)#ᵇ ∘ idC ] ⟩ ≈⟨ project₁ ⟩ π₁ ∘ ⟨ [ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] , [ idC ∘ π₂ , ((π₂ +₁ idC) ∘ f)#ᵇ ∘ idC ] ⟩ ≈⟨ project₁ ⟩
[ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] ≈⟨ []-cong₂ id-comm-sym (trans identityʳ (sym project₁)) ⟩ [ idC ∘ π₁ , ((π₁ +₁ idC) ∘ f)#ᵃ ∘ idC ] ≈⟨ []-cong₂ id-comm-sym (trans identityʳ (sym project₁)) ⟩
@ -144,7 +134,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
π₂ ∘ [ idC , ⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ] ∎) π₂ ∘ [ idC , ⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ] ∎)
)⟩ )⟩
([ idC , ⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ] ∘ f) ∎ ([ idC , ⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ] ∘ f) ∎
; #-Uniformity = λ {X Y f g h} uni → unique ; #-Uniformity = λ {X Y f g h} uni → ⟨⟩-unique
(begin (begin
π₁ ∘ ⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ≈⟨ project₁ ⟩ π₁ ∘ ⟨ ((π₁ +₁ idC) ∘ f)#ᵃ , ((π₂ +₁ idC) ∘ f)#ᵇ ⟩ ≈⟨ project₁ ⟩
(((π₁ +₁ idC) ∘ f)#ᵃ) ≈⟨ #ᵃ-Uniformity (((π₁ +₁ idC) ∘ f)#ᵃ) ≈⟨ #ᵃ-Uniformity
@ -192,7 +182,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
[ (idC +₁ i₁) ∘ ((π₂ +₁ idC) ∘ f) , i₂ ∘ h ] #ᵇ ≈⟨ #ᵇ-resp-≈ (+₁-id-swap π₂) ⟩ [ (idC +₁ i₁) ∘ ((π₂ +₁ idC) ∘ f) , i₂ ∘ h ] #ᵇ ≈⟨ #ᵇ-resp-≈ (+₁-id-swap π₂) ⟩
((π₂ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , i₂ ∘ h ])#ᵇ ∎ ((π₂ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , i₂ ∘ h ])#ᵇ ∎
Product-Elgot-Algebras : ∀ (EA EB : Elgot-Algebra D) → Product Elgot-Algebras EA EB Product-Elgot-Algebras : ∀ (EA EB : Elgot-Algebra) → Product Elgot-Algebras EA EB
Product-Elgot-Algebras EA EB = record Product-Elgot-Algebras EA EB = record
{ A×B = A×B-Helper {EA} {EB} { A×B = A×B-Helper {EA} {EB}
; π₁ = record { h = π₁ ; preserves = λ {X} {f} → project₁ } ; π₁ = record { h = π₁ ; preserves = λ {X} {f} → project₁ }
@ -209,7 +199,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
⟨ ((π₁ +₁ idC) ∘ (⟨ f , g ⟩ +₁ idC) ∘ h) #ᵃ , ((π₂ +₁ idC) ∘ (⟨ f , g ⟩ +₁ idC) ∘ h) #ᵇ ⟩ ∎ } ⟨ ((π₁ +₁ idC) ∘ (⟨ f , g ⟩ +₁ idC) ∘ h) #ᵃ , ((π₂ +₁ idC) ∘ (⟨ f , g ⟩ +₁ idC) ∘ h) #ᵇ ⟩ ∎ }
; project₁ = project₁ ; project₁ = project₁
; project₂ = project₂ ; project₂ = project₂
; unique = unique ; unique = ⟨⟩-unique
} }
where where
open Elgot-Algebra EA using (A) renaming (_# to _#ᵃ; #-Fixpoint to #ᵃ-Fixpoint; #-Uniformity to #ᵃ-Uniformity; #-Folding to #ᵃ-Folding; #-resp-≈ to #ᵃ-resp-≈) open Elgot-Algebra EA using (A) renaming (_# to _#ᵃ; #-Fixpoint to #ᵃ-Fixpoint; #-Uniformity to #ᵃ-Uniformity; #-Folding to #ᵃ-Folding; #-resp-≈ to #ᵃ-resp-≈)
@ -229,7 +219,7 @@ module _ (D : ExtensiveDistributiveCategory o e) where
```agda ```agda
-- if the carriers of the algebra form a exponential, so do the algebras -- if the carriers of the algebra form a exponential, so do the algebras
B^A-Helper : ∀ {EA : Elgot-Algebra D} {X : Obj} → Exponential C X (Elgot-Algebra.A EA) → Elgot-Algebra D B^A-Helper : ∀ {EA : Elgot-Algebra} {X : Obj} → Exponential C X (Elgot-Algebra.A EA) → Elgot-Algebra
B^A-Helper {EA} {X} exp = record B^A-Helper {EA} {X} exp = record
{ A = A^X { A = A^X
; algebra = record ; algebra = record

View file

@ -2,11 +2,7 @@
```agda ```agda
open import Level open import Level
open import Categories.Category.Core open import Categories.Category.Core
open import Categories.Category.Extensive.Bundle open import Category.Instance.AmbientCategory using (Ambient)
open import Categories.Category.Extensive
open import Categories.Category.Cocartesian
import Categories.Morphism.Reasoning as MR
open import UniformIterationAlgebra
``` ```
--> -->
@ -18,9 +14,9 @@ This file introduces the category of Uniform-Iteration Algebras
## Code ## Code
```agda ```agda
module UniformIterationAlgebras {o e} (D : ExtensiveDistributiveCategory o e) where module Category.Construction.UniformIterationAlgebras {o e} (ambient : Ambient o e) where
open ExtensiveDistributiveCategory D renaming (U to C; id to idC) open Ambient ambient
open Cocartesian (Extensive.cocartesian extensive) open import Algebra.UniformIterationAlgebra ambient
open HomReasoning open HomReasoning
open MR C open MR C
open Equiv open Equiv
@ -30,7 +26,7 @@ module UniformIterationAlgebras {o e} (D : ExtensiveDistributiveCategory o
```agda ```agda
-- iteration preversing morphism between two elgot-algebras -- iteration preversing morphism between two elgot-algebras
module _ (E₁ E₂ : Uniform-Iteration-Algebra D) where module _ (E₁ E₂ : Uniform-Iteration-Algebra) where
open Uniform-Iteration-Algebra E₁ renaming (_# to _#₁) open Uniform-Iteration-Algebra E₁ renaming (_# to _#₁)
open Uniform-Iteration-Algebra E₂ renaming (_# to _#₂; A to B) open Uniform-Iteration-Algebra E₂ renaming (_# to _#₂; A to B)
record Uniform-Iteration-Algebra-Morphism : Set (o ⊔ ⊔ e) where record Uniform-Iteration-Algebra-Morphism : Set (o ⊔ ⊔ e) where
@ -41,7 +37,7 @@ module UniformIterationAlgebras {o e} (D : ExtensiveDistributiveCategory o
-- the category of uniform-iteration algebras for a given category -- the category of uniform-iteration algebras for a given category
Uniform-Iteration-Algebras : Category (o ⊔ ⊔ e) (o ⊔ ⊔ e) e Uniform-Iteration-Algebras : Category (o ⊔ ⊔ e) (o ⊔ ⊔ e) e
Uniform-Iteration-Algebras = record Uniform-Iteration-Algebras = record
{ Obj = Uniform-Iteration-Algebra D { Obj = Uniform-Iteration-Algebra
; _⇒_ = Uniform-Iteration-Algebra-Morphism ; _⇒_ = Uniform-Iteration-Algebra-Morphism
; _≈_ = λ f g → Uniform-Iteration-Algebra-Morphism.h f ≈ Uniform-Iteration-Algebra-Morphism.h g ; _≈_ = λ f g → Uniform-Iteration-Algebra-Morphism.h f ≈ Uniform-Iteration-Algebra-Morphism.h g
; id = λ {EB} → let open Uniform-Iteration-Algebra EB in ; id = λ {EB} → let open Uniform-Iteration-Algebra EB in

View file

@ -0,0 +1,57 @@
<!--
```agda
open import Level
open import Categories.Category.Core
open import Categories.Category.Extensive using (Extensive)
open import Categories.Category.Extensive.Properties.Distributive using (Extensive×Cartesian⇒Distributive)
open import Categories.Category.Distributive using (Distributive)
open import Categories.Category.Cartesian using (Cartesian)
open import Categories.Category.BinaryProducts using (BinaryProducts)
open import Categories.Category.Cartesian.Bundle using (CartesianCategory)
open import Categories.Category.Cocartesian using (Cocartesian)
open import Categories.Object.NaturalNumbers.Parametrized using (ParametrizedNNO)
open import Categories.Object.Exponential using (Exponential)
open import Categories.Object.Terminal
import Categories.Morphism as M'
import Categories.Morphism.Reasoning as MR'
```
-->
## Summary
We work in an ambient category that
- is extensive (has finite coproducts and pullbacks along injections)
- is cartesian (has finite products, extensive + cartesian also gives a distributivity isomorphism)
- has a parametrized NNO
- has exponentials `X^`
```agda
module Category.Instance.AmbientCategory where
record Ambient (o e : Level) : Set (suc o ⊔ suc ⊔ suc e) where
field
C : Category o e
extensive : Extensive C
cartesian : Cartesian C
: ParametrizedNNO (record { U = C ; cartesian = cartesian })
_^ : ∀ X → Exponential C (ParametrizedNNO.N ) X
cartesianCategory : CartesianCategory o e
cartesianCategory = record { U = C ; cartesian = cartesian }
open Category C renaming (id to idC) public
open Extensive extensive public
open Cocartesian cocartesian renaming (+-unique to []-unique) public
open Cartesian cartesian public
-- open Terminal terminal public
-- Don't open the terminal object, because we are interested in many different terminal objects stemming from multiple categories
open BinaryProducts products renaming (η to ⁂-η; g-η to ⁂-g-η; unique to ⟨⟩-unique; unique to ⟨⟩-unique) public
open ParametrizedNNO public
distributive : Distributive C
distributive = Extensive×Cartesian⇒Distributive C extensive cartesian
open Distributive distributive hiding (cartesian; cocartesian) public
module M = M'
module MR = MR'
```

View file

@ -1,7 +1,7 @@
## Coproducts in the category of Coalgebras (needs proper imports to be compiled) ## Coproducts in the category of Coalgebras (needs proper imports to be compiled)
```agda ```agda
module Coalgebra where module Misc.Coalgebra where
``` ```
Coalg-cop : (F : Endofunctor C) → (alg₁ : F-Coalgebra F) → (alg₂ : F-Coalgebra F) → Coproduct (F-Coalgebras F) alg₁ alg₂ Coalg-cop : (F : Endofunctor C) → (alg₁ : F-Coalgebra F) → (alg₂ : F-Coalgebra F) → Coproduct (F-Coalgebras F) alg₁ alg₂

View file

@ -1,3 +1,6 @@
I thought briefly that I needed this, but I don't.
Might still be interesting!
```agda ```agda
open import Level open import Level
open import Categories.Category open import Categories.Category
@ -5,7 +8,7 @@ open import Categories.Functor
import Categories.Morphism as M import Categories.Morphism as M
import Categories.Morphism.Reasoning as MR import Categories.Morphism.Reasoning as MR
module FinalCoalgebras {o e} {C : Category o e} (F : Endofunctor C) where module Misc.FinalCoalgebras {o e} {C : Category o e} (F : Endofunctor C) where
open Category C renaming (id to idC) open Category C renaming (id to idC)
open import Categories.Object.Terminal open import Categories.Object.Terminal
open import Categories.Functor.Coalgebra open import Categories.Functor.Coalgebra

View file

@ -3,17 +3,9 @@
{-# OPTIONS --allow-unsolved-metas #-} {-# OPTIONS --allow-unsolved-metas #-}
open import Level open import Level
open import Categories.Category.Core open import Category.Instance.AmbientCategory using (Ambient)
open import Categories.Category.Extensive.Bundle
open import Categories.Category.BinaryProducts
open import Categories.Category.Cocartesian
open import Categories.Category.Cartesian
open import Categories.Category.Extensive
open import Categories.Monad open import Categories.Monad
open import Categories.Functor open import Categories.Functor
open import ElgotAlgebra
import Categories.Morphism.Reasoning as MR
``` ```
--> -->
@ -29,14 +21,12 @@ This file introduces Elgot Monads.
## Code ## Code
```agda ```agda
module Monad.ElgotMonad {o e} (ED : ExtensiveDistributiveCategory o e) where module Monad.ElgotMonad {o e} (ambient : Ambient o e) where
open ExtensiveDistributiveCategory ED renaming (U to C; id to idC) open Ambient ambient
open HomReasoning open HomReasoning
open Cocartesian (Extensive.cocartesian extensive)
open Cartesian (ExtensiveDistributiveCategory.cartesian ED)
open BinaryProducts products hiding (η)
open MR C open MR C
open Equiv open Equiv
open import Algebra.ElgotAlgebra ambient
``` ```
### *Definition 13*: Pre-Elgot Monads ### *Definition 13*: Pre-Elgot Monads
@ -48,7 +38,7 @@ module Monad.ElgotMonad {o e} (ED : ExtensiveDistributiveCategory o e) w
-- every TX needs to be equipped with an elgot algebra structure -- every TX needs to be equipped with an elgot algebra structure
field field
elgotalgebras : ∀ {X} → Elgot-Algebra-on ED (T₀ X) elgotalgebras : ∀ {X} → Elgot-Algebra-on (T₀ X)
module elgotalgebras {X} = Elgot-Algebra-on (elgotalgebras {X}) module elgotalgebras {X} = Elgot-Algebra-on (elgotalgebras {X})

View file

@ -3,16 +3,7 @@
open import Level open import Level
open import Categories.Category open import Categories.Category
open import Categories.Monad open import Categories.Monad
open import Categories.Category.Distributive
open import Categories.Category.Extensive.Bundle
open import Categories.Category.Extensive
open import Categories.Category.BinaryProducts
open import Categories.Category.Cocartesian
open import Categories.Category.Cartesian
open import Categories.Category.Cartesian.Bundle
open import Categories.Object.Terminal open import Categories.Object.Terminal
open import Categories.Object.Initial
open import Categories.Object.Coproduct
open import Categories.Category.Construction.F-Coalgebras open import Categories.Category.Construction.F-Coalgebras
open import Categories.Category.Construction.F-Algebras open import Categories.Category.Construction.F-Algebras
open import Categories.Functor.Coalgebra open import Categories.Functor.Coalgebra
@ -20,10 +11,7 @@ open import Categories.Functor
open import Categories.Functor.Algebra open import Categories.Functor.Algebra
open import Categories.Monad.Construction.Kleisli open import Categories.Monad.Construction.Kleisli
open import Categories.Category.Construction.F-Coalgebras open import Categories.Category.Construction.F-Coalgebras
open import Categories.NaturalTransformation open import Category.Instance.AmbientCategory using (Ambient)
open import FinalCoalgebras
import Categories.Morphism as M
import Categories.Morphism.Reasoning as MR
``` ```
--> -->
@ -33,19 +21,11 @@ This file introduces the delay monad ***D***
## Code ## Code
```agda ```agda
module Monad.Instance.Delay {o e} (ED : ExtensiveDistributiveCategory o e) where module Monad.Instance.Delay {o e} (ambient : Ambient o e) where
open Ambient ambient
``` ```
<!-- <!--
```agda ```agda
open ExtensiveDistributiveCategory ED renaming (U to C; id to idC)
open Cocartesian (Extensive.cocartesian extensive)
open Cartesian (ExtensiveDistributiveCategory.cartesian ED)
open BinaryProducts products
CC : CartesianCategory o e
CC = record { U = C ; cartesian = (ExtensiveDistributiveCategory.cartesian ED) }
open import Categories.Object.NaturalNumbers.Parametrized CC
open import Categories.Object.NaturalNumbers.Properties.F-Algebras using (PNNO⇒Initial₂; PNNO-Algebra) open import Categories.Object.NaturalNumbers.Properties.F-Algebras using (PNNO⇒Initial₂; PNNO-Algebra)
open M C open M C
@ -115,14 +95,11 @@ At the same time the morphism `X × N ⇒ X + X × N` is a coalgebra for the `(Y
TODO add diagram TODO add diagram
```agda ```agda
module _ ( : ParametrizedNNO) where iso : X × N ≅ X + X × N
open ParametrizedNNO iso = Lambek.lambek (record { ⊥ = PNNO-Algebra cartesianCategory coproducts X N z s ; ⊥-is-initial = PNNO⇒Initial₂ cartesianCategory coproducts X })
iso : X × N ≅ X + X × N
iso = Lambek.lambek (record { ⊥ = PNNO-Algebra CC coproducts X N z s ; ⊥-is-initial = PNNO⇒Initial₂ CC coproducts X })
ι : X × N ⇒ DX ι : X × N ⇒ DX
ι = u (! {A = record { A = X × N ; α = _≅_.from iso }}) ι = u (! {A = record { A = X × N ; α = _≅_.from iso }})
``` ```
With these definitions at hand, we can now indeed construct a monad (in extension form) as the triple `(F₀, unit, extend)`. With these definitions at hand, we can now indeed construct a monad (in extension form) as the triple `(F₀, unit, extend)`.
@ -132,9 +109,8 @@ TODO
```agda ```agda
kleisli : KleisliTriple C
monad : Monad C kleisli = record
monad = Kleisli⇒Monad C (record
{ F₀ = D₀ { F₀ = D₀
; unit = λ {X} → now X ; unit = λ {X} → now X
; extend = extend ; extend = extend
@ -143,7 +119,7 @@ TODO
; assoc = assoc' ; assoc = assoc'
; sym-assoc = sym assoc' ; sym-assoc = sym assoc'
; extend-≈ = extend-≈' ; extend-≈ = extend-≈'
}) }
where where
open Terminal open Terminal
module _ {X Y : Obj} (f : X ⇒ D₀ Y) where module _ {X Y : Obj} (f : X ⇒ D₀ Y) where
@ -285,6 +261,13 @@ TODO
⟩∘⟨refl ⟩ ⟩∘⟨refl ⟩
(u (! (coalgebras Y) {A = alg g }) ∘ idC) ∘ i₁ {B = D₀ Y} ≈⟨ identityʳ ⟩∘⟨refl ⟩ (u (! (coalgebras Y) {A = alg g }) ∘ idC) ∘ i₁ {B = D₀ Y} ≈⟨ identityʳ ⟩∘⟨refl ⟩
extend g ∎ extend g ∎
monad : Monad C
monad = Kleisli⇒Monad C kleisli
-- redundant but convenient to have
functor : Endofunctor C
functor = Monad.F monad
``` ```
### Definition 30: Search-Algebras ### Definition 30: Search-Algebras

View file

@ -0,0 +1,27 @@
<!--
```agda
open import Level
open import Categories.Functor
open import Category.Instance.AmbientCategory using (Ambient)
open import Categories.Monad.Construction.Kleisli
open import Categories.Monad.Relative renaming (Monad to RMonad)
```
-->
```agda
module Monad.Instance.Delay.Quotienting {o e} (ambient : Ambient o e) where
open Ambient ambient
open import Categories.Diagram.Coequalizer C
open import Monad.Instance.Delay ambient
module _ (D : DelayM) where
open DelayM D
open Functor functor using () renaming (F₁ to D₁)
open RMonad kleisli
module _ {X : Obj} (coeq : Coequalizer (extend (ι X)) (D₁ π₁)) where
-- TODO
```

View file

@ -1,12 +1,8 @@
<!-- <!--
```agda ```agda
open import Level open import Level
open import Categories.Category.Core
open import Categories.Category.Equivalence using (StrongEquivalence) open import Categories.Category.Equivalence using (StrongEquivalence)
open import Categories.Category.Extensive.Bundle
open import Function using (id) open import Function using (id)
open import UniformIterationAlgebras
open import UniformIterationAlgebra
open import Categories.FreeObjects.Free open import Categories.FreeObjects.Free
open import Categories.Functor.Core open import Categories.Functor.Core
open import Categories.Adjoint open import Categories.Adjoint
@ -16,6 +12,7 @@ open import Categories.NaturalTransformation.Core renaming (id to idN)
open import Categories.Monad open import Categories.Monad
open import Categories.Category.Construction.EilenbergMoore open import Categories.Category.Construction.EilenbergMoore
open import Categories.Category.Slice open import Categories.Category.Slice
open import Category.Instance.AmbientCategory using (Ambient)
``` ```
--> -->
@ -34,13 +31,15 @@ In this file I explore the monad ***K*** and its properties:
## Code ## Code
```agda ```agda
module MonadK {o e} (D : ExtensiveDistributiveCategory o e) where module Monad.Instance.K {o e} (ambient : Ambient o e) where
open ExtensiveDistributiveCategory D renaming (U to C; id to idC) open Ambient ambient
open import Category.Construction.UniformIterationAlgebras ambient
open import Algebra.UniformIterationAlgebra ambient
open Equiv open Equiv
-- TODO move this to a different file -- TODO move this to a different file
forgetfulF : Functor (Uniform-Iteration-Algebras D) C forgetfulF : Functor Uniform-Iteration-Algebras C
forgetfulF = record forgetfulF = record
{ F₀ = λ X → Uniform-Iteration-Algebra.A X { F₀ = λ X → Uniform-Iteration-Algebra.A X
; F₁ = λ f → Uniform-Iteration-Algebra-Morphism.h f ; F₁ = λ f → Uniform-Iteration-Algebra-Morphism.h f
@ -51,7 +50,7 @@ module MonadK {o e} (D : ExtensiveDistributiveCategory o e) where
-- typedef -- typedef
FreeUniformIterationAlgebra : Obj → Set (suc o ⊔ suc ⊔ suc e) FreeUniformIterationAlgebra : Obj → Set (suc o ⊔ suc ⊔ suc e)
FreeUniformIterationAlgebra X = FreeObject {C = C} {D = Uniform-Iteration-Algebras D} forgetfulF X FreeUniformIterationAlgebra X = FreeObject {C = C} {D = Uniform-Iteration-Algebras} forgetfulF X
``` ```
@ -62,7 +61,7 @@ module MonadK {o e} (D : ExtensiveDistributiveCategory o e) where
field field
algebras : ∀ X → FreeUniformIterationAlgebra X algebras : ∀ X → FreeUniformIterationAlgebra X
freeF : Functor C (Uniform-Iteration-Algebras D) freeF : Functor C Uniform-Iteration-Algebras
freeF = FO⇒Functor forgetfulF algebras freeF = FO⇒Functor forgetfulF algebras
adjoint : freeF ⊣ forgetfulF adjoint : freeF ⊣ forgetfulF