Make properties of a category an instance argument

This commit is contained in:
Frederik Hanghøj Iversen 2018-01-21 14:31:37 +01:00
parent 07e4269399
commit 4c13334277
7 changed files with 109 additions and 90 deletions

View file

@ -70,16 +70,49 @@ module _ { ' : Level} where
; 𝟙 = identity ; 𝟙 = identity
; _⊕_ = functor-comp ; _⊕_ = functor-comp
-- What gives here? Why can I not name the variables directly? -- What gives here? Why can I not name the variables directly?
; assoc = λ {_ _ _ _ f g h} assc {f = f} {g = g} {h = h} ; isCategory = {!!}
; ident = ident-r , ident-l -- ; assoc = λ {_ _ _ _ f g h} → assc {f = f} {g = g} {h = h}
-- ; ident = ident-r , ident-l
} }
module _ { : Level} (C D : Category ) where module _ { : Level} (C D : Category ) where
private private
proj₁ : Arrow CatCat (catProduct C D) C :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:
; 𝟙 = :𝟙:
; _⊕_ = _:⊕:_
}
proj₁ : Arrow CatCat :product: C
proj₁ = record { func* = fst ; func→ = fst ; ident = refl ; distrib = refl } proj₁ = record { func* = fst ; func→ = fst ; ident = refl ; distrib = refl }
proj₂ : Arrow CatCat (catProduct C D) D proj₂ : Arrow CatCat :product: D
proj₂ = record { func* = snd ; func→ = snd ; ident = refl ; distrib = refl } proj₂ = record { func* = snd ; func→ = snd ; ident = refl ; distrib = refl }
module _ {X : Object (CatCat {} {})} (x₁ : Arrow CatCat X C) (x₂ : Arrow CatCat X D) where module _ {X : Object (CatCat {} {})} (x₁ : Arrow CatCat X C) (x₂ : Arrow CatCat X D) where
@ -88,7 +121,7 @@ module _ { : Level} (C D : Category ) where
-- ident' : {c : Object X} → ((func→ x₁) {dom = c} (𝟙 X) , (func→ x₂) {dom = c} (𝟙 X)) ≡ 𝟙 (catProduct C D) -- 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₂) -- ident' {c = c} = lift-eq (ident x₁) (ident x₂)
x : Functor X (catProduct C D) x : Functor X :product:
x = record x = record
{ func* = λ x (func* x₁) x , (func* x₂) x { func* = λ x (func* x₁) x , (func* x₂) x
; func→ = λ x func→ x₁ x , func→ x₂ x ; func→ = λ x func→ x₁ x , func→ x₂ x
@ -116,7 +149,7 @@ module _ { : Level} (C D : Category ) where
product : Product { = CatCat} C D product : Product { = CatCat} C D
product = record product = record
{ obj = catProduct C D { obj = :product:
; proj₁ = proj₁ ; proj₁ = proj₁
; proj₂ = proj₂ ; proj₂ = proj₂
} }

View file

@ -160,6 +160,5 @@ Rel = record
; 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 )}
; assoc = funExt is-assoc ; isCategory = record { assoc = funExt is-assoc ; ident = funExt ident-l , funExt ident-r }
; ident = funExt ident-l , funExt ident-r
} }

View file

@ -9,21 +9,18 @@ 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.Functor
open Category
-- Sets are built-in to Agda. The set of all small sets is called Set.
Fun : { : Level} ( T U : Set ) Set
Fun T U = T U
Sets : { : Level} Category (lsuc ) Sets : { : Level} Category (lsuc )
Sets {} = record Sets {} = record
{ Object = Set { Object = Set
; Arrow = λ T U Fun {} T U ; Arrow = λ T U T U
; 𝟙 = λ x x ; 𝟙 = id
; _⊕_ = λ g f x g ( f x ) ; _⊕_ = _∘_
; assoc = refl ; isCategory = record { assoc = refl ; ident = funExt (λ _ refl) , funExt (λ _ refl) }
; ident = funExt (λ x refl) , funExt (λ x refl)
} }
where
open import Function
-- Covariant Presheaf -- Covariant Presheaf
Representable : { ' : Level} ( : Category ') Set ( lsuc ') Representable : { ' : Level} ( : Category ') Set ( lsuc ')
@ -32,13 +29,13 @@ Representable {' = '} = Functor (Sets {'})
-- The "co-yoneda" embedding. -- The "co-yoneda" embedding.
representable : { '} { : Category '} Category.Object Representable representable : { '} { : Category '} Category.Object Representable
representable { = } A = record representable { = } A = record
{ func* = λ B .Arrow A B { func* = λ B .Arrow A B
; func→ = λ f g f .⊕ g ; func→ = ._⊕_
; ident = funExt λ _ snd .ident ; ident = funExt λ _ snd ident
; distrib = funExt λ x sym .assoc ; distrib = funExt λ x sym assoc
} }
where where
open module = Category open IsCategory ( .isCategory)
-- Contravariant Presheaf -- Contravariant Presheaf
Presheaf : { '} ( : Category ') Set ( lsuc ') Presheaf : { '} ( : Category ') Set ( lsuc ')
@ -47,10 +44,10 @@ Presheaf {' = '} = Functor (Opposite ) (Sets {'})
-- Alternate name: `yoneda` -- Alternate name: `yoneda`
presheaf : { ' : Level} { : Category '} Category.Object (Opposite ) Presheaf presheaf : { ' : Level} { : Category '} Category.Object (Opposite ) Presheaf
presheaf { = } B = record presheaf { = } B = record
{ func* = λ A .Arrow A B { func* = λ A .Arrow A B
; func→ = λ f g g .⊕ f ; func→ = λ f g ._⊕_ g f
; ident = funExt λ x fst .ident ; ident = funExt λ x fst ident
; distrib = funExt λ x .assoc ; distrib = funExt λ x assoc
} }
where where
open module = Category open IsCategory ( .isCategory)

View file

@ -24,6 +24,20 @@ syntax ∃!-syntax (λ x → B) = ∃![ x ] B
postulate undefined : { : Level} {A : Set } A postulate undefined : { : Level} {A : Set } A
record IsCategory { ' : Level}
(Object : Set )
(Arrow : Object Object Set ')
(𝟙 : {o : Object} Arrow o o)
(_⊕_ : { a b c : Object } Arrow b c Arrow a b Arrow a c)
: Set (lsuc (' )) where
field
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
-- open IsCategory public
record Category ( ' : Level) : Set (lsuc (' )) where record Category ( ' : Level) : Set (lsuc (' )) where
-- adding no-eta-equality can speed up type-checking. -- adding no-eta-equality can speed up type-checking.
no-eta-equality no-eta-equality
@ -32,17 +46,14 @@ record Category ( ' : Level) : Set (lsuc (' ⊔ )) where
Arrow : Object Object Set ' Arrow : Object Object Set '
𝟙 : {o : Object} Arrow o o 𝟙 : {o : Object} Arrow o o
_⊕_ : { a b c : Object } Arrow b c Arrow a b Arrow a c _⊕_ : { 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 } {{isCategory}} : IsCategory Object Arrow 𝟙 _⊕_
h (g f) (h g) f
ident : { A B : Object } { f : Arrow A B }
f 𝟙 f × 𝟙 f f
infixl 45 _⊕_ infixl 45 _⊕_
domain : { a b : Object } Arrow a b Object domain : { a b : Object } Arrow a b Object
domain {a = a} _ = a domain {a = a} _ = a
codomain : { a b : Object } Arrow a b Object codomain : { a b : Object } Arrow a b Object
codomain {b = b} _ = b codomain {b = b} _ = b
open Category public open Category
module _ { ' : Level} { : Category '} { A B : .Object } where module _ { ' : Level} { : Category '} { A B : .Object } where
private private
@ -61,26 +72,30 @@ module _ { ' : Level} { : Category '} { A B : .Object } wher
iso-is-epi : {X} (f : .Arrow A B) Isomorphism f Epimorphism {X = X} f 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 = iso-is-epi f (f- , left-inv , right-inv) g₀ g₁ eq =
begin begin
g₀ ≡⟨ sym (fst .ident) g₀ ≡⟨ sym (fst ident)
g₀ + .𝟙 ≡⟨ cong (_+_ g₀) (sym right-inv) g₀ + .𝟙 ≡⟨ cong (_+_ g₀) (sym right-inv)
g₀ + (f + f-) ≡⟨ .assoc g₀ + (f + f-) ≡⟨ assoc
(g₀ + f) + f- ≡⟨ cong (λ x x + f-) eq (g₀ + f) + f- ≡⟨ cong (λ x x + f-) eq
(g₁ + f) + f- ≡⟨ sym .assoc (g₁ + f) + f- ≡⟨ sym assoc
g₁ + (f + f-) ≡⟨ cong (_+_ g₁) right-inv g₁ + (f + f-) ≡⟨ cong (_+_ g₁) right-inv
g₁ + .𝟙 ≡⟨ fst .ident g₁ + .𝟙 ≡⟨ fst ident
g₁ g₁
where
open IsCategory .isCategory
iso-is-mono : {X} (f : .Arrow A B ) Isomorphism f Monomorphism {X = X} f 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 = iso-is-mono f (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin begin
g₀ ≡⟨ sym (snd .ident) g₀ ≡⟨ sym (snd ident)
.𝟙 + g₀ ≡⟨ cong (λ x x + g₀) (sym left-inv) .𝟙 + g₀ ≡⟨ cong (λ x x + g₀) (sym left-inv)
(f- + f) + g₀ ≡⟨ sym .assoc (f- + f) + g₀ ≡⟨ sym assoc
f- + (f + g₀) ≡⟨ cong (_+_ f-) eq f- + (f + g₀) ≡⟨ cong (_+_ f-) eq
f- + (f + g₁) ≡⟨ .assoc f- + (f + g₁) ≡⟨ assoc
(f- + f) + g₁ ≡⟨ cong (λ x x + g₁) left-inv (f- + f) + g₁ ≡⟨ cong (λ x x + g₁) left-inv
.𝟙 + g₁ ≡⟨ snd .ident .𝟙 + g₁ ≡⟨ snd ident
g₁ g₁
where
open IsCategory .isCategory
iso-is-epi-mono : {X} (f : .Arrow A B ) Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f 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 iso-is-epi-mono f iso = iso-is-epi f iso , iso-is-mono f iso
@ -118,49 +133,27 @@ record Product { ' : Level} { : Category '} (A B : .Object)
proj₂ : .Arrow obj B proj₂ : .Arrow obj B
{{isProduct}} : IsProduct proj₁ proj₂ {{isProduct}} : IsProduct proj₁ proj₂
mutual -- Two pairs are equal if their components are equal.
catProduct : {} (C D : Category ) Category eqpair : {a b} {A : Set a} {B : Set b} {a a' : A} {b b' : B}
catProduct C D = a a' b b' (a , b) (a' , b')
eqpair eqa eqb i = eqa i , eqb i
module _ { ' : Level} ( : Category ') where
private
instance
_ : IsCategory ( .Object) (flip ( .Arrow)) ( .𝟙) (flip ( ._⊕_))
_ = record { assoc = sym assoc ; ident = swap ident }
where
open IsCategory ( .isCategory)
Opposite : Category '
Opposite =
record record
{ Object = C.Object × D.Object { Object = .Object
-- Why does "outlining with `arrowProduct` not work? ; Arrow = flip ( .Arrow)
; Arrow = λ {(c , d) (c' , d') Arrow C c c' × Arrow D d d'} ; 𝟙 = .𝟙
; 𝟙 = C.𝟙 , D.𝟙 ; _⊕_ = flip ( ._⊕_)
; _⊕_ = λ { (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'
Opposite : { '} Category ' Category '
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 -- 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:

View file

@ -34,6 +34,5 @@ module _ { ' : Level} ( : Category ') where
; Arrow = Path ; Arrow = Path
; 𝟙 = λ {o} emptyPath o ; 𝟙 = λ {o} emptyPath o
; _⊕_ = λ {a b c} concatenate {a} {b} {c} ; _⊕_ = λ {a b c} concatenate {a} {b} {c}
; assoc = p-assoc ; isCategory = record { assoc = p-assoc ; ident = ident-r , ident-l }
; ident = ident-r , ident-l
} }

View file

@ -7,14 +7,13 @@ open import Cat.Functor
open import Cat.Categories.Sets open import Cat.Categories.Sets
module _ {a a' b b'} where module _ {a a' b b'} where
Exponential : Category a a' Category b b' Category ? ? Exponential : Category a a' Category b b' Category {!!} {!!}
Exponential A B = record Exponential A B = record
{ Object = {!!} { Object = {!!}
; Arrow = {!!} ; Arrow = {!!}
; 𝟙 = {!!} ; 𝟙 = {!!}
; _⊕_ = {!!} ; _⊕_ = {!!}
; assoc = {!!} ; isCategory = ?
; ident = {!!}
} }
_⇑_ = Exponential _⇑_ = Exponential

View file

@ -49,6 +49,5 @@ module _ { ' : Level} (Ns : Set ) where
; Arrow = Mor ; Arrow = Mor
; 𝟙 = {!!} ; 𝟙 = {!!}
; _⊕_ = {!!} ; _⊕_ = {!!}
; assoc = {!!} ; isCategory = ?
; ident = {!!}
} }