cat/src/Cat/Category.agda

257 lines
9.4 KiB
Agda
Raw Normal View History

2018-02-02 14:33:54 +00:00
{-# OPTIONS --allow-unsolved-metas --cubical #-}
2017-11-10 15:00:00 +00:00
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
import Function
2017-12-12 11:39:58 +00:00
open import Cubical
2018-02-05 09:24:57 +00:00
open import Cubical.GradLemma using ( propIsEquiv )
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 RawCategory ( ' : Level) : Set (lsuc (' )) where
-- adding no-eta-equality can speed up type-checking.
-- ONLY IF you define your categories with copatterns though.
no-eta-equality
field
-- Need something like:
-- Object : Σ (Set ) isGroupoid
Object : Set
-- And:
-- Arrow : Object → Object → Σ (Set ') isSet
Arrow : Object Object Set '
𝟙 : {o : Object} Arrow o o
_∘_ : {A B C : Object} Arrow B C Arrow A B Arrow A C
infixl 10 _∘_
domain : { a b : Object } Arrow a b Object
domain {a = a} _ = a
codomain : { a b : Object } Arrow a b Object
codomain {b = b} _ = b
2018-02-02 14:33:54 +00:00
-- Thierry: All projections must be `isProp`'s
-- According to definitions 9.1.1 and 9.1.6 in the HoTT book the
-- arrows of a category form a set (arrow-is-set), and there is an
-- equivalence between the equality of objects and isomorphisms
-- (univalent).
record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where
open RawCategory
-- (Object : Set )
-- (Arrow : Object → Object → Set ')
-- (𝟙 : {o : Object} → Arrow o o)
-- (_∘_ : { a b c : Object } → Arrow b c → Arrow a b → Arrow a c)
field
assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D }
2018-02-02 14:33:54 +00:00
h (g f) (h g) f
ident : {A B : Object} {f : Arrow A B}
2018-02-02 14:33:54 +00:00
f 𝟙 f × 𝟙 f f
arrow-is-set : {A B : Object} isSet (Arrow A B)
Isomorphism : {A B} (f : Arrow A B) Set b
2018-02-02 14:33:54 +00:00
Isomorphism {A} {B} f = Σ[ g Arrow B A ] g f 𝟙 × f g 𝟙
_≅_ : (A B : Object) Set b
_≅_ A B = Σ[ f Arrow A B ] (Isomorphism f)
idIso : (A : Object) A A
idIso A = 𝟙 , (𝟙 , ident)
id-to-iso : (A B : Object) A B A B
id-to-iso A B eq = transp (\ i A eq i) (idIso A)
-- TODO: might want to implement isEquiv differently, there are 3
-- equivalent formulations in the book.
field
univalent : {A B : Object} isEquiv (A B) (A B) (id-to-iso A B)
2018-02-02 14:33:54 +00:00
module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set b
2018-02-02 14:33:54 +00:00
Epimorphism {X} f = ( g₀ g₁ : Arrow B X ) g₀ f g₁ f g₀ g₁
Monomorphism : {X : Object} (f : Arrow A B) Set b
2018-02-02 14:33:54 +00:00
Monomorphism {X} f = ( g₀ g₁ : Arrow X A ) f g₀ f g₁ g₀ g₁
module _ {a} {b} { : RawCategory a b} where
2018-02-05 09:24:57 +00:00
-- TODO, provable by using arrow-is-set and that isProp (isEquiv _ _ _)
-- This lemma will be useful to prove the equality of two categories.
IsCategory-is-prop : isProp (IsCategory )
2018-02-05 09:24:57 +00:00
IsCategory-is-prop x y i = record
{ assoc = x.arrow-is-set _ _ x.assoc y.assoc i
; ident =
( x.arrow-is-set _ _ (fst x.ident) (fst y.ident) i
, x.arrow-is-set _ _ (snd x.ident) (snd y.ident) i
)
-- ; arrow-is-set = {!λ x₁ y₁ p q → x.arrow-is-set _ _ p q!}
; arrow-is-set = λ _ _ p q
let
golden : x.arrow-is-set _ _ p q y.arrow-is-set _ _ p q
golden = λ j k l {!!}
in
golden i
; univalent = λ y₁ {!!}
}
where
module x = IsCategory x
module y = IsCategory y
Category : (a b : Level) Set (lsuc (a b))
Category a b = Σ (RawCategory a b) IsCategory
module Category {a b : Level} ( : Category a b) where
raw = fst
open RawCategory raw public
isCategory = snd
open RawCategory
-- _∈_ : ∀ {a b} ( : Category a b) → ( .fst .Object → Set b) → Set (a ⊔ b)
-- A ∈ =
2017-11-10 15:00:00 +00:00
Obj : {a b} Category a b Set a
Obj = .fst .Object
2017-11-10 15:00:00 +00:00
_[_,_] : { '} ( : Category ') (A : Obj ) (B : Obj ) Set '
[ A , B ] = .fst .Arrow A B
_[_∘_] : { '} ( : Category ') {A B C : Obj } (g : [ B , C ]) (f : [ A , B ]) [ A , C ]
[ g f ] = .fst ._∘_ g f
module _ { ' : Level} ( : Category ') {A B obj : Obj } where
IsProduct : (π₁ : [ obj , A ]) (π₂ : [ obj , B ]) Set ( ')
IsProduct π₁ π₂
= {X : Obj } (x₁ : [ X , A ]) (x₂ : [ 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 : Obj ) : Set ( ') where
no-eta-equality
field
obj : Obj
proj₁ : [ obj , A ]
proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂
2018-01-24 15:38:28 +00:00
arrowProduct : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
[ X , obj ]
2018-01-21 20:29:15 +00:00
arrowProduct π₁ π₂ = fst (isProduct π₁ π₂)
record HasProducts { ' : Level} ( : Category ') : Set ( ') where
field
product : (A B : Obj ) Product { = } A B
2018-01-24 15:38:28 +00:00
open Product
objectProduct : (A B : Obj ) Obj
2018-01-24 15:38:28 +00:00
objectProduct A B = Product.obj (product A B)
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
-- It's a "parallel" product
parallelProduct : {A A' B B' : Obj } [ A , A' ] [ B , B' ]
[ objectProduct A B , objectProduct A' B' ]
2018-01-24 15:38:28 +00:00
parallelProduct {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B')
( [ a (product A B) .proj₁ ])
( [ b (product A B) .proj₂ ])
2018-01-24 15:38:28 +00:00
module _ {a b : Level} ( : Category a b) where
private
open Category
module = RawCategory ( .fst)
OpRaw : RawCategory a b
OpRaw = record
{ Object = .Object
; Arrow = Function.flip .Arrow
; 𝟙 = .𝟙
; _∘_ = Function.flip (._∘_)
}
open IsCategory isCategory
OpIsCategory : IsCategory OpRaw
OpIsCategory = record
{ assoc = sym assoc
; ident = swap ident
; arrow-is-set = {!!}
; univalent = {!!}
}
Opposite : Category a b
Opposite = OpRaw , OpIsCategory
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) = {!!}
2018-01-25 11:01:37 +00:00
module _ { '} ( : Category ') {{hasProducts : HasProducts }} where
open HasProducts hasProducts
2018-01-24 15:38:28 +00:00
open Product hiding (obj)
private
_×p_ : (A B : Obj ) Obj
2018-01-24 15:38:28 +00:00
_×p_ A B = Product.obj (product A B)
module _ (B C : Obj ) where
IsExponential : (Cᴮ : Obj ) [ Cᴮ ×p B , C ] Set ( ')
IsExponential Cᴮ eval = (A : Obj ) (f : [ A ×p B , C ])
∃![ f~ ] ( [ eval parallelProduct f~ (Category.raw .𝟙)] f)
2018-01-24 15:38:28 +00:00
record Exponential : Set ( ') where
field
-- obj ≡ Cᴮ
obj : Obj
eval : [ obj ×p B , C ]
2018-01-24 15:38:28 +00:00
{{isExponential}} : IsExponential obj eval
-- If I make this an instance-argument then the instance resolution
-- algorithm goes into an infinite loop. Why?
2018-01-25 11:01:37 +00:00
exponentialsHaveProducts : HasProducts
exponentialsHaveProducts = hasProducts
transpose : (A : Obj ) [ A ×p B , C ] [ A , obj ]
2018-01-24 15:38:28 +00:00
transpose A f = fst (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
field
exponent : (A B : Obj ) Exponential A B
2018-01-24 15:38:28 +00:00
record CartesianClosed { ' : Level} ( : Category ') : Set ( ') where
field
{{hasProducts}} : HasProducts
{{hasExponentials}} : HasExponentials
module _ {a b : Level} ( : Category a b) where
unique = isContr
IsInitial : Obj Set (a b)
IsInitial I = {X : Obj } unique ( [ I , X ])
IsTerminal : Obj Set (a b)
-- ∃![ ? ] ?
IsTerminal T = {X : Obj } unique ( [ X , T ])
Initial : Set (a b)
Initial = Σ (Obj ) IsInitial
Terminal : Set (a b)
Terminal = Σ (Obj ) IsTerminal