Split up some of Category (#72)

This commit is contained in:
Patrick Stevens
2019-11-13 20:14:52 +00:00
committed by GitHub
parent d0f8d5ba39
commit 42f1defd2a
6 changed files with 81 additions and 39 deletions

View File

@@ -11,13 +11,7 @@ open import Categories.Definition
module Categories.Category where
postulate
extensionality : {a b : _} {S : Set a}{T : S Set b} {f g : (x : S) T x} ((x : S) f x g x) f g
dual : {a b : _} Category {a} {b} Category {a} {b}
dual record { objects = objects ; arrows = arrows ; id = id ; _∘_ = _∘_ ; rightId = rightId ; leftId = leftId ; compositionAssociative = associative } = record { objects = objects ; arrows = λ i j arrows j i ; id = id ; _∘_ = λ {x y z} g f f g ; rightId = λ {x y} f leftId f ; leftId = λ {x y} f rightId f ; associative = λ {x y z w} f g h equalityCommutative (associative h g f) }
SET : {a : _} Category {lsuc a} {a}
SET {a} = record { objects = Set a ; arrows = λ i j (i j) ; id = λ X x x ; _∘_ = λ g f x g (f x) ; rightId = λ f refl ; leftId = λ f refl ; associative = λ f g h refl }
extensionality : {a b : _} {S : Set a} {T : S Set b} {f g : (x : S) T x} ((x : S) f x g x) f g
≡Unique : {a : _} {X : Set a} {a b : X} (p1 p2 : a b) (p1 p2)
≡Unique refl refl = refl
@@ -50,30 +44,6 @@ NatPreorder = record { objects = ; arrows = λ m n → m ≤N n ; id = λ x
NatMonoid : Category {lzero} {lzero}
NatMonoid = record { objects = True ; arrows = λ _ _ ; id = λ x 0 ; _∘_ = λ f g f +N g ; rightId = λ f refl ; leftId = λ f Semiring.sumZeroRight Semiring f ; associative = λ a b c Semiring.+Associative Semiring a b c }
DISCRETE : {a : _} (X : Set a) Category {a} {a}
DISCRETE X = record { objects = X ; arrows = _≡_ ; id = λ x refl ; _∘_ = λ f g transitivity g f ; rightId = λ f ≡Unique (transitivity f refl) f ; leftId = λ f ≡Unique (transitivity refl f) f ; associative = λ f g h ≡Unique (transitivity (transitivity h g) f) (transitivity h (transitivity g f)) }
record Functor {a b c d : _} (C : Category {a} {b}) (D : Category {c} {d}) : Set (a b c d) where
field
onObj : Category.objects C Category.objects D
onArrow : {S T : Category.objects C} Category.arrows C S T Category.arrows D (onObj S) (onObj T)
mapId : {T : Category.objects C} onArrow (Category.id C T) Category.id D (onObj T)
mapCompose : {X Y Z : Category.objects C} (f : Category.arrows C X Y) (g : Category.arrows C Y Z) onArrow (Category._∘_ C g f) Category._∘_ D (onArrow g) (onArrow f)
functorCompose : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (Functor C D) (Functor B C) (Functor B D)
functorCompose G F = record { onObj = λ x Functor.onObj G (Functor.onObj F x) ; onArrow = λ f Functor.onArrow G (Functor.onArrow F f) ; mapId = λ {T} mapIdHelp G F T ; mapCompose = λ r s mapComposeHelp G F r s }
where
mapIdHelp : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (G : Functor C D) (F : Functor B C) (T : Category.objects B) Functor.onArrow G (Functor.onArrow F (Category.id B T)) Category.id D (Functor.onObj G (Functor.onObj F T))
mapIdHelp {B = B} {C} {D} record { onObj = onObjG ; onArrow = onArrowG ; mapId = mapIdG ; mapCompose = mapComposeG } record { onObj = onObj ; onArrow = onArrow ; mapId = mapId ; mapCompose = mapCompose } T rewrite mapId {T} = mapIdG {onObj T}
mapComposeHelp : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (G : Functor C D) (F : Functor B C) {S T U : Category.objects B} (r : Category.arrows B S T) (s : Category.arrows B T U) (Functor.onArrow G (Functor.onArrow F (Category._∘_ B s r))) (Category._∘_ D (Functor.onArrow G (Functor.onArrow F s)) (Functor.onArrow G (Functor.onArrow F r)))
mapComposeHelp {B = record { objects = objectsB ; arrows = arrowsB ; id = idB ; _∘_ = _∘B_ ; rightId = rightIdB ; leftId = leftIdB ; associative = associativeB }} {record { objects = objectsC ; arrows = arrowsC ; id = idC ; _∘_ = _∘C_ ; rightId = rightIdC ; leftId = leftIdC ; associative = associativeC }} {record { objects = objectsD ; arrows = arrowsD ; id = idD ; _∘_ = _∘D_ ; rightId = rightIdD ; leftId = leftIdD ; associative = associativeD }} record { onObj = onObjG ; onArrow = onArrowG ; mapId = mapIdG ; mapCompose = mapComposeG } record { onObj = onObjF ; onArrow = onArrowF ; mapId = mapIdF ; mapCompose = mapComposeF } {S} {T} {U} r s rewrite mapComposeF r s | mapComposeG (onArrowF r) (onArrowF s) = refl
idFunctor : {a b : _} (C : Category {a} {b}) Functor C C
Functor.onObj (idFunctor C) = λ x x
Functor.onArrow (idFunctor C) = λ f f
Functor.mapId (idFunctor C) = refl
Functor.mapCompose (idFunctor C) = λ f g refl
typeCastCat : {a b c d : _} {C : Category {a} {b}} {D : Category {c} {d}} (F : Functor C D) (G : Functor C D) (S T : Category.objects C) (pr : Functor.onObj F Functor.onObj G) Category.arrows D (Functor.onObj G S) (Functor.onObj G T) Category.arrows D (Functor.onObj F S) (Functor.onObj F T)
typeCastCat F G S T pr rewrite pr = refl

View File

@@ -0,0 +1,14 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Vectors
open import Semirings.Definition
open import Categories.Definition
module Categories.Dual.Definition where
dual : {a b : _} Category {a} {b} Category {a} {b}
dual record { objects = objects ; arrows = arrows ; id = id ; _∘_ = _∘_ ; rightId = rightId ; leftId = leftId ; compositionAssociative = associative } = record { objects = objects ; arrows = λ i j arrows j i ; id = id ; _∘_ = λ {x y z} g f f g ; rightId = λ {x y} f leftId f ; leftId = λ {x y} f rightId f ; compositionAssociative = λ {x y z w} f g h equalityCommutative (associative h g f) }

View File

@@ -20,11 +20,20 @@ Category.rightId (SET {a}) = λ f → refl
Category.leftId (SET {a}) = λ f refl
Category.compositionAssociative (SET {a}) = λ f g h refl
GROUP : {a : _} Category {lsuc a} {a}
Category.objects (GROUP {a}) = Group {! !} {! !}
Category.arrows (GROUP {a}) = {! !}
Category.id (GROUP {a}) = {! !}
Category._∘_ (GROUP {a}) = {! !}
Category.rightId (GROUP {a}) = {! !}
Category.leftId (GROUP {a}) = {! !}
Category.compositionAssociative (GROUP {a}) = {! !}
GROUP : {a b : _} Category {lsuc a b} {a b}
Category.objects (GROUP {a}) = Group {!!} {!!}
Category.arrows (GROUP {a}) = {!!}
Category.id (GROUP {a}) = {!!}
Category._∘_ (GROUP {a}) = {!!}
Category.rightId (GROUP {a}) = {!!}
Category.leftId (GROUP {a}) = {!!}
Category.compositionAssociative (GROUP {a}) = {!!}
DISCRETE : {a : _} (X : Set a) Category {a} {a}
Category.objects (DISCRETE X) = X
Category.arrows (DISCRETE X) = λ a b a b
Category.id (DISCRETE X) = λ x refl
Category._∘_ (DISCRETE X) = λ y=z x=y transitivity x=y y=z
Category.rightId (DISCRETE X) = {!!}
Category.leftId (DISCRETE X) = {!!}
Category.compositionAssociative (DISCRETE X) = {!!}

View File

@@ -0,0 +1,18 @@
{-# OPTIONS --warning=error --safe --without-K #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Vectors
open import Semirings.Definition
open import Categories.Definition
module Categories.Functor.Definition where
record Functor {a b c d : _} (C : Category {a} {b}) (D : Category {c} {d}) : Set (a b c d) where
field
onObj : Category.objects C Category.objects D
onArrow : {S T : Category.objects C} Category.arrows C S T Category.arrows D (onObj S) (onObj T)
mapId : {T : Category.objects C} onArrow (Category.id C T) Category.id D (onObj T)
mapCompose : {X Y Z : Category.objects C} (f : Category.arrows C X Y) (g : Category.arrows C Y Z) onArrow (Category._∘_ C g f) Category._∘_ D (onArrow g) (onArrow f)

View File

@@ -0,0 +1,26 @@
{-# OPTIONS --warning=error --without-K --safe #-}
open import LogicalFormulae
open import Agda.Primitive using (Level; lzero; lsuc; _⊔_)
open import Numbers.Naturals.Naturals
open import Numbers.Naturals.Order
open import Vectors
open import Semirings.Definition
open import Categories.Definition
open import Categories.Functor.Definition
module Categories.Functor.Lemmas where
functorCompose : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (Functor C D) (Functor B C) (Functor B D)
functorCompose G F = record { onObj = λ x Functor.onObj G (Functor.onObj F x) ; onArrow = λ f Functor.onArrow G (Functor.onArrow F f) ; mapId = λ {T} mapIdHelp G F T ; mapCompose = λ r s mapComposeHelp G F r s }
where
mapIdHelp : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (G : Functor C D) (F : Functor B C) (T : Category.objects B) Functor.onArrow G (Functor.onArrow F (Category.id B T)) Category.id D (Functor.onObj G (Functor.onObj F T))
mapIdHelp {B = B} {C} {D} record { onObj = onObjG ; onArrow = onArrowG ; mapId = mapIdG ; mapCompose = mapComposeG } record { onObj = onObj ; onArrow = onArrow ; mapId = mapId ; mapCompose = mapCompose } T rewrite mapId {T} = mapIdG {onObj T}
mapComposeHelp : {a b c d e f : _} {B : Category {a} {b}} {C : Category {c} {d}} {D : Category {e} {f}} (G : Functor C D) (F : Functor B C) {S T U : Category.objects B} (r : Category.arrows B S T) (s : Category.arrows B T U) (Functor.onArrow G (Functor.onArrow F (Category._∘_ B s r))) (Category._∘_ D (Functor.onArrow G (Functor.onArrow F s)) (Functor.onArrow G (Functor.onArrow F r)))
mapComposeHelp {B = record { objects = objectsB ; arrows = arrowsB ; id = idB ; _∘_ = _∘B_ ; rightId = rightIdB ; leftId = leftIdB ; compositionAssociative = associativeB }} {record { objects = objectsC ; arrows = arrowsC ; id = idC ; _∘_ = _∘C_ ; rightId = rightIdC ; leftId = leftIdC ; compositionAssociative = associativeC }} {record { objects = objectsD ; arrows = arrowsD ; id = idD ; _∘_ = _∘D_ ; rightId = rightIdD ; leftId = leftIdD ; compositionAssociative = associativeD }} record { onObj = onObjG ; onArrow = onArrowG ; mapId = mapIdG ; mapCompose = mapComposeG } record { onObj = onObjF ; onArrow = onArrowF ; mapId = mapIdF ; mapCompose = mapComposeF } {S} {T} {U} r s rewrite mapComposeF r s | mapComposeG (onArrowF r) (onArrowF s) = refl
idFunctor : {a b : _} (C : Category {a} {b}) Functor C C
Functor.onObj (idFunctor C) = λ x x
Functor.onArrow (idFunctor C) = λ f f
Functor.mapId (idFunctor C) = refl
Functor.mapCompose (idFunctor C) = λ f g refl

View File

@@ -75,4 +75,9 @@ open import Semirings.Solver
open import Fields.CauchyCompletion.Group
open import Fields.CauchyCompletion.Ring
open import Categories.Definition
open import Categories.Functor.Definition
open import Categories.Functor.Lemmas
open import Categories.Dual.Definition
module Everything.Safe where