bsc-leon-vatthauer/UniformIterationAlgebra.agda

29 lines
1.1 KiB
Agda
Raw Normal View History

2023-08-17 18:07:14 +02:00
open import Level
open import Categories.Category.Core
open import Categories.Category.Extensive.Bundle
open import Categories.Category.Extensive
open import Categories.Category.Cocartesian
module UniformIterationAlgebra {o e} (D : ExtensiveDistributiveCategory o e) where
open ExtensiveDistributiveCategory D renaming (U to C; id to idC)
open Cocartesian (Extensive.cocartesian extensive)
record Uniform-Iteration-Algebra-on (A : Obj) : Set (o e) where
-- iteration operator
field
_# : {X} (X A + X) (X A)
-- _# properties
field
#-Fixpoint : {X} {f : X A + X }
f # [ idC , f # ] f
#-Uniformity : {X Y} {f : X A + X} {g : Y A + Y} {h : X Y}
(idC +₁ h) f g h
f # g # h
#-resp-≈ : {X} {f g : X A + X} f g (f #) (g #)
record Uniform-Iteration-Algebra : Set (o e) where
field
A : Obj
algebra : Uniform-Iteration-Algebra-on A
open Uniform-Iteration-Algebra-on algebra public