mirror of
https://git8.cs.fau.de/theses/bsc-leon-vatthauer.git
synced 2024-05-31 07:28:34 +02:00
Added distributivity
This commit is contained in:
parent
0be1871679
commit
07dac8250d
4 changed files with 92 additions and 13 deletions
20
Distributive/Bundle.agda
Normal file
20
Distributive/Bundle.agda
Normal file
|
@ -0,0 +1,20 @@
|
||||||
|
{-# OPTIONS --without-K --safe #-}
|
||||||
|
|
||||||
|
open import Categories.Category
|
||||||
|
open import Categories.Category.Cartesian
|
||||||
|
open import Categories.Category.BinaryProducts
|
||||||
|
open import Categories.Category.Cocartesian
|
||||||
|
open import Distributive.Core using (Distributive)
|
||||||
|
import Categories.Morphism as M
|
||||||
|
|
||||||
|
module Distributive.Bundle where
|
||||||
|
|
||||||
|
open import Level
|
||||||
|
|
||||||
|
record DistributiveCategory o ℓ e : Set (suc (o ⊔ ℓ ⊔ e)) where
|
||||||
|
field
|
||||||
|
U : Category o ℓ e
|
||||||
|
distributive : Distributive U
|
||||||
|
|
||||||
|
open Category U public
|
||||||
|
open Distributive distributive public
|
28
Distributive/Core.agda
Normal file
28
Distributive/Core.agda
Normal file
|
@ -0,0 +1,28 @@
|
||||||
|
{-# OPTIONS --without-K --safe #-}
|
||||||
|
|
||||||
|
open import Categories.Category
|
||||||
|
open import Categories.Category.Cartesian
|
||||||
|
open import Categories.Category.BinaryProducts
|
||||||
|
open import Categories.Category.Cocartesian
|
||||||
|
import Categories.Morphism as M
|
||||||
|
|
||||||
|
module Distributive.Core {o ℓ e} (𝒞 : Category o ℓ e) where
|
||||||
|
|
||||||
|
open import Level
|
||||||
|
open Category 𝒞
|
||||||
|
open M 𝒞
|
||||||
|
|
||||||
|
record Distributive : Set (levelOfTerm 𝒞) where
|
||||||
|
field
|
||||||
|
cartesian : Cartesian 𝒞
|
||||||
|
cocartesian : Cocartesian 𝒞
|
||||||
|
|
||||||
|
open Cartesian cartesian
|
||||||
|
open BinaryProducts products
|
||||||
|
open Cocartesian cocartesian
|
||||||
|
|
||||||
|
distribute : ∀ {A B C} → (A × B + A × C) ⇒ (A × (B + C))
|
||||||
|
distribute = [ id ⁂ i₁ , id ⁂ i₂ ]
|
||||||
|
|
||||||
|
field
|
||||||
|
iso : ∀ {A B C} → IsIso (distribute {A} {B} {C})
|
|
@ -8,19 +8,26 @@ open import Categories.Functor.Algebra
|
||||||
open import Categories.Object.Product
|
open import Categories.Object.Product
|
||||||
open import Categories.Object.Coproduct
|
open import Categories.Object.Coproduct
|
||||||
open import Categories.Category
|
open import Categories.Category
|
||||||
|
open import Distributive.Bundle
|
||||||
|
open import Distributive.Core
|
||||||
|
open import Categories.Category.Cartesian
|
||||||
|
open import Categories.Category.BinaryProducts
|
||||||
|
open import Categories.Category.Cocartesian
|
||||||
|
|
||||||
private
|
private
|
||||||
variable
|
variable
|
||||||
o ℓ e : Level
|
o ℓ e : Level
|
||||||
|
|
||||||
|
|
||||||
module _ (C𝒞 : CocartesianCategory o ℓ e) where
|
module _ (D : DistributiveCategory o ℓ e) where
|
||||||
open CocartesianCategory C𝒞 renaming (U to 𝒞; id to idC)
|
open DistributiveCategory D renaming (U to C; id to idC)
|
||||||
|
open Cocartesian cocartesian
|
||||||
|
open Cartesian cartesian
|
||||||
|
|
||||||
--*
|
--*
|
||||||
-- F-guarded Elgot Algebra
|
-- F-guarded Elgot Algebra
|
||||||
--*
|
--*
|
||||||
module _ {F : Endofunctor 𝒞} (FA : F-Algebra F) where
|
module _ {F : Endofunctor C} (FA : F-Algebra F) where
|
||||||
record Guarded-Elgot-Algebra : Set (o ⊔ ℓ ⊔ e) where
|
record Guarded-Elgot-Algebra : Set (o ⊔ ℓ ⊔ e) where
|
||||||
open Functor F public
|
open Functor F public
|
||||||
open F-Algebra FA public
|
open F-Algebra FA public
|
||||||
|
@ -116,8 +123,8 @@ module _ (C𝒞 : CocartesianCategory o ℓ e) where
|
||||||
|
|
||||||
private
|
private
|
||||||
-- identity functor on 𝒞
|
-- identity functor on 𝒞
|
||||||
Id : Functor 𝒞 𝒞
|
Id : Functor C C
|
||||||
Id = idF {C = 𝒞}
|
Id = idF {C = C}
|
||||||
|
|
||||||
-- identity algebra
|
-- identity algebra
|
||||||
Id-Algebra : Obj → F-Algebra Id
|
Id-Algebra : Obj → F-Algebra Id
|
||||||
|
|
|
@ -2,16 +2,20 @@ open import Level renaming (suc to ℓ-suc)
|
||||||
open import Function using (_$_) renaming (id to idf; _∘_ to _∘ᶠ_)
|
open import Function using (_$_) renaming (id to idf; _∘_ to _∘ᶠ_)
|
||||||
open import Data.Product using (_,_) renaming (_×_ to _∧_)
|
open import Data.Product using (_,_) renaming (_×_ to _∧_)
|
||||||
|
|
||||||
open import Categories.Category.Cocartesian.Bundle using (CocartesianCategory)
|
open import Categories.Category.Cocartesian
|
||||||
|
open import Categories.Category.Cocartesian.Bundle
|
||||||
open import Categories.Category.Cartesian
|
open import Categories.Category.Cartesian
|
||||||
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.Object.Terminal
|
open import Categories.Object.Terminal
|
||||||
open import Categories.Object.Product
|
open import Categories.Object.Product
|
||||||
|
open import Categories.Object.Exponential
|
||||||
open import Categories.Object.Coproduct
|
open import Categories.Object.Coproduct
|
||||||
open import Categories.Category.BinaryProducts
|
open import Categories.Category.BinaryProducts
|
||||||
open import Categories.Category
|
open import Categories.Category
|
||||||
open import ElgotAlgebra
|
open import ElgotAlgebra
|
||||||
|
open import Distributive.Bundle
|
||||||
|
open import Distributive.Core
|
||||||
|
|
||||||
module ElgotAlgebras where
|
module ElgotAlgebras where
|
||||||
|
|
||||||
|
@ -19,15 +23,22 @@ private
|
||||||
variable
|
variable
|
||||||
o ℓ e : Level
|
o ℓ e : Level
|
||||||
|
|
||||||
module _ (CC : CocartesianCategory o ℓ e) where
|
module _ (D : DistributiveCategory o ℓ e) where
|
||||||
open CocartesianCategory CC renaming (U to C; id to idC)
|
open DistributiveCategory D renaming (U to C; id to idC)
|
||||||
|
open Cocartesian cocartesian
|
||||||
|
|
||||||
|
CC : CocartesianCategory o ℓ e
|
||||||
|
CC = record
|
||||||
|
{ U = C
|
||||||
|
; cocartesian = cocartesian
|
||||||
|
}
|
||||||
|
|
||||||
--*
|
--*
|
||||||
-- let's define the category of elgot-algebras
|
-- let's define the category of elgot-algebras
|
||||||
--*
|
--*
|
||||||
|
|
||||||
-- iteration preversing morphism between two elgot-algebras
|
-- iteration preversing morphism between two elgot-algebras
|
||||||
module _ (E₁ E₂ : Elgot-Algebra CC) where
|
module _ (E₁ E₂ : Elgot-Algebra D) 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
|
||||||
|
@ -38,7 +49,7 @@ module _ (CC : CocartesianCategory o ℓ e) where
|
||||||
-- the category of elgot algebras for a given (cocartesian-)category
|
-- the category of elgot algebras for a given (cocartesian-)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 CC
|
{ Obj = Elgot-Algebra D
|
||||||
; _⇒_ = 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
|
||||||
|
@ -105,7 +116,7 @@ module _ (CC : CocartesianCategory o ℓ e) where
|
||||||
open Equiv
|
open Equiv
|
||||||
|
|
||||||
-- 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 CC} → Product C (Elgot-Algebra.A EA) (Elgot-Algebra.A EB) → Elgot-Algebra CC
|
A×B-Helper : ∀ {EA EB : Elgot-Algebra D} → Product C (Elgot-Algebra.A EA) (Elgot-Algebra.A EB) → Elgot-Algebra D
|
||||||
A×B-Helper {EA} {EB} p = record
|
A×B-Helper {EA} {EB} p = record
|
||||||
{ A = A×B
|
{ A = A×B
|
||||||
; _# = λ {X : Obj} (h : X ⇒ A×B + X) → ⟨ ((π₁ +₁ idC) ∘ h)#ᵃ , ((π₂ +₁ idC) ∘ h)#ᵇ ⟩
|
; _# = λ {X : Obj} (h : X ⇒ A×B + X) → ⟨ ((π₁ +₁ idC) ∘ h)#ᵃ , ((π₂ +₁ idC) ∘ h)#ᵇ ⟩
|
||||||
|
@ -202,7 +213,7 @@ module _ (CC : CocartesianCategory o ℓ e) where
|
||||||
(([ (π₂ +₁ idC) ∘ (idC +₁ i₁) ∘ f , (π₂ +₁ idC) ∘ i₂ ∘ h ])#ᵇ) ≈⟨ sym (#ᵇ-resp-≈ ∘[]) ⟩
|
(([ (π₂ +₁ idC) ∘ (idC +₁ i₁) ∘ f , (π₂ +₁ idC) ∘ i₂ ∘ h ])#ᵇ) ≈⟨ sym (#ᵇ-resp-≈ ∘[]) ⟩
|
||||||
((π₂ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , i₂ ∘ h ])#ᵇ ∎
|
((π₂ +₁ idC) ∘ [ (idC +₁ i₁) ∘ f , i₂ ∘ h ])#ᵇ ∎
|
||||||
|
|
||||||
Product-Elgot-Algebras : ∀ (EA EB : Elgot-Algebra CC) → Product C (Elgot-Algebra.A EA) (Elgot-Algebra.A EB) → Product Elgot-Algebras EA EB
|
Product-Elgot-Algebras : ∀ (EA EB : Elgot-Algebra D) → Product C (Elgot-Algebra.A EA) (Elgot-Algebra.A EB) → Product Elgot-Algebras EA EB
|
||||||
Product-Elgot-Algebras EA EB p = record
|
Product-Elgot-Algebras EA EB p = record
|
||||||
{ A×B = A×B-Helper {EA} {EB} p
|
{ A×B = A×B-Helper {EA} {EB} p
|
||||||
; π₁ = record { h = π₁ ; preserves = λ {X} {f} → project₁ }
|
; π₁ = record { h = π₁ ; preserves = λ {X} {f} → project₁ }
|
||||||
|
@ -241,3 +252,16 @@ module _ (CC : CocartesianCategory o ℓ e) where
|
||||||
open Cartesian CaC using (terminal; products)
|
open Cartesian CaC using (terminal; products)
|
||||||
open BinaryProducts products using (product)
|
open BinaryProducts products using (product)
|
||||||
open Equiv
|
open Equiv
|
||||||
|
|
||||||
|
-- if the carriers of the algebra form a exponential, so do the algebras
|
||||||
|
B^A-Helper : ∀ {EA EB : Elgot-Algebra D} → Exponential C (Elgot-Algebra.A EA) (Elgot-Algebra.A EB) → Elgot-Algebra D
|
||||||
|
B^A-Helper {EA} {EB} e = record
|
||||||
|
{ A = B^A
|
||||||
|
; _# = λ {X} f → λg {! !} {! !}
|
||||||
|
; #-Fixpoint = {! !}
|
||||||
|
; #-Uniformity = {! !}
|
||||||
|
; #-Folding = {! !}
|
||||||
|
; #-resp-≈ = {! !}
|
||||||
|
}
|
||||||
|
where
|
||||||
|
open Exponential e
|
Loading…
Reference in a new issue