cat/src/Cat/Category.agda

143 lines
5.1 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
2017-12-12 11:39:58 +00:00
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
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
postulate undefined : { : Level} {A : Set } A
record Category { '} : Set (lsuc (' )) where
constructor category
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
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
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 public
module _ { ' : Level} { : Category {} {'}} { A B : Object } where
private
open module = Category
_+_ = ._⊕_
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₁
iso-is-epi : {X} (f : .Arrow A B) Isomorphism f Epimorphism {X = X} f
-- Idea: Pre-compose with f- on both sides of the equality of eq to get
-- g₀ + f + f- ≡ g₁ + f + f-
-- which by left-inv reduces to the goal.
iso-is-epi f (f- , left-inv , right-inv) g₀ g₁ eq =
trans (sym (fst .ident))
( trans (cong (_+_ g₀) (sym right-inv))
( trans .assoc
( trans (cong (λ x x + f-) eq)
( trans (sym .assoc)
( trans (cong (_+_ g₁) right-inv) (fst .ident))
)
)
)
)
iso-is-mono : {X} (f : .Arrow A B ) Isomorphism f Monomorphism {X = X} f
-- For the next goal we do something similar: Post-compose with f- and use
-- right-inv to get the goal.
iso-is-mono f (f- , (left-inv , right-inv)) g₀ g₁ eq =
trans (sym (snd .ident))
( trans (cong (λ x x + g₀) (sym left-inv))
( trans (sym .assoc)
( trans (cong (_+_ f-) eq)
( trans .assoc
( trans (cong (λ x x + g₁) left-inv) (snd .ident)
)
)
)
)
)
iso-is-epi-mono : {X} (f : .Arrow A B ) Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
iso-is-epi-mono f iso = iso-is-epi f iso , iso-is-mono f iso
{-
epi-mono-is-not-iso : { '} ¬ (( : Category {} {'}) {A B X : Object } (f : Arrow A B ) Epimorphism { = } {X = X} f Monomorphism { = } {X = X} f Isomorphism { = } f)
epi-mono-is-not-iso f =
let k = f {!!} {!!} {!!} {!!}
in {!!}
-}
-- Isomorphism of objects
2017-11-10 15:00:00 +00:00
_≅_ : { ' : Level } { : Category {} {'} } ( A B : Object ) Set '
_≅_ { = } A B = Σ[ f .Arrow A B ] (Isomorphism { = } f)
where
open module = Category
Product : { : Level} ( C D : Category {} {} ) Category {} {}
Product C D =
record
{ Object = C.Object × D.Object
; Arrow = λ { (c , d) (c' , d')
let carr = C.Arrow c c'
darr = D.Arrow d d'
in carr × darr}
; 𝟙 = C.𝟙 , D.𝟙
; _⊕_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) bc∈C C.⊕ ab∈C , bc∈D D.⊕ ab∈D}
; assoc = eqpair C.assoc D.assoc
; ident =
let (Cl , Cr) = C.ident
(Dl , Dr) = D.ident
in eqpair Cl Dl , eqpair Cr Dr
}
where
open module C = Category C
open module D = Category D
-- Two pairs are equal if their components are equal.
eqpair : { : Level} { A : Set } { B : Set } { a a' : A } { b b' : B } a a' b b' (a , b) (a' , b')
eqpair {a = a} {b = b} eqa eqb = subst eqa (subst eqb (refl {x = (a , b)}))
Opposite : { '} Category {} {'} Category {} {'}
Opposite =
record
{ Object = .Object
; Arrow = λ A B .Arrow B A
; 𝟙 = .𝟙
; _⊕_ = λ g f f .⊕ g
; assoc = sym .assoc
; ident = swap .ident
}
where
open module = Category
2018-01-17 11:10:18 +00:00
Hom : { ' : Level} ( : Category {} {'}) (A B : Object ) Set '
Hom A B = Arrow A B
2017-11-10 15:00:00 +00:00
module _ { ' : Level} { : Category {} {'}} where
private
Obj = Object
Arr = Arrow
_+_ = _⊕_
HomFromArrow : (A : Obj) {B B' : Obj} (g : Arr B B')
2018-01-17 11:10:18 +00:00
Hom A B Hom A B'
2017-11-10 15:00:00 +00:00
HomFromArrow _A g = λ f g + f