```agda module Algebra.Properties {o ℓ e} (ambient : Ambient o ℓ e) where open Ambient ambient open import Algebra.ElgotAlgebra ambient open import Algebra.UniformIterationAlgebra ambient open import Category.Construction.ElgotAlgebras ambient open import Category.Construction.UniformIterationAlgebras ambient open Equiv ``` # Properties of algebras This file contains some typedefs and records concerning different algebras. ## Free Uniform-Iteration Algebras ```agda uniformForgetfulF : Functor Uniform-Iteration-Algebras C uniformForgetfulF = record { F₀ = λ X → Uniform-Iteration-Algebra.A X ; F₁ = λ f → Uniform-Iteration-Algebra-Morphism.h f ; identity = refl ; homomorphism = refl ; F-resp-≈ = λ x → x } -- typedef FreeUniformIterationAlgebra : Obj → Set (suc o ⊔ suc ℓ ⊔ suc e) FreeUniformIterationAlgebra X = FreeObject {C = C} {D = Uniform-Iteration-Algebras} uniformForgetfulF X ``` ## Stable Free Uniform-Iteration Algebras ```agda record IsStableFreeUniformIterationAlgebra {Y : Obj} (freeAlgebra : FreeUniformIterationAlgebra Y) : Set (suc o ⊔ suc ℓ ⊔ suc e) where open FreeObject freeAlgebra using (η) renaming (FX to FY) open Uniform-Iteration-Algebra FY using (_#) field -- TODO awkward notation... [_,_]♯ : ∀ {X : Obj} (A : Uniform-Iteration-Algebra) (f : X × Y ⇒ Uniform-Iteration-Algebra.A A) → X × Uniform-Iteration-Algebra.A FY ⇒ Uniform-Iteration-Algebra.A A ♯-law : ∀ {X : Obj} {A : Uniform-Iteration-Algebra} (f : X × Y ⇒ Uniform-Iteration-Algebra.A A) → f ≈ [ A , f ]♯ ∘ (idC ⁂ η) ♯-preserving : ∀ {X : Obj} {B : Uniform-Iteration-Algebra} (f : X × Y ⇒ Uniform-Iteration-Algebra.A B) {Z : Obj} (h : Z ⇒ Uniform-Iteration-Algebra.A FY + Z) → [ B , f ]♯ ∘ (idC ⁂ h #) ≈ Uniform-Iteration-Algebra._# B (([ B , f ]♯ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) ♯-unique : ∀ {X : Obj} {B : Uniform-Iteration-Algebra} (f : X × Y ⇒ Uniform-Iteration-Algebra.A B) (g : X × Uniform-Iteration-Algebra.A FY ⇒ Uniform-Iteration-Algebra.A B) → f ≈ g ∘ (idC ⁂ η) → (∀ {Z : Obj} (h : Z ⇒ Uniform-Iteration-Algebra.A FY + Z) → g ∘ (idC ⁂ h #) ≈ Uniform-Iteration-Algebra._# B ((g +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))) → g ≈ [ B , f ]♯ record StableFreeUniformIterationAlgebra : Set (suc o ⊔ suc ℓ ⊔ suc e) where field Y : Obj freeAlgebra : FreeUniformIterationAlgebra Y isStableFreeUniformIterationAlgebra : IsStableFreeUniformIterationAlgebra freeAlgebra open IsStableFreeUniformIterationAlgebra isStableFreeUniformIterationAlgebra public ``` ## Free Elgot Algebras ```agda elgotForgetfulF : Functor Elgot-Algebras C elgotForgetfulF = record { F₀ = λ X → Elgot-Algebra.A X ; F₁ = λ f → Elgot-Algebra-Morphism.h f ; identity = refl ; homomorphism = refl ; F-resp-≈ = λ x → x } -- typedef FreeElgotAlgebra : Obj → Set (suc o ⊔ suc ℓ ⊔ suc e) FreeElgotAlgebra X = FreeObject {C = C} {D = Elgot-Algebras} elgotForgetfulF X ``` ## Stable Free Elgot Algebras -- TODO this is duplicated from uniform iteration and since free elgots coincide with free uniform iterations this can later be removed. ```agda record IsStableFreeElgotAlgebra {Y : Obj} (freeElgot : FreeElgotAlgebra Y) : Set (suc o ⊔ suc ℓ ⊔ suc e) where open FreeObject freeElgot using (η) renaming (FX to FY) open Elgot-Algebra FY using (_#) field -- TODO awkward notation... [_,_]♯ : ∀ {X : Obj} (A : Elgot-Algebra) (f : X × Y ⇒ Elgot-Algebra.A A) → X × Elgot-Algebra.A FY ⇒ Elgot-Algebra.A A ♯-law : ∀ {X : Obj} {A : Elgot-Algebra} (f : X × Y ⇒ Elgot-Algebra.A A) → f ≈ [ A , f ]♯ ∘ (idC ⁂ η) ♯-preserving : ∀ {X : Obj} {B : Elgot-Algebra} (f : X × Y ⇒ Elgot-Algebra.A B) {Z : Obj} (h : Z ⇒ Elgot-Algebra.A FY + Z) → [ B , f ]♯ ∘ (idC ⁂ h #) ≈ Elgot-Algebra._# B (([ B , f ]♯ +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h)) ♯-unique : ∀ {X : Obj} {B : Elgot-Algebra} (f : X × Y ⇒ Elgot-Algebra.A B) (g : X × Elgot-Algebra.A FY ⇒ Elgot-Algebra.A B) → f ≈ g ∘ (idC ⁂ η) → (∀ {Z : Obj} (h : Z ⇒ Elgot-Algebra.A FY + Z) → g ∘ (idC ⁂ h #) ≈ Elgot-Algebra._# B ((g +₁ idC) ∘ distributeˡ⁻¹ ∘ (idC ⁂ h))) → g ≈ [ B , f ]♯ record StableFreeElgotAlgebra : Set (suc o ⊔ suc ℓ ⊔ suc e) where field Y : Obj freeElgot : FreeElgotAlgebra Y isStableFreeElgotAlgebra : IsStableFreeElgotAlgebra freeElgot open IsStableFreeElgotAlgebra isStableFreeElgotAlgebra public ``` ## Free Elgot to Free Uniform Iteration ```agda elgot-to-uniformF : Functor Elgot-Algebras Uniform-Iteration-Algebras elgot-to-uniformF = record { F₀ = λ X → record { A = Elgot-Algebra.A X ; algebra = record { _# = λ f → (X Elgot-Algebra.#) f ; #-Fixpoint = Elgot-Algebra.#-Fixpoint X ; #-Uniformity = Elgot-Algebra.#-Uniformity X ; #-resp-≈ = Elgot-Algebra.#-resp-≈ X }} ; F₁ = λ f → record { h = Elgot-Algebra-Morphism.h f ; preserves = Elgot-Algebra-Morphism.preserves f } ; identity = refl ; homomorphism = refl ; F-resp-≈ = λ x → x } ``` FreeElgots⇒FreeUniformIterations : (∀ X → FreeElgotAlgebra X) → (∀ X → FreeUniformIterationAlgebra X) FreeElgots⇒FreeUniformIterations free-elgots X = record { FX = F₀ (FreeObject.FX (free-elgots X)) ; η = FreeObject.η (free-elgots X) ; _* = λ {FY} f → F₁ (FreeObject._* (free-elgots X) {A = record { A = Uniform-Iteration-Algebra.A FY ; algebra = record { _# = FY Uniform-Iteration-Algebra.# ; #-Fixpoint = Uniform-Iteration-Algebra.#-Fixpoint FY ; #-Uniformity = Uniform-Iteration-Algebra.#-Uniformity FY ; #-Folding = {! !} ; #-resp-≈ = Uniform-Iteration-Algebra.#-resp-≈ FY } }} f) -- FreeObject.FX (free-elgots (Uniform-Iteration-Algebra.A FY)) ; *-lift = {! !} ; *-uniq = {! !} } where open Functor elgot-to-uniformF open Functor (FO⇒Functor elgotForgetfulF free-elgots) using () renaming (F₀ to FO₀; F₁ to FO₁)