mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
60 lines
No EOL
3 KiB
Agda
60 lines
No EOL
3 KiB
Agda
open import Level
|
||
open import Categories.Category.Core
|
||
open import Categories.Category.Extensive.Bundle
|
||
open import Categories.Category.Extensive
|
||
open import Categories.Category.Cocartesian
|
||
import Categories.Morphism.Reasoning as MR
|
||
|
||
open import UniformIterationAlgebra
|
||
|
||
module UniformIterationAlgebras {o ℓ e} (D : ExtensiveDistributiveCategory o ℓ e) where
|
||
open ExtensiveDistributiveCategory D renaming (U to C; id to idC)
|
||
open Cocartesian (Extensive.cocartesian extensive)
|
||
open HomReasoning
|
||
open MR C
|
||
open Equiv
|
||
|
||
-- iteration preversing morphism between two elgot-algebras
|
||
module _ (E₁ E₂ : Uniform-Iteration-Algebra D) where
|
||
open Uniform-Iteration-Algebra E₁ renaming (_# to _#₁)
|
||
open Uniform-Iteration-Algebra E₂ renaming (_# to _#₂; A to B)
|
||
record Uniform-Iteration-Algebra-Morphism : Set (o ⊔ ℓ ⊔ e) where
|
||
field
|
||
h : A ⇒ B
|
||
preserves : ∀ {X} {f : X ⇒ A + X} → h ∘ (f #₁) ≈ ((h +₁ idC) ∘ f)#₂
|
||
|
||
-- the category of uniform-iteration algebras for a given category
|
||
Uniform-Iteration-Algebras : Category (o ⊔ ℓ ⊔ e) (o ⊔ ℓ ⊔ e) e
|
||
Uniform-Iteration-Algebras = record
|
||
{ Obj = Uniform-Iteration-Algebra D
|
||
; _⇒_ = Uniform-Iteration-Algebra-Morphism
|
||
; _≈_ = λ f g → Uniform-Iteration-Algebra-Morphism.h f ≈ Uniform-Iteration-Algebra-Morphism.h g
|
||
; id = λ {EB} → let open Uniform-Iteration-Algebra EB in
|
||
record { h = idC; preserves = λ {X : Obj} {f : X ⇒ A + X} → begin
|
||
idC ∘ f # ≈⟨ identityˡ ⟩
|
||
f # ≈⟨ #-resp-≈ (introˡ (coproduct.unique id-comm-sym id-comm-sym)) ⟩
|
||
((idC +₁ idC) ∘ f) # ∎ }
|
||
; _∘_ = λ {EA} {EB} {EC} f g → let
|
||
open Uniform-Iteration-Algebra-Morphism f renaming (h to hᶠ; preserves to preservesᶠ)
|
||
open Uniform-Iteration-Algebra-Morphism g renaming (h to hᵍ; preserves to preservesᵍ)
|
||
open Uniform-Iteration-Algebra EA using (A) renaming (_# to _#ᵃ)
|
||
open Uniform-Iteration-Algebra EB using () renaming (_# to _#ᵇ; A to B)
|
||
open Uniform-Iteration-Algebra EC using () renaming (_# to _#ᶜ; A to C; #-resp-≈ to #ᶜ-resp-≈)
|
||
in record { h = hᶠ ∘ hᵍ; preserves = λ {X} {f : X ⇒ A + X} → begin
|
||
(hᶠ ∘ hᵍ) ∘ (f #ᵃ) ≈⟨ pullʳ preservesᵍ ⟩
|
||
(hᶠ ∘ (((hᵍ +₁ idC) ∘ f) #ᵇ)) ≈⟨ preservesᶠ ⟩
|
||
(((hᶠ +₁ idC) ∘ (hᵍ +₁ idC) ∘ f) #ᶜ) ≈⟨ #ᶜ-resp-≈ (pullˡ (trans +₁∘+₁ (+₁-cong₂ refl (identity²)))) ⟩
|
||
((hᶠ ∘ hᵍ +₁ idC) ∘ f) #ᶜ ∎ }
|
||
; identityˡ = identityˡ
|
||
; identityʳ = identityʳ
|
||
; identity² = identity²
|
||
; assoc = assoc
|
||
; sym-assoc = sym-assoc
|
||
; equiv = record
|
||
{ refl = refl
|
||
; sym = sym
|
||
; trans = trans
|
||
}
|
||
; ∘-resp-≈ = ∘-resp-≈
|
||
}
|
||
where open Uniform-Iteration-Algebra-Morphism |