2018-02-02 14:33:54 +00:00
|
|
|
|
{-# OPTIONS --allow-unsolved-metas --cubical #-}
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
2018-01-08 21:48:59 +00:00
|
|
|
|
module Cat.Category where
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
|
|
|
|
open import Agda.Primitive
|
|
|
|
|
open import Data.Unit.Base
|
2018-01-20 23:21:25 +00:00
|
|
|
|
open import Data.Product renaming
|
|
|
|
|
( proj₁ to fst
|
|
|
|
|
; proj₂ to snd
|
|
|
|
|
; ∃! to ∃!≈
|
|
|
|
|
)
|
2017-11-15 21:56:04 +00:00
|
|
|
|
open import Data.Empty
|
2018-01-30 18:19:16 +00:00
|
|
|
|
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
|
|
|
|
|
2018-01-20 23:21:25 +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
|
|
|
|
|
|
2018-02-02 14:33:54 +00:00
|
|
|
|
-- Thierry: All projections must be `isProp`'s
|
2018-02-01 14:20:20 +00:00
|
|
|
|
|
|
|
|
|
-- 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).
|
2018-01-21 13:31:37 +00:00
|
|
|
|
record IsCategory {ℓ ℓ' : Level}
|
|
|
|
|
(Object : Set ℓ)
|
|
|
|
|
(Arrow : Object → Object → Set ℓ')
|
|
|
|
|
(𝟙 : {o : Object} → Arrow o o)
|
2018-02-02 14:33:54 +00:00
|
|
|
|
(_∘_ : { a b c : Object } → Arrow b c → Arrow a b → Arrow a c)
|
2018-01-21 13:31:37 +00:00
|
|
|
|
: Set (lsuc (ℓ' ⊔ ℓ)) where
|
|
|
|
|
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
|
2018-01-21 13:31:37 +00:00
|
|
|
|
ident : {A B : Object} {f : Arrow A B}
|
2018-02-02 14:33:54 +00:00
|
|
|
|
→ f ∘ 𝟙 ≡ f × 𝟙 ∘ f ≡ f
|
2018-02-01 14:20:20 +00:00
|
|
|
|
arrow-is-set : ∀ {A B : Object} → isSet (Arrow A B)
|
|
|
|
|
|
|
|
|
|
Isomorphism : ∀ {A B} → (f : Arrow A B) → Set ℓ'
|
2018-02-02 14:33:54 +00:00
|
|
|
|
Isomorphism {A} {B} f = Σ[ g ∈ Arrow B A ] g ∘ f ≡ 𝟙 × f ∘ g ≡ 𝟙
|
2018-02-01 14:20:20 +00:00
|
|
|
|
|
|
|
|
|
_≅_ : (A B : Object) → Set ℓ'
|
|
|
|
|
_≅_ 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 ℓ'
|
|
|
|
|
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₁
|
|
|
|
|
|
2018-02-05 09:24:57 +00:00
|
|
|
|
module _ {ℓ} {ℓ'} {Object : Set ℓ}
|
2018-02-01 14:20:20 +00:00
|
|
|
|
{Arrow : Object → Object → Set ℓ'}
|
|
|
|
|
{𝟙 : {o : Object} → Arrow o o}
|
|
|
|
|
{_⊕_ : { a b c : Object } → Arrow b c → Arrow a b → Arrow a c}
|
|
|
|
|
where
|
2018-02-05 09:24:57 +00:00
|
|
|
|
-- TODO, provable by using arrow-is-set and that isProp (isEquiv _ _ _)
|
2018-02-01 14:20:20 +00:00
|
|
|
|
-- This lemma will be useful to prove the equality of two categories.
|
|
|
|
|
IsCategory-is-prop : isProp (IsCategory Object Arrow 𝟙 _⊕_)
|
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
|
2018-01-21 13:31:37 +00:00
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
record Category (ℓ ℓ' : Level) : Set (lsuc (ℓ' ⊔ ℓ)) where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- adding no-eta-equality can speed up type-checking.
|
2018-02-01 14:20:20 +00:00
|
|
|
|
-- ONLY IF you define your categories with copatterns though.
|
2018-01-20 23:21:25 +00:00
|
|
|
|
no-eta-equality
|
2017-11-10 15:00:00 +00:00
|
|
|
|
field
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- Need something like:
|
|
|
|
|
-- Object : Σ (Set ℓ) isGroupoid
|
2017-11-10 15:00:00 +00:00
|
|
|
|
Object : Set ℓ
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- And:
|
|
|
|
|
-- Arrow : Object → Object → Σ (Set ℓ') isSet
|
2017-11-10 15:00:00 +00:00
|
|
|
|
Arrow : Object → Object → Set ℓ'
|
|
|
|
|
𝟙 : {o : Object} → Arrow o o
|
2018-01-30 18:19:16 +00:00
|
|
|
|
_∘_ : {A B C : Object} → Arrow B C → Arrow A B → Arrow A C
|
|
|
|
|
{{isCategory}} : IsCategory Object Arrow 𝟙 _∘_
|
|
|
|
|
infixl 10 _∘_
|
2017-12-02 00:36:16 +00:00
|
|
|
|
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
|
|
|
|
|
2018-01-21 13:31:37 +00:00
|
|
|
|
open Category
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
2018-01-30 17:26:11 +00:00
|
|
|
|
_[_,_] : ∀ {ℓ ℓ'} → (ℂ : Category ℓ ℓ') → (A : ℂ .Object) → (B : ℂ .Object) → Set ℓ'
|
|
|
|
|
_[_,_] = Arrow
|
|
|
|
|
|
|
|
|
|
_[_∘_] : ∀ {ℓ ℓ'} → (ℂ : Category ℓ ℓ') → {A B C : ℂ .Object} → (g : ℂ [ B , C ]) → (f : ℂ [ A , B ]) → ℂ [ A , C ]
|
2018-01-30 18:19:16 +00:00
|
|
|
|
_[_∘_] = _∘_
|
2018-01-30 17:26:11 +00:00
|
|
|
|
|
2018-01-21 14:01:01 +00:00
|
|
|
|
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)
|
2018-01-30 18:19:16 +00:00
|
|
|
|
→ ∃![ x ] (ℂ [ π₁ ∘ x ] ≡ x₁ × ℂ [ π₂ ∘ x ] ≡ x₂)
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-21 00:03:40 +00:00
|
|
|
|
-- Tip from Andrea; Consider this style for efficiency:
|
|
|
|
|
-- record IsProduct {ℓ ℓ' : Level} (ℂ : Category {ℓ} {ℓ'})
|
|
|
|
|
-- {A B obj : Object ℂ} (π₁ : Arrow ℂ obj A) (π₂ : Arrow ℂ obj B) : Set (ℓ ⊔ ℓ') where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- field
|
2018-01-21 00:03:40 +00:00
|
|
|
|
-- isProduct : ∀ {X : ℂ .Object} (x₁ : ℂ .Arrow X A) (x₂ : ℂ .Arrow X B)
|
|
|
|
|
-- → ∃![ x ] (ℂ ._⊕_ π₁ x ≡ x₁ × ℂ. _⊕_ π₂ x ≡ x₂)
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
record Product {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} (A B : ℂ .Object) : Set (ℓ ⊔ ℓ') where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
no-eta-equality
|
|
|
|
|
field
|
2018-01-21 00:03:40 +00:00
|
|
|
|
obj : ℂ .Object
|
|
|
|
|
proj₁ : ℂ .Arrow obj A
|
|
|
|
|
proj₂ : ℂ .Arrow obj B
|
2018-01-20 23:21:25 +00:00
|
|
|
|
{{isProduct}} : IsProduct ℂ proj₁ proj₂
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-01-21 20:29:15 +00:00
|
|
|
|
arrowProduct : ∀ {X} → (π₁ : Arrow ℂ X A) (π₂ : Arrow ℂ X B)
|
|
|
|
|
→ Arrow ℂ X obj
|
|
|
|
|
arrowProduct π₁ π₂ = fst (isProduct π₁ π₂)
|
|
|
|
|
|
|
|
|
|
record HasProducts {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') : Set (ℓ ⊔ ℓ') where
|
|
|
|
|
field
|
|
|
|
|
product : ∀ (A B : ℂ .Object) → Product {ℂ = ℂ} A B
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-24 15:38:28 +00:00
|
|
|
|
open Product
|
|
|
|
|
|
|
|
|
|
objectProduct : (A B : ℂ .Object) → ℂ .Object
|
|
|
|
|
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' : ℂ .Object} → ℂ .Arrow A A' → ℂ .Arrow B B'
|
|
|
|
|
→ ℂ .Arrow (objectProduct A B) (objectProduct A' B')
|
|
|
|
|
parallelProduct {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B')
|
2018-01-30 18:19:16 +00:00
|
|
|
|
(ℂ [ a ∘ (product A B) .proj₁ ])
|
|
|
|
|
(ℂ [ b ∘ (product A B) .proj₂ ])
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
2018-01-21 13:31:37 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') where
|
|
|
|
|
Opposite : Category ℓ ℓ'
|
|
|
|
|
Opposite =
|
2018-01-20 23:21:25 +00:00
|
|
|
|
record
|
2018-01-21 13:31:37 +00:00
|
|
|
|
{ Object = ℂ .Object
|
2018-01-30 18:19:16 +00:00
|
|
|
|
; Arrow = Function.flip (ℂ .Arrow)
|
2018-01-21 13:31:37 +00:00
|
|
|
|
; 𝟙 = ℂ .𝟙
|
2018-01-30 18:19:16 +00:00
|
|
|
|
; _∘_ = Function.flip (ℂ ._∘_)
|
2018-02-01 14:20:20 +00:00
|
|
|
|
; isCategory = record { assoc = sym assoc ; ident = swap ident
|
|
|
|
|
; arrow-is-set = {!!}
|
|
|
|
|
; univalent = {!!} }
|
2018-01-20 23:21:25 +00:00
|
|
|
|
}
|
2018-01-21 14:01:01 +00:00
|
|
|
|
where
|
|
|
|
|
open IsCategory (ℂ .isCategory)
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
2018-01-20 23:21:25 +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-21 00:11:08 +00:00
|
|
|
|
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
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
HomFromArrow : (A : ℂ .Object) → {B B' : ℂ .Object} → (g : ℂ .Arrow B B')
|
2018-01-17 11:10:18 +00:00
|
|
|
|
→ Hom ℂ A B → Hom ℂ A B'
|
2018-01-30 18:19:16 +00:00
|
|
|
|
HomFromArrow _A = ℂ ._∘_
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
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 : ℂ .Object) → ℂ .Object
|
|
|
|
|
_×p_ A B = Product.obj (product A B)
|
|
|
|
|
|
|
|
|
|
module _ (B C : ℂ .Category.Object) where
|
|
|
|
|
IsExponential : (Cᴮ : ℂ .Object) → ℂ .Arrow (Cᴮ ×p B) C → Set (ℓ ⊔ ℓ')
|
|
|
|
|
IsExponential Cᴮ eval = ∀ (A : ℂ .Object) (f : ℂ .Arrow (A ×p B) C)
|
2018-01-30 18:19:16 +00:00
|
|
|
|
→ ∃![ f~ ] (ℂ [ eval ∘ parallelProduct f~ (ℂ .𝟙)] ≡ f)
|
2018-01-24 15:38:28 +00:00
|
|
|
|
|
|
|
|
|
record Exponential : Set (ℓ ⊔ ℓ') where
|
|
|
|
|
field
|
|
|
|
|
-- obj ≡ Cᴮ
|
|
|
|
|
obj : ℂ .Object
|
|
|
|
|
eval : ℂ .Arrow ( obj ×p B ) C
|
|
|
|
|
{{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
|
2018-01-24 15:38:28 +00:00
|
|
|
|
transpose : (A : ℂ .Object) → ℂ .Arrow (A ×p B) C → ℂ .Arrow A obj
|
|
|
|
|
transpose A f = fst (isExponential A f)
|
|
|
|
|
|
|
|
|
|
record HasExponentials {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') {{_ : HasProducts ℂ}} : Set (ℓ ⊔ ℓ') where
|
|
|
|
|
field
|
|
|
|
|
exponent : (A B : ℂ .Object) → Exponential ℂ A B
|
|
|
|
|
|
|
|
|
|
record CartesianClosed {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') : Set (ℓ ⊔ ℓ') where
|
|
|
|
|
field
|
|
|
|
|
{{hasProducts}} : HasProducts ℂ
|
|
|
|
|
{{hasExponentials}} : HasExponentials ℂ
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
|
|
|
|
module _ {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where
|
|
|
|
|
unique = isContr
|
|
|
|
|
|
|
|
|
|
IsInitial : ℂ .Object → Set (ℓa ⊔ ℓb)
|
|
|
|
|
IsInitial I = {X : ℂ .Object} → unique (ℂ .Arrow I X)
|
|
|
|
|
|
|
|
|
|
IsTerminal : ℂ .Object → Set (ℓa ⊔ ℓb)
|
|
|
|
|
-- ∃![ ? ] ?
|
|
|
|
|
IsTerminal T = {X : ℂ .Object} → unique (ℂ .Arrow X T)
|
|
|
|
|
|
|
|
|
|
Initial : Set (ℓa ⊔ ℓb)
|
|
|
|
|
Initial = Σ (ℂ .Object) IsInitial
|
|
|
|
|
|
|
|
|
|
Terminal : Set (ℓa ⊔ ℓb)
|
|
|
|
|
Terminal = Σ (ℂ .Object) IsTerminal
|