Changes to the category of categories
This commit is contained in:
parent
e8215b2c05
commit
e8ac6786ff
|
@ -9,7 +9,9 @@ open import Function
|
||||||
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
|
open import Data.Product renaming (proj₁ to fst ; proj₂ to snd)
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
open import Cat.Functor
|
open import Cat.Category.Functor
|
||||||
|
open import Cat.Category.Product
|
||||||
|
open import Cat.Category.Exponential
|
||||||
|
|
||||||
open import Cat.Equality
|
open import Cat.Equality
|
||||||
open Equality.Data.Product
|
open Equality.Data.Product
|
||||||
|
@ -26,12 +28,12 @@ module _ (ℓ ℓ' : Level) where
|
||||||
eq* : func* (H ∘f (G ∘f F)) ≡ func* ((H ∘f G) ∘f F)
|
eq* : func* (H ∘f (G ∘f F)) ≡ func* ((H ∘f G) ∘f F)
|
||||||
eq* = refl
|
eq* = refl
|
||||||
eq→ : PathP
|
eq→ : PathP
|
||||||
(λ i → {A B : 𝔸 .Object} → 𝔸 [ A , B ] → 𝔻 [ eq* i A , eq* i B ])
|
(λ i → {A B : Object 𝔸} → 𝔸 [ A , B ] → 𝔻 [ eq* i A , eq* i B ])
|
||||||
(func→ (H ∘f (G ∘f F))) (func→ ((H ∘f G) ∘f F))
|
(func→ (H ∘f (G ∘f F))) (func→ ((H ∘f G) ∘f F))
|
||||||
eq→ = refl
|
eq→ = refl
|
||||||
postulate
|
postulate
|
||||||
eqI
|
eqI
|
||||||
: (λ i → ∀ {A : 𝔸 .Object} → eq→ i (𝔸 .𝟙 {A}) ≡ 𝔻 .𝟙 {eq* i A})
|
: (λ i → ∀ {A : Object 𝔸} → eq→ i (𝟙 𝔸 {A}) ≡ 𝟙 𝔻 {eq* i A})
|
||||||
[ (H ∘f (G ∘f F)) .isFunctor .ident
|
[ (H ∘f (G ∘f F)) .isFunctor .ident
|
||||||
≡ ((H ∘f G) ∘f F) .isFunctor .ident
|
≡ ((H ∘f G) ∘f F) .isFunctor .ident
|
||||||
]
|
]
|
||||||
|
@ -58,12 +60,12 @@ module _ (ℓ ℓ' : Level) where
|
||||||
eq→ = refl
|
eq→ = refl
|
||||||
postulate
|
postulate
|
||||||
eqI-r
|
eqI-r
|
||||||
: (λ i → {c : ℂ .Object} → (λ _ → 𝔻 [ func* F c , func* F c ])
|
: (λ i → {c : Object ℂ} → (λ _ → 𝔻 [ func* F c , func* F c ])
|
||||||
[ func→ F (ℂ .𝟙) ≡ 𝔻 .𝟙 ])
|
[ func→ F (𝟙 ℂ) ≡ 𝟙 𝔻 ])
|
||||||
[(F ∘f identity) .isFunctor .ident ≡ F .isFunctor .ident ]
|
[(F ∘f identity) .isFunctor .ident ≡ F .isFunctor .ident ]
|
||||||
eqD-r : PathP
|
eqD-r : PathP
|
||||||
(λ i →
|
(λ i →
|
||||||
{A B C : ℂ .Object} {f : ℂ .Arrow A B} {g : ℂ .Arrow B C} →
|
{A B C : Object ℂ} {f : ℂ [ A , B ]} {g : ℂ [ B , C ]} →
|
||||||
eq→ i (ℂ [ g ∘ f ]) ≡ 𝔻 [ eq→ i g ∘ eq→ i f ])
|
eq→ i (ℂ [ g ∘ f ]) ≡ 𝔻 [ eq→ i g ∘ eq→ i f ])
|
||||||
((F ∘f identity) .isFunctor .distrib) (F .isFunctor .distrib)
|
((F ∘f identity) .isFunctor .distrib) (F .isFunctor .distrib)
|
||||||
ident-r : F ∘f identity ≡ F
|
ident-r : F ∘f identity ≡ F
|
||||||
|
@ -73,40 +75,50 @@ module _ (ℓ ℓ' : Level) where
|
||||||
postulate
|
postulate
|
||||||
eq* : (identity ∘f F) .func* ≡ F .func*
|
eq* : (identity ∘f F) .func* ≡ F .func*
|
||||||
eq→ : PathP
|
eq→ : PathP
|
||||||
(λ i → {x y : Object ℂ} → ℂ .Arrow x y → 𝔻 .Arrow (eq* i x) (eq* i y))
|
(λ i → {x y : Object ℂ} → ℂ [ x , y ] → 𝔻 [ eq* i x , eq* i y ])
|
||||||
((identity ∘f F) .func→) (F .func→)
|
((identity ∘f F) .func→) (F .func→)
|
||||||
eqI : (λ i → ∀ {A : ℂ .Object} → eq→ i (ℂ .𝟙 {A}) ≡ 𝔻 .𝟙 {eq* i A})
|
eqI : (λ i → ∀ {A : Object ℂ} → eq→ i (𝟙 ℂ {A}) ≡ 𝟙 𝔻 {eq* i A})
|
||||||
[ ((identity ∘f F) .isFunctor .ident) ≡ (F .isFunctor .ident) ]
|
[ ((identity ∘f F) .isFunctor .ident) ≡ (F .isFunctor .ident) ]
|
||||||
eqD : PathP (λ i → {A B C : ℂ .Object} {f : ℂ .Arrow A B} {g : ℂ .Arrow B C}
|
eqD : PathP (λ i → {A B C : Object ℂ} {f : ℂ [ A , B ]} {g : ℂ [ B , C ]}
|
||||||
→ eq→ i (ℂ [ g ∘ f ]) ≡ 𝔻 [ eq→ i g ∘ eq→ i f ])
|
→ eq→ i (ℂ [ g ∘ f ]) ≡ 𝔻 [ eq→ i g ∘ eq→ i f ])
|
||||||
((identity ∘f F) .isFunctor .distrib) (F .isFunctor .distrib)
|
((identity ∘f F) .isFunctor .distrib) (F .isFunctor .distrib)
|
||||||
-- (λ z → eq* i z) (eq→ i)
|
-- (λ z → eq* i z) (eq→ i)
|
||||||
ident-l : identity ∘f F ≡ F
|
ident-l : identity ∘f F ≡ F
|
||||||
ident-l = Functor≡ eq* eq→ λ i → record { ident = eqI i ; distrib = eqD i }
|
ident-l = Functor≡ eq* eq→ λ i → record { ident = eqI i ; distrib = eqD i }
|
||||||
|
|
||||||
Cat : Category (lsuc (ℓ ⊔ ℓ')) (ℓ ⊔ ℓ')
|
RawCat : RawCategory (lsuc (ℓ ⊔ ℓ')) (ℓ ⊔ ℓ')
|
||||||
Cat =
|
RawCat =
|
||||||
record
|
record
|
||||||
{ Object = Category ℓ ℓ'
|
{ Object = Category ℓ ℓ'
|
||||||
; Arrow = Functor
|
; Arrow = Functor
|
||||||
; 𝟙 = identity
|
; 𝟙 = identity
|
||||||
; _∘_ = _∘f_
|
; _∘_ = _∘f_
|
||||||
-- What gives here? Why can I not name the variables directly?
|
-- What gives here? Why can I not name the variables directly?
|
||||||
; isCategory = record
|
-- ; isCategory = record
|
||||||
{ assoc = λ {_ _ _ _ F G H} → assc {F = F} {G = G} {H = H}
|
-- { assoc = λ {_ _ _ _ F G H} → assc {F = F} {G = G} {H = H}
|
||||||
; ident = ident-r , ident-l
|
-- ; ident = ident-r , ident-l
|
||||||
}
|
-- }
|
||||||
}
|
}
|
||||||
|
open IsCategory
|
||||||
|
instance
|
||||||
|
:isCategory: : IsCategory RawCat
|
||||||
|
assoc :isCategory: {f = F} {G} {H} = assc {F = F} {G = G} {H = H}
|
||||||
|
ident :isCategory: = ident-r , ident-l
|
||||||
|
arrow-is-set :isCategory: = {!!}
|
||||||
|
univalent :isCategory: = {!!}
|
||||||
|
|
||||||
|
Cat : Category (lsuc (ℓ ⊔ ℓ')) (ℓ ⊔ ℓ')
|
||||||
|
raw Cat = RawCat
|
||||||
|
|
||||||
module _ {ℓ ℓ' : Level} where
|
module _ {ℓ ℓ' : Level} where
|
||||||
module _ (ℂ 𝔻 : Category ℓ ℓ') where
|
module _ (ℂ 𝔻 : Category ℓ ℓ') where
|
||||||
private
|
private
|
||||||
Catt = Cat ℓ ℓ'
|
Catt = Cat ℓ ℓ'
|
||||||
:Object: = ℂ .Object × 𝔻 .Object
|
:Object: = Object ℂ × Object 𝔻
|
||||||
:Arrow: : :Object: → :Object: → Set ℓ'
|
:Arrow: : :Object: → :Object: → Set ℓ'
|
||||||
:Arrow: (c , d) (c' , d') = Arrow ℂ c c' × Arrow 𝔻 d d'
|
:Arrow: (c , d) (c' , d') = Arrow ℂ c c' × Arrow 𝔻 d d'
|
||||||
:𝟙: : {o : :Object:} → :Arrow: o o
|
:𝟙: : {o : :Object:} → :Arrow: o o
|
||||||
:𝟙: = ℂ .𝟙 , 𝔻 .𝟙
|
:𝟙: = 𝟙 ℂ , 𝟙 𝔻
|
||||||
_:⊕:_ :
|
_:⊕:_ :
|
||||||
{a b c : :Object:} →
|
{a b c : :Object:} →
|
||||||
:Arrow: b c →
|
:Arrow: b c →
|
||||||
|
@ -114,25 +126,35 @@ module _ {ℓ ℓ' : Level} where
|
||||||
:Arrow: a c
|
:Arrow: a c
|
||||||
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) → ℂ [ bc∈C ∘ ab∈C ] , 𝔻 [ bc∈D ∘ ab∈D ]}
|
_:⊕:_ = λ { (bc∈C , bc∈D) (ab∈C , ab∈D) → ℂ [ bc∈C ∘ ab∈C ] , 𝔻 [ bc∈D ∘ ab∈D ]}
|
||||||
|
|
||||||
|
:rawProduct: : RawCategory ℓ ℓ'
|
||||||
|
RawCategory.Object :rawProduct: = :Object:
|
||||||
|
RawCategory.Arrow :rawProduct: = :Arrow:
|
||||||
|
RawCategory.𝟙 :rawProduct: = :𝟙:
|
||||||
|
RawCategory._∘_ :rawProduct: = _:⊕:_
|
||||||
|
|
||||||
|
module C = IsCategory (ℂ .isCategory)
|
||||||
|
module D = IsCategory (𝔻 .isCategory)
|
||||||
|
postulate
|
||||||
|
issSet : {A B : RawCategory.Object :rawProduct:} → isSet (RawCategory.Arrow :rawProduct: A B)
|
||||||
instance
|
instance
|
||||||
:isCategory: : IsCategory :Object: :Arrow: :𝟙: _:⊕:_
|
:isCategory: : IsCategory :rawProduct:
|
||||||
:isCategory: = record
|
-- :isCategory: = record
|
||||||
{ assoc = Σ≡ C.assoc D.assoc
|
-- { assoc = Σ≡ C.assoc D.assoc
|
||||||
; ident
|
-- ; ident
|
||||||
|
-- = Σ≡ (fst C.ident) (fst D.ident)
|
||||||
|
-- , Σ≡ (snd C.ident) (snd D.ident)
|
||||||
|
-- ; arrow-is-set = issSet
|
||||||
|
-- ; univalent = {!!}
|
||||||
|
-- }
|
||||||
|
IsCategory.assoc :isCategory: = Σ≡ C.assoc D.assoc
|
||||||
|
IsCategory.ident :isCategory:
|
||||||
= Σ≡ (fst C.ident) (fst D.ident)
|
= Σ≡ (fst C.ident) (fst D.ident)
|
||||||
, Σ≡ (snd C.ident) (snd D.ident)
|
, Σ≡ (snd C.ident) (snd D.ident)
|
||||||
}
|
IsCategory.arrow-is-set :isCategory: = issSet
|
||||||
where
|
IsCategory.univalent :isCategory: = {!!}
|
||||||
open module C = IsCategory (ℂ .isCategory)
|
|
||||||
open module D = IsCategory (𝔻 .isCategory)
|
|
||||||
|
|
||||||
:product: : Category ℓ ℓ'
|
:product: : Category ℓ ℓ'
|
||||||
:product: = record
|
raw :product: = :rawProduct:
|
||||||
{ Object = :Object:
|
|
||||||
; Arrow = :Arrow:
|
|
||||||
; 𝟙 = :𝟙:
|
|
||||||
; _∘_ = _:⊕:_
|
|
||||||
}
|
|
||||||
|
|
||||||
proj₁ : Catt [ :product: , ℂ ]
|
proj₁ : Catt [ :product: , ℂ ]
|
||||||
proj₁ = record { func* = fst ; func→ = fst ; isFunctor = record { ident = refl ; distrib = refl } }
|
proj₁ = record { func* = fst ; func→ = fst ; isFunctor = record { ident = refl ; distrib = refl } }
|
||||||
|
@ -143,28 +165,32 @@ module _ {ℓ ℓ' : Level} where
|
||||||
module _ {X : Object Catt} (x₁ : Catt [ X , ℂ ]) (x₂ : Catt [ X , 𝔻 ]) where
|
module _ {X : Object Catt} (x₁ : Catt [ X , ℂ ]) (x₂ : Catt [ X , 𝔻 ]) where
|
||||||
open Functor
|
open Functor
|
||||||
|
|
||||||
x : Functor X :product:
|
postulate x : Functor X :product:
|
||||||
x = record
|
-- x = record
|
||||||
{ func* = λ x → x₁ .func* x , x₂ .func* x
|
-- { func* = λ x → x₁ .func* x , x₂ .func* x
|
||||||
; func→ = λ x → func→ x₁ x , func→ x₂ x
|
-- ; func→ = λ x → func→ x₁ x , func→ x₂ x
|
||||||
; isFunctor = record
|
-- ; isFunctor = record
|
||||||
{ ident = Σ≡ x₁.ident x₂.ident
|
-- { ident = Σ≡ x₁.ident x₂.ident
|
||||||
; distrib = Σ≡ x₁.distrib x₂.distrib
|
-- ; distrib = Σ≡ x₁.distrib x₂.distrib
|
||||||
}
|
-- }
|
||||||
}
|
-- }
|
||||||
where
|
-- where
|
||||||
open module x₁ = IsFunctor (x₁ .isFunctor)
|
-- open module x₁ = IsFunctor (x₁ .isFunctor)
|
||||||
open module x₂ = IsFunctor (x₂ .isFunctor)
|
-- open module x₂ = IsFunctor (x₂ .isFunctor)
|
||||||
|
|
||||||
isUniqL : Catt [ proj₁ ∘ x ] ≡ x₁
|
-- Turned into postulate after:
|
||||||
isUniqL = Functor≡ eq* eq→ eqIsF -- Functor≡ refl refl λ i → record { ident = refl i ; distrib = refl i }
|
-- > commit e8215b2c051062c6301abc9b3f6ec67106259758 (HEAD -> dev, github/dev)
|
||||||
where
|
-- > Author: Frederik Hanghøj Iversen <fhi.1990@gmail.com>
|
||||||
eq* : (Catt [ proj₁ ∘ x ]) .func* ≡ x₁ .func*
|
-- > Date: Mon Feb 5 14:59:53 2018 +0100
|
||||||
eq* = refl
|
postulate isUniqL : Catt [ proj₁ ∘ x ] ≡ x₁
|
||||||
eq→ : (λ i → {A : X .Object} {B : X .Object} → X [ A , B ] → ℂ [ eq* i A , eq* i B ])
|
-- isUniqL = Functor≡ eq* eq→ {!!}
|
||||||
[ (Catt [ proj₁ ∘ x ]) .func→ ≡ x₁ .func→ ]
|
-- where
|
||||||
eq→ = refl
|
-- eq* : (Catt [ proj₁ ∘ x ]) .func* ≡ x₁ .func*
|
||||||
postulate eqIsF : (Catt [ proj₁ ∘ x ]) .isFunctor ≡ x₁ .isFunctor
|
-- eq* = {!refl!}
|
||||||
|
-- eq→ : (λ i → {A : Object X} {B : Object X} → X [ A , B ] → ℂ [ eq* i A , eq* i B ])
|
||||||
|
-- [ (Catt [ proj₁ ∘ x ]) .func→ ≡ x₁ .func→ ]
|
||||||
|
-- eq→ = refl
|
||||||
|
-- postulate eqIsF : (Catt [ proj₁ ∘ x ]) .isFunctor ≡ x₁ .isFunctor
|
||||||
-- eqIsF = IsFunctor≡ {!refl!} {!!}
|
-- eqIsF = IsFunctor≡ {!refl!} {!!}
|
||||||
|
|
||||||
postulate isUniqR : Catt [ proj₂ ∘ x ] ≡ x₂
|
postulate isUniqR : Catt [ proj₂ ∘ x ] ≡ x₂
|
||||||
|
@ -202,55 +228,55 @@ module _ (ℓ : Level) where
|
||||||
Catℓ = Cat ℓ ℓ
|
Catℓ = Cat ℓ ℓ
|
||||||
module _ (ℂ 𝔻 : Category ℓ ℓ) where
|
module _ (ℂ 𝔻 : Category ℓ ℓ) where
|
||||||
private
|
private
|
||||||
:obj: : Cat ℓ ℓ .Object
|
:obj: : Object (Cat ℓ ℓ)
|
||||||
:obj: = Fun {ℂ = ℂ} {𝔻 = 𝔻}
|
:obj: = Fun {ℂ = ℂ} {𝔻 = 𝔻}
|
||||||
|
|
||||||
:func*: : Functor ℂ 𝔻 × ℂ .Object → 𝔻 .Object
|
:func*: : Functor ℂ 𝔻 × Object ℂ → Object 𝔻
|
||||||
:func*: (F , A) = F .func* A
|
:func*: (F , A) = F .func* A
|
||||||
|
|
||||||
module _ {dom cod : Functor ℂ 𝔻 × ℂ .Object} where
|
module _ {dom cod : Functor ℂ 𝔻 × Object ℂ} where
|
||||||
private
|
private
|
||||||
F : Functor ℂ 𝔻
|
F : Functor ℂ 𝔻
|
||||||
F = proj₁ dom
|
F = proj₁ dom
|
||||||
A : ℂ .Object
|
A : Object ℂ
|
||||||
A = proj₂ dom
|
A = proj₂ dom
|
||||||
|
|
||||||
G : Functor ℂ 𝔻
|
G : Functor ℂ 𝔻
|
||||||
G = proj₁ cod
|
G = proj₁ cod
|
||||||
B : ℂ .Object
|
B : Object ℂ
|
||||||
B = proj₂ cod
|
B = proj₂ cod
|
||||||
|
|
||||||
:func→: : (pobj : NaturalTransformation F G × ℂ .Arrow A B)
|
:func→: : (pobj : NaturalTransformation F G × ℂ [ A , B ])
|
||||||
→ 𝔻 .Arrow (F .func* A) (G .func* B)
|
→ 𝔻 [ F .func* A , G .func* B ]
|
||||||
:func→: ((θ , θNat) , f) = result
|
:func→: ((θ , θNat) , f) = result
|
||||||
where
|
where
|
||||||
θA : 𝔻 .Arrow (F .func* A) (G .func* A)
|
θA : 𝔻 [ F .func* A , G .func* A ]
|
||||||
θA = θ A
|
θA = θ A
|
||||||
θB : 𝔻 .Arrow (F .func* B) (G .func* B)
|
θB : 𝔻 [ F .func* B , G .func* B ]
|
||||||
θB = θ B
|
θB = θ B
|
||||||
F→f : 𝔻 .Arrow (F .func* A) (F .func* B)
|
F→f : 𝔻 [ F .func* A , F .func* B ]
|
||||||
F→f = F .func→ f
|
F→f = F .func→ f
|
||||||
G→f : 𝔻 .Arrow (G .func* A) (G .func* B)
|
G→f : 𝔻 [ G .func* A , G .func* B ]
|
||||||
G→f = G .func→ f
|
G→f = G .func→ f
|
||||||
l : 𝔻 .Arrow (F .func* A) (G .func* B)
|
l : 𝔻 [ F .func* A , G .func* B ]
|
||||||
l = 𝔻 [ θB ∘ F→f ]
|
l = 𝔻 [ θB ∘ F→f ]
|
||||||
r : 𝔻 .Arrow (F .func* A) (G .func* B)
|
r : 𝔻 [ F .func* A , G .func* B ]
|
||||||
r = 𝔻 [ G→f ∘ θA ]
|
r = 𝔻 [ G→f ∘ θA ]
|
||||||
-- There are two choices at this point,
|
-- There are two choices at this point,
|
||||||
-- but I suppose the whole point is that
|
-- but I suppose the whole point is that
|
||||||
-- by `θNat f` we have `l ≡ r`
|
-- by `θNat f` we have `l ≡ r`
|
||||||
-- lem : 𝔻 [ θ B ∘ F .func→ f ] ≡ 𝔻 [ G .func→ f ∘ θ A ]
|
-- lem : 𝔻 [ θ B ∘ F .func→ f ] ≡ 𝔻 [ G .func→ f ∘ θ A ]
|
||||||
-- lem = θNat f
|
-- lem = θNat f
|
||||||
result : 𝔻 .Arrow (F .func* A) (G .func* B)
|
result : 𝔻 [ F .func* A , G .func* B ]
|
||||||
result = l
|
result = l
|
||||||
|
|
||||||
_×p_ = product
|
_×p_ = product
|
||||||
|
|
||||||
module _ {c : Functor ℂ 𝔻 × ℂ .Object} where
|
module _ {c : Functor ℂ 𝔻 × Object ℂ} where
|
||||||
private
|
private
|
||||||
F : Functor ℂ 𝔻
|
F : Functor ℂ 𝔻
|
||||||
F = proj₁ c
|
F = proj₁ c
|
||||||
C : ℂ .Object
|
C : Object ℂ
|
||||||
C = proj₂ c
|
C = proj₂ c
|
||||||
|
|
||||||
-- NaturalTransformation F G × ℂ .Arrow A B
|
-- NaturalTransformation F G × ℂ .Arrow A B
|
||||||
|
@ -259,19 +285,19 @@ module _ (ℓ : Level) where
|
||||||
-- where
|
-- where
|
||||||
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
|
-- open module 𝔻 = IsCategory (𝔻 .isCategory)
|
||||||
-- Unfortunately the equational version has some ambigous arguments.
|
-- Unfortunately the equational version has some ambigous arguments.
|
||||||
:ident: : :func→: {c} {c} (identityNat F , ℂ .𝟙 {o = proj₂ c}) ≡ 𝔻 .𝟙
|
:ident: : :func→: {c} {c} (identityNat F , 𝟙 ℂ {o = proj₂ c}) ≡ 𝟙 𝔻
|
||||||
:ident: = begin
|
:ident: = begin
|
||||||
:func→: {c} {c} ((:obj: ×p ℂ) .Product.obj .𝟙 {c}) ≡⟨⟩
|
:func→: {c} {c} (𝟙 (Product.obj (:obj: ×p ℂ)) {c}) ≡⟨⟩
|
||||||
:func→: {c} {c} (identityNat F , ℂ .𝟙) ≡⟨⟩
|
:func→: {c} {c} (identityNat F , 𝟙 ℂ) ≡⟨⟩
|
||||||
𝔻 [ identityTrans F C ∘ F .func→ (ℂ .𝟙)] ≡⟨⟩
|
𝔻 [ identityTrans F C ∘ F .func→ (𝟙 ℂ)] ≡⟨⟩
|
||||||
𝔻 [ 𝔻 .𝟙 ∘ F .func→ (ℂ .𝟙)] ≡⟨ proj₂ 𝔻.ident ⟩
|
𝔻 [ 𝟙 𝔻 ∘ F .func→ (𝟙 ℂ)] ≡⟨ proj₂ 𝔻.ident ⟩
|
||||||
F .func→ (ℂ .𝟙) ≡⟨ F.ident ⟩
|
F .func→ (𝟙 ℂ) ≡⟨ F.ident ⟩
|
||||||
𝔻 .𝟙 ∎
|
𝟙 𝔻 ∎
|
||||||
where
|
where
|
||||||
open module 𝔻 = IsCategory (𝔻 .isCategory)
|
open module 𝔻 = IsCategory (𝔻 .isCategory)
|
||||||
open module F = IsFunctor (F .isFunctor)
|
open module F = IsFunctor (F .isFunctor)
|
||||||
|
|
||||||
module _ {F×A G×B H×C : Functor ℂ 𝔻 × ℂ .Object} where
|
module _ {F×A G×B H×C : Functor ℂ 𝔻 × Object ℂ} where
|
||||||
F = F×A .proj₁
|
F = F×A .proj₁
|
||||||
A = F×A .proj₂
|
A = F×A .proj₂
|
||||||
G = G×B .proj₁
|
G = G×B .proj₁
|
||||||
|
@ -279,27 +305,27 @@ module _ (ℓ : Level) where
|
||||||
H = H×C .proj₁
|
H = H×C .proj₁
|
||||||
C = H×C .proj₂
|
C = H×C .proj₂
|
||||||
-- Not entirely clear what this is at this point:
|
-- Not entirely clear what this is at this point:
|
||||||
_P⊕_ = (:obj: ×p ℂ) .Product.obj .Category._∘_ {F×A} {G×B} {H×C}
|
_P⊕_ = Category._∘_ (Product.obj (:obj: ×p ℂ)) {F×A} {G×B} {H×C}
|
||||||
module _
|
module _
|
||||||
-- NaturalTransformation F G × ℂ .Arrow A B
|
-- NaturalTransformation F G × ℂ .Arrow A B
|
||||||
{θ×f : NaturalTransformation F G × ℂ .Arrow A B}
|
{θ×f : NaturalTransformation F G × ℂ [ A , B ]}
|
||||||
{η×g : NaturalTransformation G H × ℂ .Arrow B C} where
|
{η×g : NaturalTransformation G H × ℂ [ B , C ]} where
|
||||||
private
|
private
|
||||||
θ : Transformation F G
|
θ : Transformation F G
|
||||||
θ = proj₁ (proj₁ θ×f)
|
θ = proj₁ (proj₁ θ×f)
|
||||||
θNat : Natural F G θ
|
θNat : Natural F G θ
|
||||||
θNat = proj₂ (proj₁ θ×f)
|
θNat = proj₂ (proj₁ θ×f)
|
||||||
f : ℂ .Arrow A B
|
f : ℂ [ A , B ]
|
||||||
f = proj₂ θ×f
|
f = proj₂ θ×f
|
||||||
η : Transformation G H
|
η : Transformation G H
|
||||||
η = proj₁ (proj₁ η×g)
|
η = proj₁ (proj₁ η×g)
|
||||||
ηNat : Natural G H η
|
ηNat : Natural G H η
|
||||||
ηNat = proj₂ (proj₁ η×g)
|
ηNat = proj₂ (proj₁ η×g)
|
||||||
g : ℂ .Arrow B C
|
g : ℂ [ B , C ]
|
||||||
g = proj₂ η×g
|
g = proj₂ η×g
|
||||||
|
|
||||||
ηθNT : NaturalTransformation F H
|
ηθNT : NaturalTransformation F H
|
||||||
ηθNT = Fun .Category._∘_ {F} {G} {H} (η , ηNat) (θ , θNat)
|
ηθNT = Category._∘_ Fun {F} {G} {H} (η , ηNat) (θ , θNat)
|
||||||
|
|
||||||
ηθ = proj₁ ηθNT
|
ηθ = proj₁ ηθNT
|
||||||
ηθNat = proj₂ ηθNT
|
ηθNat = proj₂ ηθNT
|
||||||
|
@ -341,17 +367,28 @@ module _ (ℓ : Level) where
|
||||||
}
|
}
|
||||||
|
|
||||||
module _ (𝔸 : Category ℓ ℓ) (F : Functor ((𝔸 ×p ℂ) .Product.obj) 𝔻) where
|
module _ (𝔸 : Category ℓ ℓ) (F : Functor ((𝔸 ×p ℂ) .Product.obj) 𝔻) where
|
||||||
open HasProducts (hasProducts {ℓ} {ℓ}) using (parallelProduct)
|
open HasProducts (hasProducts {ℓ} {ℓ}) renaming (_|×|_ to parallelProduct)
|
||||||
|
|
||||||
postulate
|
postulate
|
||||||
transpose : Functor 𝔸 :obj:
|
transpose : Functor 𝔸 :obj:
|
||||||
eq : Catℓ [ :eval: ∘ (parallelProduct transpose (Catℓ .𝟙 {o = ℂ})) ] ≡ F
|
eq : Catℓ [ :eval: ∘ (parallelProduct transpose (𝟙 Catℓ {o = ℂ})) ] ≡ F
|
||||||
|
-- eq : Catℓ [ :eval: ∘ (HasProducts._|×|_ hasProducts transpose (𝟙 Catℓ {o = ℂ})) ] ≡ F
|
||||||
|
-- eq' : (Catℓ [ :eval: ∘
|
||||||
|
-- (record { product = product } HasProducts.|×| transpose)
|
||||||
|
-- (𝟙 Catℓ)
|
||||||
|
-- ])
|
||||||
|
-- ≡ F
|
||||||
|
|
||||||
catTranspose : ∃![ F~ ] (Catℓ [ :eval: ∘ (parallelProduct F~ (Catℓ .𝟙 {o = ℂ}))] ≡ F )
|
-- For some reason after `e8215b2c051062c6301abc9b3f6ec67106259758`
|
||||||
catTranspose = transpose , eq
|
-- `catTranspose` makes Agda hang. catTranspose : ∃![ F~ ] (Catℓ [
|
||||||
|
-- :eval: ∘ (parallelProduct F~ (𝟙 Catℓ {o = ℂ}))] ≡ F) catTranspose =
|
||||||
|
-- transpose , eq
|
||||||
|
|
||||||
:isExponential: : IsExponential Catℓ ℂ 𝔻 :obj: :eval:
|
:isExponential: : IsExponential Catℓ ℂ 𝔻 :obj: :eval:
|
||||||
:isExponential: = catTranspose
|
:isExponential: = {!catTranspose!}
|
||||||
|
where
|
||||||
|
open HasProducts (hasProducts {ℓ} {ℓ}) using (_|×|_)
|
||||||
|
-- :isExponential: = λ 𝔸 F → transpose 𝔸 F , eq' 𝔸 F
|
||||||
|
|
||||||
-- :exponent: : Exponential (Cat ℓ ℓ) A B
|
-- :exponent: : Exponential (Cat ℓ ℓ) A B
|
||||||
:exponent: : Exponential Catℓ ℂ 𝔻
|
:exponent: : Exponential Catℓ ℂ 𝔻
|
||||||
|
|
|
@ -76,8 +76,10 @@ record IsCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc
|
||||||
|
|
||||||
-- TODO: might want to implement isEquiv differently, there are 3
|
-- TODO: might want to implement isEquiv differently, there are 3
|
||||||
-- equivalent formulations in the book.
|
-- equivalent formulations in the book.
|
||||||
|
Univalent : Set (ℓa ⊔ ℓb)
|
||||||
|
Univalent = {A B : Object} → isEquiv (A ≡ B) (A ≅ B) (id-to-iso A B)
|
||||||
field
|
field
|
||||||
univalent : {A B : Object} → isEquiv (A ≡ B) (A ≅ B) (id-to-iso A B)
|
univalent : Univalent
|
||||||
|
|
||||||
module _ {A B : Object} where
|
module _ {A B : Object} where
|
||||||
Epimorphism : {X : Object } → (f : Arrow A B) → Set ℓb
|
Epimorphism : {X : Object } → (f : Arrow A B) → Set ℓb
|
||||||
|
|
|
@ -19,7 +19,7 @@ module _ {ℓ ℓ'} (ℂ : Category ℓ ℓ') {{hasProducts : HasProducts ℂ}}
|
||||||
module _ (B C : Object ℂ) where
|
module _ (B C : Object ℂ) where
|
||||||
IsExponential : (Cᴮ : Object ℂ) → ℂ [ Cᴮ ×p B , C ] → Set (ℓ ⊔ ℓ')
|
IsExponential : (Cᴮ : Object ℂ) → ℂ [ Cᴮ ×p B , C ] → Set (ℓ ⊔ ℓ')
|
||||||
IsExponential Cᴮ eval = ∀ (A : Object ℂ) (f : ℂ [ A ×p B , C ])
|
IsExponential Cᴮ eval = ∀ (A : Object ℂ) (f : ℂ [ A ×p B , C ])
|
||||||
→ ∃![ f~ ] (ℂ [ eval ∘ parallelProduct f~ (Category.𝟙 ℂ)] ≡ f)
|
→ ∃![ f~ ] (ℂ [ eval ∘ f~ |×| Category.𝟙 ℂ ] ≡ f)
|
||||||
|
|
||||||
record Exponential : Set (ℓ ⊔ ℓ') where
|
record Exponential : Set (ℓ ⊔ ℓ') where
|
||||||
field
|
field
|
||||||
|
|
|
@ -1,8 +1,8 @@
|
||||||
module Cat.Category.Product where
|
module Cat.Category.Product where
|
||||||
|
|
||||||
open import Agda.Primitive
|
open import Agda.Primitive
|
||||||
open import Data.Product
|
|
||||||
open import Cubical
|
open import Cubical
|
||||||
|
open import Data.Product as P hiding (_×_)
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
|
|
||||||
|
@ -12,14 +12,16 @@ module _ {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') {A B obj : Object ℂ} whe
|
||||||
IsProduct : (π₁ : ℂ [ obj , A ]) (π₂ : ℂ [ obj , B ]) → Set (ℓ ⊔ ℓ')
|
IsProduct : (π₁ : ℂ [ obj , A ]) (π₂ : ℂ [ obj , B ]) → Set (ℓ ⊔ ℓ')
|
||||||
IsProduct π₁ π₂
|
IsProduct π₁ π₂
|
||||||
= ∀ {X : Object ℂ} (x₁ : ℂ [ X , A ]) (x₂ : ℂ [ X , B ])
|
= ∀ {X : Object ℂ} (x₁ : ℂ [ X , A ]) (x₂ : ℂ [ X , B ])
|
||||||
→ ∃![ x ] (ℂ [ π₁ ∘ x ] ≡ x₁ × ℂ [ π₂ ∘ x ] ≡ x₂)
|
→ ∃![ x ] (ℂ [ π₁ ∘ x ] ≡ x₁ P.× ℂ [ π₂ ∘ x ] ≡ x₂)
|
||||||
|
|
||||||
-- Tip from Andrea; Consider this style for efficiency:
|
-- Tip from Andrea; Consider this style for efficiency:
|
||||||
-- record IsProduct {ℓ ℓ' : Level} (ℂ : Category {ℓ} {ℓ'})
|
-- record IsProduct {ℓa ℓb : Level} (ℂ : Category ℓa ℓb)
|
||||||
-- {A B obj : Object ℂ} (π₁ : Arrow ℂ obj A) (π₂ : Arrow ℂ obj B) : Set (ℓ ⊔ ℓ') where
|
-- {A B obj : Object ℂ} (π₁ : Arrow ℂ obj A) (π₂ : Arrow ℂ obj B) : Set (ℓa ⊔ ℓb) where
|
||||||
-- field
|
-- field
|
||||||
-- isProduct : ∀ {X : ℂ .Object} (x₁ : ℂ .Arrow X A) (x₂ : ℂ .Arrow X B)
|
-- issProduct : ∀ {X : Object ℂ} (x₁ : ℂ [ X , A ]) (x₂ : ℂ [ X , B ])
|
||||||
-- → ∃![ x ] (ℂ ._⊕_ π₁ x ≡ x₁ × ℂ. _⊕_ π₂ x ≡ x₂)
|
-- → ∃![ x ] (ℂ [ π₁ ∘ x ] ≡ x₁ P.× ℂ [ π₂ ∘ x ] ≡ x₂)
|
||||||
|
|
||||||
|
-- open IsProduct
|
||||||
|
|
||||||
record Product {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} (A B : Object ℂ) : Set (ℓ ⊔ ℓ') where
|
record Product {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} (A B : Object ℂ) : Set (ℓ ⊔ ℓ') where
|
||||||
no-eta-equality
|
no-eta-equality
|
||||||
|
@ -29,9 +31,9 @@ record Product {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} (A B : Object ℂ) :
|
||||||
proj₂ : ℂ [ obj , B ]
|
proj₂ : ℂ [ obj , B ]
|
||||||
{{isProduct}} : IsProduct ℂ proj₁ proj₂
|
{{isProduct}} : IsProduct ℂ proj₁ proj₂
|
||||||
|
|
||||||
arrowProduct : ∀ {X} → (π₁ : ℂ [ X , A ]) (π₂ : ℂ [ X , B ])
|
_P[_×_] : ∀ {X} → (π₁ : ℂ [ X , A ]) (π₂ : ℂ [ X , B ])
|
||||||
→ ℂ [ X , obj ]
|
→ ℂ [ X , obj ]
|
||||||
arrowProduct π₁ π₂ = proj₁ (isProduct π₁ π₂)
|
_P[_×_] π₁ π₂ = proj₁ (isProduct π₁ π₂)
|
||||||
|
|
||||||
record HasProducts {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') : Set (ℓ ⊔ ℓ') where
|
record HasProducts {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') : Set (ℓ ⊔ ℓ') where
|
||||||
field
|
field
|
||||||
|
@ -39,12 +41,14 @@ record HasProducts {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') : Set (ℓ ⊔
|
||||||
|
|
||||||
open Product
|
open Product
|
||||||
|
|
||||||
objectProduct : (A B : Object ℂ) → Object ℂ
|
_×_ : (A B : Object ℂ) → Object ℂ
|
||||||
objectProduct A B = Product.obj (product A B)
|
A × B = Product.obj (product A B)
|
||||||
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
|
-- The product mentioned in awodey in Def 6.1 is not the regular product of arrows.
|
||||||
-- It's a "parallel" product
|
-- It's a "parallel" product
|
||||||
parallelProduct : {A A' B B' : Object ℂ} → ℂ [ A , A' ] → ℂ [ B , B' ]
|
_|×|_ : {A A' B B' : Object ℂ} → ℂ [ A , A' ] → ℂ [ B , B' ]
|
||||||
→ ℂ [ objectProduct A B , objectProduct A' B' ]
|
→ ℂ [ A × B , A' × B' ]
|
||||||
parallelProduct {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B')
|
_|×|_ {A = A} {A' = A'} {B = B} {B' = B'} a b
|
||||||
(ℂ [ a ∘ (product A B) .proj₁ ])
|
= product A' B'
|
||||||
(ℂ [ b ∘ (product A B) .proj₂ ])
|
P[ ℂ [ a ∘ (product A B) .proj₁ ]
|
||||||
|
× ℂ [ b ∘ (product A B) .proj₂ ]
|
||||||
|
]
|
||||||
|
|
|
@ -60,7 +60,7 @@ open Functor
|
||||||
-- open import Cat.Categories.Fun
|
-- open import Cat.Categories.Fun
|
||||||
-- open import Cat.Categories.Sets
|
-- open import Cat.Categories.Sets
|
||||||
-- -- module Cat = Cat.Categories.Cat
|
-- -- module Cat = Cat.Categories.Cat
|
||||||
-- open Exponential
|
-- open import Cat.Category.Exponential
|
||||||
-- private
|
-- private
|
||||||
-- Catℓ = Cat ℓ ℓ
|
-- Catℓ = Cat ℓ ℓ
|
||||||
-- prshf = presheaf {ℂ = ℂ}
|
-- prshf = presheaf {ℂ = ℂ}
|
||||||
|
|
Loading…
Reference in a new issue