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-05 10:43:38 +00:00
|
|
|
|
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
|
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-02-05 10:43:38 +00:00
|
|
|
|
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)
|
2018-01-21 13:31:37 +00:00
|
|
|
|
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)
|
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
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 ≡ 𝟙
|
2018-02-01 14:20:20 +00:00
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
_≅_ : (A B : Object) → Set ℓb
|
2018-02-01 14:20:20 +00:00
|
|
|
|
_≅_ 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
|
2018-02-05 10:43:38 +00:00
|
|
|
|
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₁
|
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
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₁
|
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
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 _ _ _)
|
2018-02-01 14:20:20 +00:00
|
|
|
|
-- This lemma will be useful to prove the equality of two categories.
|
2018-02-05 10:43:38 +00:00
|
|
|
|
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
|
2018-01-21 13:31:37 +00:00
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
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 ℂ
|
|
|
|
|
isCategory = snd ℂ
|
|
|
|
|
|
2018-02-05 11:21:39 +00:00
|
|
|
|
private
|
|
|
|
|
module ℂ = RawCategory raw
|
|
|
|
|
|
|
|
|
|
Object : Set ℓa
|
|
|
|
|
Object = ℂ.Object
|
|
|
|
|
|
|
|
|
|
Arrow = ℂ.Arrow
|
|
|
|
|
|
|
|
|
|
𝟙 = ℂ.𝟙
|
|
|
|
|
|
|
|
|
|
_∘_ = ℂ._∘_
|
2018-02-05 10:43:38 +00:00
|
|
|
|
|
2018-02-05 11:21:39 +00:00
|
|
|
|
_[_,_] : (A : Object) → (B : Object) → Set ℓb
|
|
|
|
|
_[_,_] = ℂ.Arrow
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
2018-02-05 11:21:39 +00:00
|
|
|
|
_[_∘_] : {A B C : Object} → (g : ℂ.Arrow B C) → (f : ℂ.Arrow A B) → ℂ.Arrow A C
|
|
|
|
|
_[_∘_] = ℂ._∘_
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
2018-02-05 11:21:39 +00:00
|
|
|
|
open Category using ( Object ; _[_,_] ; _[_∘_])
|
2018-01-30 17:26:11 +00:00
|
|
|
|
|
2018-02-05 10:43:38 +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 = {!!}
|
2018-01-20 23:21:25 +00:00
|
|
|
|
}
|
2018-02-05 10:43:38 +00:00
|
|
|
|
Opposite : Category ℓa ℓb
|
|
|
|
|
Opposite = OpRaw , OpIsCategory
|
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-02-02 13:47:33 +00:00
|
|
|
|
module _ {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where
|
|
|
|
|
unique = isContr
|
|
|
|
|
|
2018-02-05 11:21:39 +00:00
|
|
|
|
IsInitial : Object ℂ → Set (ℓa ⊔ ℓb)
|
|
|
|
|
IsInitial I = {X : Object ℂ} → unique (ℂ [ I , X ])
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
2018-02-05 11:21:39 +00:00
|
|
|
|
IsTerminal : Object ℂ → Set (ℓa ⊔ ℓb)
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- ∃![ ? ] ?
|
2018-02-05 11:21:39 +00:00
|
|
|
|
IsTerminal T = {X : Object ℂ} → unique (ℂ [ X , T ])
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
|
|
|
|
Initial : Set (ℓa ⊔ ℓb)
|
2018-02-05 11:21:39 +00:00
|
|
|
|
Initial = Σ (Object ℂ) IsInitial
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
|
|
|
|
Terminal : Set (ℓa ⊔ ℓb)
|
2018-02-05 11:21:39 +00:00
|
|
|
|
Terminal = Σ (Object ℂ) IsTerminal
|