2018-01-17 22:00:27 +00:00
|
|
|
|
{-# OPTIONS --cubical --allow-unsolved-metas #-}
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-17 22:00:27 +00:00
|
|
|
|
module Cat.Categories.Cat where
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
|
|
|
|
open import Agda.Primitive
|
|
|
|
|
open import Cubical
|
|
|
|
|
open import Function
|
|
|
|
|
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
|
|
|
|
|
|
2018-01-17 22:00:27 +00:00
|
|
|
|
open import Cat.Category
|
|
|
|
|
open import Cat.Functor
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-21 14:03:00 +00:00
|
|
|
|
-- Tip from Andrea:
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- Use co-patterns - they help with showing more understandable types in goals.
|
|
|
|
|
lift-eq : ∀ {ℓ} {A B : Set ℓ} {a a' : A} {b b' : B} → a ≡ a' → b ≡ b' → (a , b) ≡ (a' , b')
|
|
|
|
|
fst (lift-eq a b i) = a i
|
|
|
|
|
snd (lift-eq a b i) = b i
|
|
|
|
|
|
2018-01-21 14:03:00 +00:00
|
|
|
|
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
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
|
|
|
|
open Functor
|
|
|
|
|
open Category
|
2018-01-21 00:11:08 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} {A B : Category ℓ ℓ'} where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
lift-eq-functors : {f g : Functor A B}
|
2018-01-21 18:23:24 +00:00
|
|
|
|
→ (eq* : f .func* ≡ g .func*)
|
|
|
|
|
→ (eq→ : PathP (λ i → ∀ {x y} → A .Arrow x y → B .Arrow (eq* i x) (eq* i y))
|
|
|
|
|
(f .func→) (g .func→))
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- → (eq→ : Functor.func→ f ≡ {!!}) -- Functor.func→ g)
|
|
|
|
|
-- Use PathP
|
|
|
|
|
-- directly to show heterogeneous equalities by using previous
|
|
|
|
|
-- equalities (i.e. continuous paths) to create new continuous paths.
|
|
|
|
|
→ (eqI : PathP (λ i → ∀ {c : A .Object} → eq→ i (A .𝟙 {c}) ≡ B .𝟙 {eq* i c})
|
|
|
|
|
(ident f) (ident g))
|
|
|
|
|
→ (eqD : PathP (λ i → { c c' c'' : A .Object} {a : A .Arrow c c'} {a' : A .Arrow c' c''}
|
2018-01-21 18:23:24 +00:00
|
|
|
|
→ eq→ i (A ._⊕_ a' a) ≡ B ._⊕_ (eq→ i a') (eq→ i a))
|
|
|
|
|
(distrib f) (distrib g))
|
2018-01-20 23:21:25 +00:00
|
|
|
|
→ f ≡ g
|
|
|
|
|
lift-eq-functors eq* eq→ eqI eqD i = record { func* = eq* i ; func→ = eq→ i ; ident = eqI i ; distrib = eqD i }
|
|
|
|
|
|
2018-01-08 21:29:29 +00:00
|
|
|
|
-- The category of categories
|
|
|
|
|
module _ {ℓ ℓ' : Level} where
|
|
|
|
|
private
|
2018-01-21 00:11:08 +00:00
|
|
|
|
module _ {A B C D : Category ℓ ℓ'} {f : Functor A B} {g : Functor B C} {h : Functor C D} where
|
2018-01-21 18:23:24 +00:00
|
|
|
|
eq* : func* (h ∘f (g ∘f f)) ≡ func* ((h ∘f g) ∘f f)
|
|
|
|
|
eq* = refl
|
|
|
|
|
eq→ : PathP
|
|
|
|
|
(λ i → {x y : A .Object} → A .Arrow x y → D .Arrow (eq* i x) (eq* i y))
|
|
|
|
|
(func→ (h ∘f (g ∘f f))) (func→ ((h ∘f g) ∘f f))
|
|
|
|
|
eq→ = refl
|
|
|
|
|
id-l = (h ∘f (g ∘f f)) .ident -- = func→ (h ∘f (g ∘f f)) (𝟙 A) ≡ 𝟙 D
|
|
|
|
|
id-r = ((h ∘f g) ∘f f) .ident -- = func→ ((h ∘f g) ∘f f) (𝟙 A) ≡ 𝟙 D
|
|
|
|
|
postulate eqI : PathP
|
|
|
|
|
(λ i → ∀ {c : A .Object} → eq→ i (A .𝟙 {c}) ≡ D .𝟙 {eq* i c})
|
|
|
|
|
(ident ((h ∘f (g ∘f f))))
|
|
|
|
|
(ident ((h ∘f g) ∘f f))
|
|
|
|
|
postulate eqD : PathP (λ i → { c c' c'' : A .Object} {a : A .Arrow c c'} {a' : A .Arrow c' c''}
|
|
|
|
|
→ eq→ i (A ._⊕_ a' a) ≡ D ._⊕_ (eq→ i a') (eq→ i a))
|
|
|
|
|
(distrib (h ∘f (g ∘f f))) (distrib ((h ∘f g) ∘f f))
|
|
|
|
|
-- eqD = {!!}
|
|
|
|
|
|
|
|
|
|
assc : h ∘f (g ∘f f) ≡ (h ∘f g) ∘f f
|
|
|
|
|
assc = lift-eq-functors eq* eq→ eqI eqD
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
module _ {A B : Category ℓ ℓ'} {f : Functor A B} where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
lem : (func* f) ∘ (func* (identity {C = A})) ≡ func* f
|
2018-01-08 21:29:29 +00:00
|
|
|
|
lem = refl
|
2018-01-21 14:21:50 +00:00
|
|
|
|
-- lemmm : func→ {C = A} {D = B} (f ∘f identity) ≡ func→ f
|
2018-01-20 23:21:25 +00:00
|
|
|
|
lemmm : PathP
|
|
|
|
|
(λ i →
|
|
|
|
|
{x y : Object A} → Arrow A x y → Arrow B (func* f x) (func* f y))
|
2018-01-21 14:21:50 +00:00
|
|
|
|
(func→ (f ∘f identity)) (func→ f)
|
2018-01-20 23:21:25 +00:00
|
|
|
|
lemmm = refl
|
|
|
|
|
postulate lemz : PathP (λ i → {c : A .Object} → PathP (λ _ → Arrow B (func* f c) (func* f c)) (func→ f (A .𝟙)) (B .𝟙))
|
2018-01-21 14:21:50 +00:00
|
|
|
|
(ident (f ∘f identity)) (ident f)
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- lemz = {!!}
|
2018-01-21 14:21:50 +00:00
|
|
|
|
postulate ident-r : f ∘f identity ≡ f
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- ident-r = lift-eq-functors lem lemmm {!lemz!} {!!}
|
2018-01-21 14:21:50 +00:00
|
|
|
|
postulate ident-l : identity ∘f f ≡ f
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- ident-l = lift-eq-functors lem lemmm {!refl!} {!!}
|
2018-01-08 21:29:29 +00:00
|
|
|
|
|
2018-01-21 14:23:40 +00:00
|
|
|
|
Cat : Category (lsuc (ℓ ⊔ ℓ')) (ℓ ⊔ ℓ')
|
|
|
|
|
Cat =
|
2018-01-08 21:29:29 +00:00
|
|
|
|
record
|
2018-01-21 00:11:08 +00:00
|
|
|
|
{ Object = Category ℓ ℓ'
|
2018-01-08 21:29:29 +00:00
|
|
|
|
; Arrow = Functor
|
|
|
|
|
; 𝟙 = identity
|
2018-01-21 14:21:50 +00:00
|
|
|
|
; _⊕_ = _∘f_
|
2018-01-20 23:21:25 +00:00
|
|
|
|
-- What gives here? Why can I not name the variables directly?
|
2018-01-21 18:23:24 +00:00
|
|
|
|
; isCategory = record
|
|
|
|
|
{ assoc = λ {_ _ _ _ f g h} → assc {f = f} {g = g} {h = h}
|
|
|
|
|
; ident = ident-r , ident-l
|
|
|
|
|
}
|
2018-01-08 21:29:29 +00:00
|
|
|
|
}
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-21 13:31:37 +00:00
|
|
|
|
module _ {ℓ : Level} (C D : Category ℓ ℓ) where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
private
|
2018-01-21 13:31:37 +00:00
|
|
|
|
:Object: = C .Object × D .Object
|
|
|
|
|
:Arrow: : :Object: → :Object: → Set ℓ
|
|
|
|
|
:Arrow: (c , d) (c' , d') = Arrow C c c' × Arrow D d d'
|
|
|
|
|
:𝟙: : {o : :Object:} → :Arrow: o o
|
|
|
|
|
:𝟙: = C .𝟙 , D .𝟙
|
|
|
|
|
_:⊕:_ :
|
|
|
|
|
{a b c : :Object:} →
|
|
|
|
|
:Arrow: b c →
|
|
|
|
|
:Arrow: a b →
|
|
|
|
|
:Arrow: a c
|
|
|
|
|
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) → (C ._⊕_) bc∈C ab∈C , D ._⊕_ bc∈D ab∈D}
|
|
|
|
|
|
|
|
|
|
instance
|
|
|
|
|
:isCategory: : IsCategory :Object: :Arrow: :𝟙: _:⊕:_
|
|
|
|
|
:isCategory: = record
|
|
|
|
|
{ assoc = eqpair C.assoc D.assoc
|
|
|
|
|
; ident
|
|
|
|
|
= eqpair (fst C.ident) (fst D.ident)
|
|
|
|
|
, eqpair (snd C.ident) (snd D.ident)
|
|
|
|
|
}
|
|
|
|
|
where
|
|
|
|
|
open module C = IsCategory (C .isCategory)
|
|
|
|
|
open module D = IsCategory (D .isCategory)
|
|
|
|
|
|
|
|
|
|
:product: : Category ℓ ℓ
|
|
|
|
|
:product: = record
|
|
|
|
|
{ Object = :Object:
|
|
|
|
|
; Arrow = :Arrow:
|
|
|
|
|
; 𝟙 = :𝟙:
|
|
|
|
|
; _⊕_ = _:⊕:_
|
|
|
|
|
}
|
|
|
|
|
|
2018-01-21 14:23:40 +00:00
|
|
|
|
proj₁ : Arrow Cat :product: C
|
2018-01-20 23:21:25 +00:00
|
|
|
|
proj₁ = record { func* = fst ; func→ = fst ; ident = refl ; distrib = refl }
|
|
|
|
|
|
2018-01-21 14:23:40 +00:00
|
|
|
|
proj₂ : Arrow Cat :product: D
|
2018-01-20 23:21:25 +00:00
|
|
|
|
proj₂ = record { func* = snd ; func→ = snd ; ident = refl ; distrib = refl }
|
|
|
|
|
|
2018-01-21 14:23:40 +00:00
|
|
|
|
module _ {X : Object (Cat {ℓ} {ℓ})} (x₁ : Arrow Cat X C) (x₂ : Arrow Cat X D) where
|
2018-01-20 23:21:25 +00:00
|
|
|
|
open Functor
|
|
|
|
|
|
|
|
|
|
-- ident' : {c : Object X} → ((func→ x₁) {dom = c} (𝟙 X) , (func→ x₂) {dom = c} (𝟙 X)) ≡ 𝟙 (catProduct C D)
|
|
|
|
|
-- ident' {c = c} = lift-eq (ident x₁) (ident x₂)
|
|
|
|
|
|
2018-01-21 13:31:37 +00:00
|
|
|
|
x : Functor X :product:
|
2018-01-20 23:21:25 +00:00
|
|
|
|
x = record
|
|
|
|
|
{ func* = λ x → (func* x₁) x , (func* x₂) x
|
|
|
|
|
; func→ = λ x → func→ x₁ x , func→ x₂ x
|
|
|
|
|
; ident = lift-eq (ident x₁) (ident x₂)
|
|
|
|
|
; distrib = lift-eq (distrib x₁) (distrib x₂)
|
|
|
|
|
}
|
|
|
|
|
|
|
|
|
|
-- Need to "lift equality of functors"
|
|
|
|
|
-- If I want to do this like I do it for pairs it's gonna be a pain.
|
2018-01-21 18:23:24 +00:00
|
|
|
|
postulate isUniqL : (Cat ⊕ proj₁) x ≡ x₁
|
|
|
|
|
-- isUniqL = lift-eq-functors refl refl {!!} {!!}
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-21 18:23:24 +00:00
|
|
|
|
postulate isUniqR : (Cat ⊕ proj₂) x ≡ x₂
|
|
|
|
|
-- isUniqR = lift-eq-functors refl refl {!!} {!!}
|
2018-01-20 23:21:25 +00:00
|
|
|
|
|
2018-01-21 14:23:40 +00:00
|
|
|
|
isUniq : (Cat ⊕ proj₁) x ≡ x₁ × (Cat ⊕ proj₂) x ≡ x₂
|
2018-01-20 23:21:25 +00:00
|
|
|
|
isUniq = isUniqL , isUniqR
|
|
|
|
|
|
2018-01-21 14:23:40 +00:00
|
|
|
|
uniq : ∃![ x ] ((Cat ⊕ proj₁) x ≡ x₁ × (Cat ⊕ proj₂) x ≡ x₂)
|
2018-01-20 23:21:25 +00:00
|
|
|
|
uniq = x , isUniq
|
|
|
|
|
|
|
|
|
|
instance
|
2018-01-21 14:23:40 +00:00
|
|
|
|
isProduct : IsProduct Cat proj₁ proj₂
|
2018-01-20 23:21:25 +00:00
|
|
|
|
isProduct = uniq
|
|
|
|
|
|
2018-01-21 14:23:40 +00:00
|
|
|
|
product : Product {ℂ = Cat} C D
|
2018-01-20 23:21:25 +00:00
|
|
|
|
product = record
|
2018-01-21 13:31:37 +00:00
|
|
|
|
{ obj = :product:
|
2018-01-20 23:21:25 +00:00
|
|
|
|
; proj₁ = proj₁
|
|
|
|
|
; proj₂ = proj₂
|
|
|
|
|
}
|