Skip to content
New issue

Have a question about this project? Sign up for a free GitHub account to open an issue and contact its maintainers and the community.

By clicking “Sign up for GitHub”, you agree to our terms of service and privacy statement. We’ll occasionally send you account related emails.

Already on GitHub? Sign in to your account

Add monoidal & enriched categories #665

Merged
merged 4 commits into from
Dec 17, 2021
Merged
Show file tree
Hide file tree
Changes from 2 commits
Commits
File filter

Filter by extension

Filter by extension

Conversations
Failed to load comments.
Loading
Jump to
Jump to file
Failed to load files.
Loading
Diff view
Diff view
3 changes: 3 additions & 0 deletions Cubical/Categories/Category/Base.agda
Original file line number Diff line number Diff line change
Expand Up @@ -27,6 +27,9 @@ record Category ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where
_∘_ : ∀ {x y z} (g : Hom[ y , z ]) (f : Hom[ x , y ]) → Hom[ x , z ]
g ∘ f = f ⋆ g

infixr 9 _⋆_
infixr 9 _∘_

open Category

-- Helpful syntax/notation
Expand Down
64 changes: 64 additions & 0 deletions Cubical/Categories/Constructions/Product.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,64 @@
-- Product of two categories
{-# OPTIONS --safe #-}

module Cubical.Categories.Constructions.Product where

open import Cubical.Categories.Category.Base
open import Cubical.Categories.Functor.Base
open import Cubical.Data.Sigma renaming (_×_ to _×'_)
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Prelude

private
variable
ℓC ℓC' ℓD ℓD' ℓE ℓE' : Level


open Category

_×_ : (C : Category ℓC ℓC') → (D : Category ℓD ℓD')
→ Category (ℓ-max ℓC ℓD) (ℓ-max ℓC' ℓD')

(C × D) .ob = (ob C) ×' (ob D)
(C × D) .Hom[_,_] (c , d) (c' , d') = (C [ c , c' ]) ×' (D [ d , d' ])
(C × D) .id = (id C , id D)
(C × D) ._⋆_ _ _ = (_ ⋆⟨ C ⟩ _ , _ ⋆⟨ D ⟩ _)
(C × D) .⋆IdL _ = ≡-× (⋆IdL C _) (⋆IdL D _)
(C × D) .⋆IdR _ = ≡-× (⋆IdR C _) (⋆IdR D _)
(C × D) .⋆Assoc _ _ _ = ≡-× (⋆Assoc C _ _ _) (⋆Assoc D _ _ _)
(C × D) .isSetHom = isSet× (isSetHom C) (isSetHom D)

infixr 5 _×_


-- Some useful functors
module _ (C : Category ℓC ℓC')
(D : Category ℓD ℓD') where
open Functor

module _ (E : Category ℓE ℓE') where
-- Associativity of product
×C-assoc : Functor (C × (D × E)) ((C × D) × E)
×C-assoc .F-ob (c , (d , e)) = ((c , d), e)
×C-assoc .F-hom (f , (g , h)) = ((f , g), h)
×C-assoc .F-id = refl
×C-assoc .F-seq _ _ = refl

-- Left/right injections into product
linj : (d : ob D) → Functor C (C × D)
linj d .F-ob c = (c , d)
linj d .F-hom f = (f , id D)
linj d .F-id = refl
linj d .F-seq f g = ≡-× refl (sym (⋆IdL D _))

rinj : (c : ob C) → Functor D (C × D)
rinj c .F-ob d = (c , d)
rinj c .F-hom f = (id C , f)
rinj c .F-id = refl
rinj c .F-seq f g = ≡-× (sym (⋆IdL C _)) refl

{-
TODO:
- define inverse to `assoc`, prove isomorphism
- prove product is commutative up to isomorphism
-}
3 changes: 2 additions & 1 deletion Cubical/Categories/Functor.agda
Original file line number Diff line number Diff line change
Expand Up @@ -3,5 +3,6 @@
module Cubical.Categories.Functor where

open import Cubical.Categories.Functor.Base public
open import Cubical.Categories.Functor.Properties public
open import Cubical.Categories.Functor.Compose public
open import Cubical.Categories.Functor.Product public
open import Cubical.Categories.Functor.Properties public
26 changes: 26 additions & 0 deletions Cubical/Categories/Functor/Product.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,26 @@
-- Product of two functors
{-# OPTIONS --safe #-}

module Cubical.Categories.Functor.Product where

open import Cubical.Categories.Category.Base
open import Cubical.Categories.Constructions.Product
open import Cubical.Categories.Functor.Base
open import Cubical.Data.Sigma.Properties
open import Cubical.Foundations.Prelude

private
variable
ℓA ℓA' ℓB ℓB' ℓC ℓC' ℓD ℓD' : Level
A : Category ℓA ℓA'
B : Category ℓB ℓB'
C : Category ℓC ℓC'
D : Category ℓD ℓD'

open Functor

_×F_ : Functor A C → Functor B D → Functor (A × B) (C × D)
(G ×F H) .F-ob (a , b) = (G ⟅ a ⟆ , H ⟅ b ⟆)
(G ×F H) .F-hom (g , h) = (G ⟪ g ⟫ , H ⟪ h ⟫)
(G ×F H) .F-id = ≡-× (G .F-id) (H .F-id)
(G ×F H) .F-seq _ _ = ≡-× (G .F-seq _ _) (H .F-seq _ _)
7 changes: 7 additions & 0 deletions Cubical/Categories/Monoidal.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,7 @@
-- Monoidal categories
{-# OPTIONS --safe #-}

module Cubical.Categories.Monoidal where

open import Cubical.Categories.Monoidal.Base public
open import Cubical.Categories.Monoidal.Enriched public
121 changes: 121 additions & 0 deletions Cubical/Categories/Monoidal/Base.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,121 @@
-- Monoidal categories
{-# OPTIONS --safe #-}

module Cubical.Categories.Monoidal.Base where

open import Cubical.Categories.Category.Base
open import Cubical.Categories.Constructions.Product
open import Cubical.Categories.Functor.Base
open import Cubical.Categories.Functor.Product
open import Cubical.Categories.Morphism
open import Cubical.Categories.NaturalTransformation.Base
open import Cubical.Foundations.Prelude


module _ {ℓ ℓ' : Level} (C : Category ℓ ℓ') where
open Category C

private
record TensorStr : Type (ℓ-max ℓ ℓ') where
field
─⊗─ : Functor (C × C) C
i : ob

open Functor

-- Useful tensor product notation
_⊗_ : ob → ob → ob
x ⊗ y = ─⊗─ .F-ob (x , y)

_⊗ₕ_ : ∀ {x y z w} → Hom[ x , y ] → Hom[ z , w ]
→ Hom[ x ⊗ z , y ⊗ w ]
f ⊗ₕ g = ─⊗─ .F-hom (f , g)


record StrictMonStr : Type (ℓ-max ℓ ℓ') where
field
tenstr : TensorStr

open TensorStr tenstr public

field
-- Axioms - strict
assoc : ∀ x y z → x ⊗ (y ⊗ z) ≡ (x ⊗ y) ⊗ z
idl : ∀ x → i ⊗ x ≡ x
idr : ∀ x → x ⊗ i ≡ x


record MonoidalStr : Type (ℓ-max ℓ ℓ') where
field
tenstr : TensorStr

open TensorStr tenstr public

private
-- Give some names to things
x⊗[y⊗z] : Functor (C × C × C) C
x⊗[y⊗z] = ─⊗─ ∘F (𝟙⟨ C ⟩ ×F ─⊗─)

[x⊗y]⊗z : Functor (C × C × C) C
[x⊗y]⊗z = ─⊗─ ∘F (─⊗─ ×F 𝟙⟨ C ⟩) ∘F (×C-assoc C C C)

x = 𝟙⟨ C ⟩
i⊗x = ─⊗─ ∘F (rinj C C i)
x⊗i = ─⊗─ ∘F (linj C C i)

field
-- "Axioms" - up to natural isomorphism
α : x⊗[y⊗z] ≅ᶜ [x⊗y]⊗z
η : i⊗x ≅ᶜ x
ρ : x⊗i ≅ᶜ x

open NatIso

-- More nice notations
α⟨_,_,_⟩ : (x y z : ob) → Hom[ x ⊗ (y ⊗ z) , (x ⊗ y) ⊗ z ]
α⟨ x , y , z ⟩ = α .trans ⟦ ( x , y , z ) ⟧

η⟨_⟩ : (x : ob) → Hom[ i ⊗ x , x ]
η⟨ x ⟩ = η .trans ⟦ x ⟧

ρ⟨_⟩ : (x : ob) → Hom[ x ⊗ i , x ]
ρ⟨ x ⟩ = ρ .trans ⟦ x ⟧

field
-- Coherence conditions
pentagon : ∀ w x y z →
id ⊗ₕ α⟨ x , y , z ⟩ ⋆ α⟨ w , x ⊗ y , z ⟩ ⋆ α⟨ w , x , y ⟩ ⊗ₕ id
≡ α⟨ w , x , y ⊗ z ⟩ ⋆ α⟨ w ⊗ x , y , z ⟩

triangle : ∀ x y →
α⟨ x , i , y ⟩ ⋆ ρ⟨ x ⟩ ⊗ₕ id ≡ id ⊗ₕ η⟨ y ⟩

open isIso

-- Inverses of α, η, ρ for convenience
α⁻¹⟨_,_,_⟩ : (x y z : ob) → Hom[ (x ⊗ y) ⊗ z , x ⊗ (y ⊗ z) ]
α⁻¹⟨ x , y , z ⟩ = α .nIso (x , y , z) .inv

η⁻¹⟨_⟩ : (x : ob) → Hom[ x , i ⊗ x ]
η⁻¹⟨ x ⟩ = η .nIso (x) .inv

ρ⁻¹⟨_⟩ : (x : ob) → Hom[ x , x ⊗ i ]
ρ⁻¹⟨ x ⟩ = ρ .nIso (x) .inv


record StrictMonCategory ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where
field
C : Category ℓ ℓ'
sms : StrictMonStr C

open Category C public
open StrictMonStr sms public


record MonoidalCategory ℓ ℓ' : Type (ℓ-suc (ℓ-max ℓ ℓ')) where
field
C : Category ℓ ℓ'
monstr : MonoidalStr C

open Category C public
open MonoidalStr monstr public
29 changes: 29 additions & 0 deletions Cubical/Categories/Monoidal/Enriched.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,29 @@
-- Enriched categories
{-# OPTIONS --safe #-}

module Cubical.Categories.Monoidal.Enriched where

open import Cubical.Categories.Monoidal.Base
open import Cubical.Foundations.Prelude

module _ {ℓV ℓV' : Level} (V : MonoidalCategory ℓV ℓV') (ℓE : Level) where

open MonoidalCategory V
renaming (ob to obV; Hom[_,_] to V[_,_]; id to idV; _⋆_ to _⋆V_)

record EnrichedCategory : Type (ℓ-max (ℓ-max ℓV ℓV') (ℓ-suc ℓE)) where
field
ob : Type ℓE
Hom[_,_] : ob → ob → obV
id : ∀ {x} → V[ i , Hom[ x , x ] ]
seq : ∀ x y z → V[ Hom[ x , y ] ⊗ Hom[ y , z ] , Hom[ x , z ] ]

-- Axioms
⋆IdL : ∀ x y → η⟨ _ ⟩ ≡ (id {x} ⊗ₕ idV) ⋆V (seq x x y)
⋆IdR : ∀ x y → ρ⟨ _ ⟩ ≡ (idV ⊗ₕ id {y}) ⋆V (seq x y y)
⋆Assoc : ∀ x y z w →
α⟨ _ , _ , _ ⟩ ⋆V ((seq x y z) ⊗ₕ idV) ⋆V (seq x z w)
≡ (idV ⊗ₕ (seq y z w)) ⋆V (seq x y w)


-- TODO: define underlying category using Hom[ x , y ] := V[ i , Hom[ x , y ] ]