Split Category into RawCategory and IsCategory

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-05 11:43:38 +01:00
parent fecb4dc1ce
commit 22a9a71870
8 changed files with 189 additions and 158 deletions

View file

@ -12,7 +12,7 @@ open import Function
open import Relation.Nullary open import Relation.Nullary
open import Relation.Nullary.Decidable open import Relation.Nullary.Decidable
open import Cat.Category hiding (Hom) open import Cat.Category
open import Cat.Functor open import Cat.Functor
open import Cat.Equality open import Cat.Equality
open Equality.Data.Product open Equality.Data.Product
@ -66,12 +66,14 @@ module _ { ' : Level} (Ns : Set ) where
Hom = Σ Hom' rules Hom = Σ Hom' rules
-- The category of names and substitutions -- The category of names and substitutions
: Category -- o (lsuc lzero ⊔ o) Raw : RawCategory -- o (lsuc lzero ⊔ o)
= record Raw = record
{ Object = FiniteDecidableSubset { Object = FiniteDecidableSubset
-- { Object = Ns → Bool -- { Object = Ns → Bool
; Arrow = Hom ; Arrow = Hom
; 𝟙 = λ { {o} inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} } } ; 𝟙 = λ { {o} inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} } }
; _∘_ = {!!} ; _∘_ = {!!}
; isCategory = {!!}
} }
postulate RawIsCategory : IsCategory Raw
: Category
= Raw , RawIsCategory

View file

@ -13,36 +13,42 @@ open Equality.Data.Product
module _ (a b : Level) where module _ (a b : Level) where
private private
Obj = Σ[ A Set a ] (A Set b) Obj' = Σ[ A Set a ] (A Set b)
Arr : Obj Obj Set (a b) Arr : Obj' Obj' Set (a b)
Arr (A , B) (A' , B') = Σ[ f (A A') ] ({x : A} B x B' (f x)) Arr (A , B) (A' , B') = Σ[ f (A A') ] ({x : A} B x B' (f x))
one : {o : Obj} Arr o o one : {o : Obj'} Arr o o
proj₁ one = λ x x proj₁ one = λ x x
proj₂ one = λ b b proj₂ one = λ b b
_∘_ : {a b c : Obj} Arr b c Arr a b Arr a c _∘_ : {a b c : Obj'} Arr b c Arr a b Arr a c
(g , g') (f , f') = g Function.∘ f , g' Function.∘ f' (g , g') (f , f') = g Function.∘ f , g' Function.∘ f'
_⟨_∘_⟩ : {a b : Obj} (c : Obj) Arr b c Arr a b Arr a c _⟨_∘_⟩ : {a b : Obj'} (c : Obj') Arr b c Arr a b Arr a c
c g f = _∘_ {c = c} g f c g f = _∘_ {c = c} g f
module _ {A B C D : Obj} {f : Arr A B} {g : Arr B C} {h : Arr C D} where module _ {A B C D : Obj'} {f : Arr A B} {g : Arr B C} {h : Arr C D} where
assoc : (D h C g f ) D D h g f assoc : (D h C g f ) D D h g f
assoc = Σ≡ refl refl assoc = Σ≡ refl refl
module _ {A B : Obj} {f : Arr A B} where module _ {A B : Obj'} {f : Arr A B} where
ident : B f one f × B one {B} f f ident : B f one f × B one {B} f f
ident = (Σ≡ refl refl) , Σ≡ refl refl ident = (Σ≡ refl refl) , Σ≡ refl refl
RawFam : RawCategory (lsuc (a b)) (a b)
RawFam = record
{ Object = Obj'
; Arrow = Arr
; 𝟙 = one
; _∘_ = λ {a b c} _∘_ {a} {b} {c}
}
instance instance
isCategory : IsCategory Obj Arr one (λ {a b c} _∘_ {a} {b} {c}) isCategory : IsCategory RawFam
isCategory = record isCategory = record
{ assoc = λ {A} {B} {C} {D} {f} {g} {h} assoc {D = D} {f} {g} {h} { assoc = λ {A} {B} {C} {D} {f} {g} {h} assoc {D = D} {f} {g} {h}
; ident = λ {A} {B} {f} ident {A} {B} {f = f} ; ident = λ {A} {B} {f} ident {A} {B} {f = f}
; arrow-is-set = ?
; univalent = ?
} }
Fam : Category (lsuc (a b)) (a b) Fam : Category (lsuc (a b)) (a b)
Fam = record Fam = RawFam , isCategory
{ Object = Obj
; Arrow = Arr
; 𝟙 = one
; _∘_ = λ {a b c} _∘_ {a} {b} {c}
}

View file

@ -53,7 +53,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
𝔻 [ F→ f identityTrans F A ] 𝔻 [ F→ f identityTrans F A ]
where where
F→ = F .func→ F→ = F .func→
module 𝔻 = IsCategory (𝔻 .isCategory) module 𝔻 = IsCategory (isCategory 𝔻)
identityNat : (F : Functor 𝔻) NaturalTransformation F F identityNat : (F : Functor 𝔻) NaturalTransformation F F
identityNat F = identityTrans F , identityNatural F identityNat F = identityTrans F , identityNatural F
@ -75,7 +75,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
𝔻 [ H .func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩ 𝔻 [ H .func→ f 𝔻 [ θ A η A ] ] ≡⟨⟩
𝔻 [ H .func→ f (θ ∘nt η) A ] 𝔻 [ H .func→ f (θ ∘nt η) A ]
where where
open IsCategory (𝔻 .isCategory) open IsCategory (isCategory 𝔻)
NatComp = _:⊕:_ NatComp = _:⊕:_
@ -96,29 +96,33 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
× (_:⊕:_ {A} {B} {B} (identityNat B) f) f × (_:⊕:_ {A} {B} {B} (identityNat B) f) f
:ident: = ident-r , ident-l :ident: = ident-r , ident-l
instance
:isCategory: : IsCategory (Functor 𝔻) NaturalTransformation
(λ {F} identityNat F) (λ {a} {b} {c} _:⊕:_ {a} {b} {c})
:isCategory: = record
{ assoc = λ {A B C D} :assoc: {A} {B} {C} {D}
; ident = λ {A B} :ident: {A} {B}
}
-- Functor categories. Objects are functors, arrows are natural transformations. -- Functor categories. Objects are functors, arrows are natural transformations.
Fun : Category (c c' d d') (c c' d') RawFun : RawCategory (c c' d d') (c c' d')
Fun = record RawFun = record
{ Object = Functor 𝔻 { Object = Functor 𝔻
; Arrow = NaturalTransformation ; Arrow = NaturalTransformation
; 𝟙 = λ {F} identityNat F ; 𝟙 = λ {F} identityNat F
; _∘_ = λ {F G H} _:⊕:_ {F} {G} {H} ; _∘_ = λ {F G H} _:⊕:_ {F} {G} {H}
} }
instance
:isCategory: : IsCategory RawFun
:isCategory: = record
{ assoc = λ {A B C D} :assoc: {A} {B} {C} {D}
; ident = λ {A B} :ident: {A} {B}
; arrow-is-set = ?
; univalent = ?
}
Fun : Category (c c' d d') (c c' d')
Fun = RawFun , :isCategory:
module _ { ' : Level} ( : Category ') where module _ { ' : Level} ( : Category ') where
open import Cat.Categories.Sets open import Cat.Categories.Sets
-- Restrict the functors to Presheafs. -- Restrict the functors to Presheafs.
Presh : Category ( lsuc ') ( ') RawPresh : RawCategory ( lsuc ') ( ')
Presh = record RawPresh = record
{ Object = Presheaf { Object = Presheaf
; Arrow = NaturalTransformation ; Arrow = NaturalTransformation
; 𝟙 = λ {F} identityNat F ; 𝟙 = λ {F} identityNat F

View file

@ -154,16 +154,18 @@ module _ {A B C D : Set} {S : Subset (A × B)} {R : Subset (B × C)} {Q : Subset
(Σ[ b B ] (a , b) S × (Σ[ c C ] (b , c) R × (c , d) Q)) (Σ[ b B ] (a , b) S × (Σ[ c C ] (b , c) R × (c , d) Q))
is-assoc = equivToPath equi is-assoc = equivToPath equi
Rel : Category (lsuc lzero) (lsuc lzero) RawRel : RawCategory (lsuc lzero) (lsuc lzero)
Rel = record RawRel = record
{ Object = Set { Object = Set
; Arrow = λ S R Subset (S × R) ; Arrow = λ S R Subset (S × R)
; 𝟙 = λ {S} Diag S ; 𝟙 = λ {S} Diag S
; _∘_ = λ {A B C} S R λ {( a , c ) Σ[ b B ] ( (a , b) R × (b , c) S )} ; _∘_ = λ {A B C} S R λ {( a , c ) Σ[ b B ] ( (a , b) R × (b , c) S )}
; isCategory = record }
{ assoc = funExt is-assoc
; ident = funExt ident-l , funExt ident-r RawIsCategoryRel : IsCategory RawRel
; arrow-is-set = {!!} RawIsCategoryRel = record
; univalent = {!!} { assoc = funExt is-assoc
} ; ident = funExt ident-l , funExt ident-r
; arrow-is-set = {!!}
; univalent = {!!}
} }

View file

@ -4,22 +4,31 @@ module Cat.Categories.Sets where
open import Cubical open import Cubical
open import Agda.Primitive open import Agda.Primitive
open import Data.Product open import Data.Product
import Function
open import Cat.Category open import Cat.Category
open import Cat.Functor open import Cat.Functor
open Category open Category
module _ { : Level} where module _ { : Level} where
Sets : Category (lsuc ) SetsRaw : RawCategory (lsuc )
Sets = record SetsRaw = record
{ Object = Set { Object = Set
; Arrow = λ T U T U ; Arrow = λ T U T U
; 𝟙 = id ; 𝟙 = Function.id
; _∘_ = _∘_ ; _∘_ = Function._∘_
; isCategory = record { assoc = refl ; ident = funExt (λ _ refl) , funExt (λ _ refl) } }
SetsIsCategory : IsCategory SetsRaw
SetsIsCategory = record
{ assoc = refl
; ident = funExt (λ _ refl) , funExt (λ _ refl)
; arrow-is-set = {!!}
; univalent = {!!}
} }
where
open import Function Sets : Category (lsuc )
Sets = SetsRaw , SetsIsCategory
private private
module _ {X A B : Set } (f : X A) (g : X B) where module _ {X A B : Set } (f : X A) (g : X B) where
@ -55,7 +64,7 @@ representable { = } A = record
} }
} }
where where
open IsCategory ( .isCategory) open IsCategory (isCategory )
-- Contravariant Presheaf -- Contravariant Presheaf
Presheaf : { '} ( : Category ') Set ( lsuc ') Presheaf : { '} ( : Category ') Set ( lsuc ')
@ -72,4 +81,4 @@ presheaf { = } B = record
} }
} }
where where
open IsCategory ( .isCategory) open IsCategory (isCategory )

View file

@ -23,18 +23,37 @@ open import Cubical.GradLemma using ( propIsEquiv )
syntax ∃!-syntax (λ x B) = ∃![ x ] B syntax ∃!-syntax (λ x B) = ∃![ x ] B
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
-- Thierry: All projections must be `isProp`'s -- Thierry: All projections must be `isProp`'s
-- According to definitions 9.1.1 and 9.1.6 in the HoTT book the -- 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 -- arrows of a category form a set (arrow-is-set), and there is an
-- equivalence between the equality of objects and isomorphisms -- equivalence between the equality of objects and isomorphisms
-- (univalent). -- (univalent).
record IsCategory { ' : Level} record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where
(Object : Set ) open RawCategory
(Arrow : Object Object Set ') -- (Object : Set )
(𝟙 : {o : Object} Arrow o o) -- (Arrow : Object → Object → Set ')
(_∘_ : { a b c : Object } Arrow b c Arrow a b Arrow a c) -- (𝟙 : {o : Object} → Arrow o o)
: Set (lsuc (' )) where -- (_∘_ : { a b c : Object } → Arrow b c → Arrow a b → Arrow a c)
field field
assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D } assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D }
h (g f) (h g) f h (g f) (h g) f
@ -42,10 +61,10 @@ record IsCategory { ' : Level}
f 𝟙 f × 𝟙 f f f 𝟙 f × 𝟙 f f
arrow-is-set : {A B : Object} isSet (Arrow A B) arrow-is-set : {A B : Object} isSet (Arrow A B)
Isomorphism : {A B} (f : Arrow A B) Set ' Isomorphism : {A B} (f : Arrow A B) Set b
Isomorphism {A} {B} f = Σ[ g Arrow B A ] g f 𝟙 × f g 𝟙 Isomorphism {A} {B} f = Σ[ g Arrow B A ] g f 𝟙 × f g 𝟙
_≅_ : (A B : Object) Set ' _≅_ : (A B : Object) Set b
_≅_ A B = Σ[ f Arrow A B ] (Isomorphism f) _≅_ A B = Σ[ f Arrow A B ] (Isomorphism f)
idIso : (A : Object) A A idIso : (A : Object) A A
@ -61,20 +80,16 @@ record IsCategory { ' : Level}
univalent : {A B : Object} isEquiv (A B) (A B) (id-to-iso A B) univalent : {A B : Object} isEquiv (A B) (A B) (id-to-iso A B)
module _ {A B : Object} where module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set ' Epimorphism : {X : Object } (f : Arrow A B) Set b
Epimorphism {X} f = ( g₀ g₁ : Arrow B X ) g₀ f g₁ f g₀ g₁ Epimorphism {X} f = ( g₀ g₁ : Arrow B X ) g₀ f g₁ f g₀ g₁
Monomorphism : {X : Object} (f : Arrow A B) Set ' Monomorphism : {X : Object} (f : Arrow A B) Set b
Monomorphism {X} f = ( g₀ g₁ : Arrow X A ) f g₀ f g₁ g₀ g₁ Monomorphism {X} f = ( g₀ g₁ : Arrow X A ) f g₀ f g₁ g₀ g₁
module _ {} {'} {Object : Set } module _ {a} {b} { : RawCategory a b} where
{Arrow : Object Object Set '}
{𝟙 : {o : Object} Arrow o o}
{_⊕_ : { a b c : Object } Arrow b c Arrow a b Arrow a c}
where
-- TODO, provable by using arrow-is-set and that isProp (isEquiv _ _ _) -- TODO, provable by using arrow-is-set and that isProp (isEquiv _ _ _)
-- This lemma will be useful to prove the equality of two categories. -- This lemma will be useful to prove the equality of two categories.
IsCategory-is-prop : isProp (IsCategory Object Arrow 𝟙 _⊕_) IsCategory-is-prop : isProp (IsCategory )
IsCategory-is-prop x y i = record IsCategory-is-prop x y i = record
{ assoc = x.arrow-is-set _ _ x.assoc y.assoc i { assoc = x.arrow-is-set _ _ x.assoc y.assoc i
; ident = ; ident =
@ -94,38 +109,32 @@ module _ {} {'} {Object : Set }
module x = IsCategory x module x = IsCategory x
module y = IsCategory y module y = IsCategory y
record Category ( ' : Level) : Set (lsuc (' )) where Category : (a b : Level) Set (lsuc (a b))
-- adding no-eta-equality can speed up type-checking. Category a b = Σ (RawCategory a b) IsCategory
-- 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
{{isCategory}} : IsCategory Object Arrow 𝟙 _∘_
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
open Category module Category {a b : Level} ( : Category a b) where
raw = fst
open RawCategory raw public
isCategory = snd
_[_,_] : { '} ( : Category ') (A : .Object) (B : .Object) Set ' open RawCategory
_[_,_] = Arrow
_[_∘_] : { '} ( : Category ') {A B C : .Object} (g : [ B , C ]) (f : [ A , B ]) [ A , C ] -- _∈_ : ∀ {a b} ( : Category a b) → ( .fst .Object → Set b) → Set (a ⊔ b)
_[_∘_] = _∘_ -- A ∈ =
module _ { ' : Level} ( : Category ') {A B obj : Object } where Obj : {a b} Category a b Set a
IsProduct : (π₁ : Arrow obj A) (π₂ : Arrow obj B) Set ( ') Obj = .fst .Object
_[_,_] : { '} ( : Category ') (A : Obj ) (B : Obj ) Set '
[ A , B ] = .fst .Arrow A B
_[_∘_] : { '} ( : Category ') {A B C : Obj } (g : [ B , C ]) (f : [ A , B ]) [ A , C ]
[ g f ] = .fst ._∘_ g f
module _ { ' : Level} ( : Category ') {A B obj : Obj } where
IsProduct : (π₁ : [ obj , A ]) (π₂ : [ obj , B ]) Set ( ')
IsProduct π₁ π₂ IsProduct π₁ π₂
= {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B) = {X : Obj } (x₁ : [ X , A ]) (x₂ : [ X , B ])
∃![ x ] ( [ π₁ x ] x₁ × [ π₂ x ] x₂) ∃![ x ] ( [ π₁ x ] x₁ × [ π₂ x ] x₂)
-- Tip from Andrea; Consider this style for efficiency: -- Tip from Andrea; Consider this style for efficiency:
@ -135,48 +144,55 @@ module _ { ' : Level} ( : Category ') {A B obj : Object } whe
-- isProduct : ∀ {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B) -- isProduct : ∀ {X : .Object} (x₁ : .Arrow X A) (x₂ : .Arrow X B)
-- → ∃![ x ] ( ._⊕_ π₁ x ≡ x₁ × . _⊕_ π₂ x ≡ x₂) -- → ∃![ x ] ( ._⊕_ π₁ x ≡ x₁ × . _⊕_ π₂ x ≡ x₂)
record Product { ' : Level} { : Category '} (A B : .Object) : Set ( ') where record Product { ' : Level} { : Category '} (A B : Obj ) : Set ( ') where
no-eta-equality no-eta-equality
field field
obj : .Object obj : Obj
proj₁ : .Arrow obj A proj₁ : [ obj , A ]
proj₂ : .Arrow obj B proj₂ : [ obj , B ]
{{isProduct}} : IsProduct proj₁ proj₂ {{isProduct}} : IsProduct proj₁ proj₂
arrowProduct : {X} (π₁ : Arrow X A) (π₂ : Arrow X B) arrowProduct : {X} (π₁ : [ X , A ]) (π₂ : [ X , B ])
Arrow X obj [ X , obj ]
arrowProduct π₁ π₂ = fst (isProduct π₁ π₂) arrowProduct π₁ π₂ = fst (isProduct π₁ π₂)
record HasProducts { ' : Level} ( : Category ') : Set ( ') where record HasProducts { ' : Level} ( : Category ') : Set ( ') where
field field
product : (A B : .Object) Product { = } A B product : (A B : Obj ) Product { = } A B
open Product open Product
objectProduct : (A B : .Object) .Object objectProduct : (A B : Obj ) Obj
objectProduct A B = Product.obj (product A B) objectProduct 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} .Arrow A A' .Arrow B B' parallelProduct : {A A' B B' : Obj } [ A , A' ] [ B , B' ]
.Arrow (objectProduct A B) (objectProduct A' B') [ objectProduct A B , objectProduct A' B' ]
parallelProduct {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B') parallelProduct {A = A} {A' = A'} {B = B} {B' = B'} a b = arrowProduct (product A' B')
( [ a (product A B) .proj₁ ]) ( [ a (product A B) .proj₁ ])
( [ b (product A B) .proj₂ ]) ( [ b (product A B) .proj₂ ])
module _ { ' : Level} ( : Category ') where module _ {a b : Level} ( : Category a b) where
Opposite : Category ' private
Opposite = open Category
record module = RawCategory ( .fst)
{ Object = .Object OpRaw : RawCategory a b
; Arrow = Function.flip ( .Arrow) OpRaw = record
; 𝟙 = .𝟙 { Object = .Object
; _∘_ = Function.flip ( ._∘_) ; Arrow = Function.flip .Arrow
; isCategory = record { assoc = sym assoc ; ident = swap ident ; 𝟙 = .𝟙
; arrow-is-set = {!!} ; _∘_ = Function.flip (._∘_)
; univalent = {!!} }
} }
where open IsCategory isCategory
open IsCategory ( .isCategory) OpIsCategory : IsCategory OpRaw
OpIsCategory = record
{ assoc = sym assoc
; ident = swap ident
; arrow-is-set = {!!}
; univalent = {!!}
}
Opposite : Category a b
Opposite = OpRaw , OpIsCategory
-- A consequence of no-eta-equality; `Opposite-is-involution` is no longer -- A consequence of no-eta-equality; `Opposite-is-involution` is no longer
-- definitional - i.e.; you must match on the fields: -- definitional - i.e.; you must match on the fields:
@ -189,42 +205,34 @@ module _ { ' : Level} ( : Category ') where
-- assoc (Opposite-is-involution i) = {!!} -- assoc (Opposite-is-involution i) = {!!}
-- ident (Opposite-is-involution i) = {!!} -- ident (Opposite-is-involution i) = {!!}
Hom : { ' : Level} ( : Category ') (A B : Object ) Set '
Hom A B = Arrow A B
module _ { ' : Level} { : Category '} where
HomFromArrow : (A : .Object) {B B' : .Object} (g : .Arrow B B')
Hom A B Hom A B'
HomFromArrow _A = ._∘_
module _ { '} ( : Category ') {{hasProducts : HasProducts }} where module _ { '} ( : Category ') {{hasProducts : HasProducts }} where
open HasProducts hasProducts open HasProducts hasProducts
open Product hiding (obj) open Product hiding (obj)
private private
_×p_ : (A B : .Object) .Object _×p_ : (A B : Obj ) Obj
_×p_ A B = Product.obj (product A B) _×p_ A B = Product.obj (product A B)
module _ (B C : .Category.Object) where module _ (B C : Obj ) where
IsExponential : (Cᴮ : .Object) .Arrow (Cᴮ ×p B) C Set ( ') IsExponential : (Cᴮ : Obj ) [ Cᴮ ×p B , C ] Set ( ')
IsExponential Cᴮ eval = (A : .Object) (f : .Arrow (A ×p B) C) IsExponential Cᴮ eval = (A : Obj ) (f : [ A ×p B , C ])
∃![ f~ ] ( [ eval parallelProduct f~ ( .𝟙)] f) ∃![ f~ ] ( [ eval parallelProduct f~ (Category.raw .𝟙)] f)
record Exponential : Set ( ') where record Exponential : Set ( ') where
field field
-- obj ≡ Cᴮ -- obj ≡ Cᴮ
obj : .Object obj : Obj
eval : .Arrow ( obj ×p B ) C eval : [ obj ×p B , C ]
{{isExponential}} : IsExponential obj eval {{isExponential}} : IsExponential obj eval
-- If I make this an instance-argument then the instance resolution -- If I make this an instance-argument then the instance resolution
-- algorithm goes into an infinite loop. Why? -- algorithm goes into an infinite loop. Why?
exponentialsHaveProducts : HasProducts exponentialsHaveProducts : HasProducts
exponentialsHaveProducts = hasProducts exponentialsHaveProducts = hasProducts
transpose : (A : .Object) .Arrow (A ×p B) C .Arrow A obj transpose : (A : Obj ) [ A ×p B , C ] [ A , obj ]
transpose A f = fst (isExponential A f) transpose A f = fst (isExponential A f)
record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where record HasExponentials { ' : Level} ( : Category ') {{_ : HasProducts }} : Set ( ') where
field field
exponent : (A B : .Object) Exponential A B exponent : (A B : Obj ) Exponential A B
record CartesianClosed { ' : Level} ( : Category ') : Set ( ') where record CartesianClosed { ' : Level} ( : Category ') : Set ( ') where
field field
@ -234,15 +242,15 @@ record CartesianClosed { ' : Level} ( : Category ') : Set (
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
unique = isContr unique = isContr
IsInitial : .Object Set (a b) IsInitial : Obj Set (a b)
IsInitial I = {X : .Object} unique ( .Arrow I X) IsInitial I = {X : Obj } unique ( [ I , X ])
IsTerminal : .Object Set (a b) IsTerminal : Obj Set (a b)
-- ∃![ ? ] ? -- ∃![ ? ] ?
IsTerminal T = {X : .Object} unique ( .Arrow X T) IsTerminal T = {X : Obj } unique ( [ X , T ])
Initial : Set (a b) Initial : Set (a b)
Initial = Σ ( .Object) IsInitial Initial = Σ (Obj ) IsInitial
Terminal : Set (a b) Terminal : Set (a b)
Terminal = Σ ( .Object) IsTerminal Terminal = Σ (Obj ) IsTerminal

View file

@ -10,33 +10,33 @@ open import Cat.Category as C
module _ { ' : Level} ( : Category ') where module _ { ' : Level} ( : Category ') where
private private
open module = Category open module = Category
Obj = .Object
postulate postulate
Path : ( a b : Obj ) Set ' Path : ( a b : Obj ) Set '
emptyPath : (o : Obj) Path o o emptyPath : (o : Obj ) Path o o
concatenate : {a b c : Obj} Path b c Path a b Path a c concatenate : {a b c : Obj } Path b c Path a b Path a c
private private
module _ {A B C D : Obj} {r : Path A B} {q : Path B C} {p : Path C D} where module _ {A B C D : Obj } {r : Path A B} {q : Path B C} {p : Path C D} where
postulate postulate
p-assoc : concatenate {A} {C} {D} p (concatenate {A} {B} {C} q r) p-assoc : concatenate {A} {C} {D} p (concatenate {A} {B} {C} q r)
concatenate {A} {B} {D} (concatenate {B} {C} {D} p q) r concatenate {A} {B} {D} (concatenate {B} {C} {D} p q) r
module _ {A B : Obj} {p : Path A B} where module _ {A B : Obj } {p : Path A B} where
postulate postulate
ident-r : concatenate {A} {A} {B} p (emptyPath A) p ident-r : concatenate {A} {A} {B} p (emptyPath A) p
ident-l : concatenate {A} {B} {B} (emptyPath B) p p ident-l : concatenate {A} {B} {B} (emptyPath B) p p
Free : Category ' RawFree : RawCategory '
Free = record RawFree = record
{ Object = Obj { Object = Obj
; Arrow = Path ; Arrow = Path
; 𝟙 = λ {o} emptyPath o ; 𝟙 = λ {o} emptyPath o
; _∘_ = λ {a b c} concatenate {a} {b} {c} ; _∘_ = λ {a b c} concatenate {a} {b} {c}
; isCategory = record }
{ assoc = p-assoc RawIsCategoryFree : IsCategory RawFree
; ident = ident-r , ident-l RawIsCategoryFree = record
; arrow-is-set = {!!} { assoc = p-assoc
; univalent = {!!} ; ident = ident-r , ident-l
} ; arrow-is-set = {!!}
; univalent = {!!}
} }

View file

@ -10,11 +10,11 @@ open Category hiding (_∘_)
module _ {c c' d d'} ( : Category c c') (𝔻 : Category d d') where module _ {c c' d d'} ( : Category c c') (𝔻 : Category d d') where
record IsFunctor record IsFunctor
(func* : .Object 𝔻 .Object) (func* : Obj Obj 𝔻)
(func→ : {A B : .Object} [ A , B ] 𝔻 [ func* A , func* B ]) (func→ : {A B : Obj } [ A , B ] 𝔻 [ func* A , func* B ])
: Set (c c' d d') where : Set (c c' d d') where
field field
ident : { c : .Object } func→ ( .𝟙 {c}) 𝔻 .𝟙 {func* c} ident : {c : Obj } func→ ( .𝟙 {c}) 𝔻 .𝟙 {func* c}
-- TODO: Avoid use of ugly explicit arguments somehow. -- TODO: Avoid use of ugly explicit arguments somehow.
-- This guy managed to do it: -- This guy managed to do it:
-- https://github.com/copumpkin/categories/blob/master/Categories/Functor/Core.agda -- https://github.com/copumpkin/categories/blob/master/Categories/Functor/Core.agda