cat/src/Cat/Category.agda

123 lines
4.5 KiB
Agda
Raw Normal View History

2017-11-10 15:00:00 +00:00
{-# OPTIONS --cubical #-}
module Cat.Category where
2017-11-10 15:00:00 +00:00
open import Agda.Primitive
open import Data.Unit.Base
open import Data.Product renaming
( proj₁ to fst
; proj₂ to snd
; ∃! to ∃!≈
)
2017-11-15 21:56:04 +00:00
open import Data.Empty
2017-12-12 11:39:58 +00:00
open import Function
open import Cubical
2017-11-10 15:00:00 +00:00
∃! : {a b} {A : Set a}
(A Set b) Set (a b)
∃! = ∃!≈ _≡_
∃!-syntax : {a b} {A : Set a} (A Set b) Set (a b)
∃!-syntax =
syntax ∃!-syntax (λ x B) = ∃![ x ] B
record IsCategory { ' : Level}
(Object : Set )
(Arrow : Object Object Set ')
(𝟙 : {o : Object} Arrow o o)
(_⊕_ : { a b c : Object } Arrow b c Arrow a b Arrow a c)
: Set (lsuc (' )) where
field
assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D }
h (g f) (h g) f
ident : {A B : Object} {f : Arrow A B}
f 𝟙 f × 𝟙 f f
-- open IsCategory public
record Category ( ' : Level) : Set (lsuc (' )) where
-- adding no-eta-equality can speed up type-checking.
no-eta-equality
2017-11-10 15:00:00 +00:00
field
Object : Set
Arrow : Object Object Set '
𝟙 : {o : Object} Arrow o o
_⊕_ : { a b c : Object } Arrow b c Arrow a b Arrow a c
{{isCategory}} : IsCategory Object Arrow 𝟙 _⊕_
2017-11-10 15:00:00 +00:00
infixl 45 _⊕_
domain : { a b : Object } Arrow a b Object
domain {a = a} _ = a
codomain : { a b : Object } Arrow a b Object
codomain {b = b} _ = b
2017-11-10 15:00:00 +00:00
open Category
2017-11-10 15:00:00 +00:00
module _ { ' : Level} { : Category '} where
module _ { A B : .Object } where
Isomorphism : (f : .Arrow A B) Set '
Isomorphism f = Σ[ g .Arrow B A ] ._⊕_ g f .𝟙 × ._⊕_ f g .𝟙
Epimorphism : {X : .Object } (f : .Arrow A B) Set '
Epimorphism {X} f = ( g₀ g₁ : .Arrow B X ) ._⊕_ g₀ f ._⊕_ g₁ f g₀ g₁
Monomorphism : {X : .Object} (f : .Arrow A B) Set '
Monomorphism {X} f = ( g₀ g₁ : .Arrow X A ) ._⊕_ f g₀ ._⊕_ f g₁ g₀ g₁
-- Isomorphism of objects
_≅_ : (A B : Object ) Set '
_≅_ A B = Σ[ f .Arrow A B ] (Isomorphism f)
module _ { ' : Level} ( : Category ') {A B obj : Object } where
IsProduct : (π₁ : Arrow obj A) (π₂ : Arrow obj B) Set ( ')
IsProduct π₁ π₂
= {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
∃![ x ] ( ._⊕_ π₁ x x₁ × ._⊕_ π₂ x x₂)
-- Tip from Andrea; Consider this style for efficiency:
-- record IsProduct { ' : Level} ( : Category {} {'})
-- {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) : Set (') where
-- field
-- isProduct : ∀ {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
-- → ∃![ x ] ( ._⊕_ π₁ x ≡ x₁ × . _⊕_ π₂ x ≡ x₂)
record Product { ' : Level} { : Category '} (A B : .Object) : Set ( ') where
no-eta-equality
field
obj : .Object
proj₁ : .Arrow obj A
proj₂ : .Arrow obj B
{{isProduct}} : IsProduct proj₁ proj₂
module _ { ' : Level} ( : Category ') where
Opposite : Category '
Opposite =
record
{ Object = .Object
; Arrow = flip ( .Arrow)
; 𝟙 = .𝟙
; _⊕_ = flip ( ._⊕_)
; isCategory = record { assoc = sym assoc ; ident = swap ident }
}
where
open IsCategory ( .isCategory)
2017-11-10 15:00:00 +00:00
-- A consequence of no-eta-equality; `Opposite-is-involution` is no longer
-- definitional - i.e.; you must match on the fields:
--
-- Opposite-is-involution : ∀ { '} → {C : Category {} {'}} → Opposite (Opposite C) ≡ C
-- Object (Opposite-is-involution {C = C} i) = Object C
-- Arrow (Opposite-is-involution i) = {!!}
-- 𝟙 (Opposite-is-involution i) = {!!}
-- _⊕_ (Opposite-is-involution i) = {!!}
-- assoc (Opposite-is-involution i) = {!!}
-- ident (Opposite-is-involution i) = {!!}
Hom : { ' : Level} ( : Category ') (A B : Object ) Set '
2018-01-17 11:10:18 +00:00
Hom A B = Arrow A B
2017-11-10 15:00:00 +00:00
module _ { ' : Level} { : Category '} where
HomFromArrow : (A : .Object) {B B' : .Object} (g : .Arrow B B')
2018-01-17 11:10:18 +00:00
Hom A B Hom A B'
HomFromArrow _A = _⊕_