cat/src/Cat/Category.agda

183 lines
7.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
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
2017-11-10 15:00:00 +00:00
postulate undefined : { : Level} {A : Set } A
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
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
2017-11-10 15:00:00 +00:00
private
open module = Category
_+_ = ._⊕_
Isomorphism : (f : .Arrow A B) Set '
Isomorphism f = Σ[ g .Arrow B A ] g .⊕ f .𝟙 × f + g .𝟙
2017-11-10 15:00:00 +00:00
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
iso-is-epi f (f- , left-inv , right-inv) g₀ g₁ eq =
begin
g₀ ≡⟨ sym (fst .ident)
g₀ + .𝟙 ≡⟨ cong (_+_ g₀) (sym right-inv)
g₀ + (f + f-) ≡⟨ .assoc
(g₀ + f) + f- ≡⟨ cong (λ x x + f-) eq
(g₁ + f) + f- ≡⟨ sym .assoc
g₁ + (f + f-) ≡⟨ cong (_+_ g₁) right-inv
g₁ + .𝟙 ≡⟨ fst .ident
g₁
2017-11-10 15:00:00 +00:00
iso-is-mono : {X} (f : .Arrow A B ) Isomorphism f Monomorphism {X = X} f
iso-is-mono f (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin
g₀ ≡⟨ sym (snd .ident)
.𝟙 + g₀ ≡⟨ cong (λ x x + g₀) (sym left-inv)
(f- + f) + g₀ ≡⟨ sym .assoc
f- + (f + g₀) ≡⟨ cong (_+_ f-) eq
f- + (f + g₁) ≡⟨ .assoc
(f- + f) + g₁ ≡⟨ cong (λ x x + g₁) left-inv
.𝟙 + g₁ ≡⟨ snd .ident
g₁
2017-11-10 15:00:00 +00:00
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
_≅_ : { '} { : Category '} (A B : Object ) Set '
_≅_ { = } A B = Σ[ f .Arrow A B ] (Isomorphism { = } f)
2017-11-10 15:00:00 +00:00
IsProduct : { '} ( : Category ') {A B obj : Object } (π₁ : Arrow obj A) (π₂ : Arrow obj B) Set ( ')
IsProduct {A = A} {B = B} π₁ π₂
= {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
∃![ x ] (π₁ .⊕ x x₁ × π₂ .⊕ x x₂)
2017-11-10 15:00:00 +00:00
where
open module = Category
-- 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₂
mutual
catProduct : {} (C D : Category ) Category
catProduct C D =
record
{ Object = C.Object × D.Object
-- Why does "outlining with `arrowProduct` not work?
; Arrow = λ {(c , d) (c' , d') Arrow C c c' × Arrow D d d'}
; 𝟙 = 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 : {a b} {A : Set a} {B : Set b} {a a' : A} {b b' : B}
a a' b b' (a , b) (a' , b')
eqpair eqa eqb i = eqa i , eqb i
-- arrowProduct : ∀ {} {C D : Category {} {}} → (Object C) × (Object D) → (Object C) × (Object D) → Set
-- arrowProduct = {!!}
-- Arrows in the product-category
arrowProduct : {} {C D : Category } (c d : Object (catProduct C D)) Set
arrowProduct {C = C} {D = D} (c , d) (c' , d') = Arrow C c c' × Arrow D d d'
2017-11-10 15:00:00 +00:00
Opposite : { '} Category ' Category '
2017-11-10 15:00:00 +00:00
Opposite =
record
{ Object = .Object
; Arrow = λ A B .Arrow B A
; 𝟙 = .𝟙
; _⊕_ = λ g f f .⊕ g
; assoc = sym .assoc
; ident = swap .ident
}
where
open module = Category
-- 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 = _⊕_