Skip to content
Open
Show file tree
Hide file tree
Changes from all commits
Commits
Show all changes
45 commits
Select commit Hold shift + click to select a range
c744da6
Create folder for Dagger Categories
anshwad10 Mar 2, 2025
5c8fb22
Add files via upload
anshwad10 Mar 2, 2025
c6f6978
Update Base.agda
anshwad10 Mar 2, 2025
5d787ae
Add files via upload
anshwad10 Mar 2, 2025
10a563e
Update Properties.agda
anshwad10 Mar 2, 2025
35c2601
add syntax for dagger, fix whitespace, other minor improvements
anshwad10 Mar 2, 2025
30658a5
fix whitespace, add more lemmas for univalence
anshwad10 Mar 2, 2025
0bf35c1
Add files via upload
anshwad10 Mar 2, 2025
97dd20e
Dagger functors
anshwad10 Mar 2, 2025
8083a73
`areInv` is a proposition
anshwad10 Mar 2, 2025
3a41cf4
Update Properties.agda
anshwad10 Mar 2, 2025
3f1e345
fix whitespace
anshwad10 Mar 2, 2025
a3daeed
fix whitespace
anshwad10 Mar 2, 2025
f35dca3
remove trailing newlines
anshwad10 Mar 2, 2025
c5e0208
Update Morphism.agda
anshwad10 Mar 2, 2025
b57d802
fix type error
anshwad10 Mar 2, 2025
127a1c7
fix import
anshwad10 Mar 2, 2025
4cd983f
Create Dagger.agda
anshwad10 Mar 2, 2025
c220ad1
Update Properties.agda
anshwad10 Mar 3, 2025
f3bd6f1
undo Update Properties.agda
anshwad10 Mar 3, 2025
bbaa94f
Update Functor.agda
anshwad10 Mar 3, 2025
7d12126
add binary products of dagger categories
anshwad10 Mar 3, 2025
0fd5bb7
Add files via upload
anshwad10 Mar 3, 2025
62ab610
Update Functors.agda
anshwad10 Mar 3, 2025
a82cb33
Update BinProduct.agda
anshwad10 Mar 3, 2025
87678c8
fix whitespace
anshwad10 Mar 3, 2025
251aa71
fix whitespace
anshwad10 Mar 3, 2025
ee66405
fix whitespace
anshwad10 Mar 3, 2025
4327fef
Update Functor.agda
anshwad10 Mar 4, 2025
a76600a
Update Functor.agda
anshwad10 Mar 4, 2025
a5faa06
import Functors.Constant
anshwad10 Mar 4, 2025
c79c8c0
import natural transformations
anshwad10 Mar 4, 2025
9711894
add hidden argument puns
anshwad10 Mar 4, 2025
40f1026
Update Properties.agda
anshwad10 Mar 4, 2025
0a3bbf0
Update Properties.agda
anshwad10 Mar 4, 2025
38c176b
Merge branch 'agda:master' into dagger-cat
anshwad10 May 14, 2025
ff32908
Every dagger category is equivalent to its opposite
anshwad10 May 15, 2025
c0cdca4
fix whitespace
anshwad10 May 15, 2025
ab77792
Credit Karvonen, rename PIso
anshwad10 Jul 23, 2025
8b1b6da
Consistent naming
anshwad10 Jul 24, 2025
da49b83
fix whitespace
anshwad10 Jul 24, 2025
3c8bb63
Update BinProduct.agda
anshwad10 Jul 24, 2025
4a1c3b8
Merge branch 'dagger-cat' of https://github.com/anshwad10/cubical int…
anshwad10 Jul 24, 2025
e41b143
Update Functors.agda
anshwad10 Jul 24, 2025
5e97299
Update Dagger.agda
anshwad10 Jul 24, 2025
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
10 changes: 10 additions & 0 deletions Cubical/Categories/Dagger.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,10 @@
-- Dagger categories
{-# OPTIONS --safe #-}

module Cubical.Categories.Dagger where

open import Cubical.Categories.Dagger.Base public
open import Cubical.Categories.Dagger.Properties public
open import Cubical.Categories.Dagger.Functor public
open import Cubical.Categories.Dagger.Instances.BinProduct public
open import Cubical.Categories.Dagger.Instances.Functors public
64 changes: 64 additions & 0 deletions Cubical/Categories/Dagger/Base.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,64 @@
{-# OPTIONS --safe #-}

module Cubical.Categories.Dagger.Base where

open import Cubical.Foundations.Prelude

open import Cubical.Categories.Category

private variable
ℓ ℓ' : Level

module _ (C : Category ℓ ℓ') where
open Category C
private variable
x y z w : ob

record IsDagger (_† : {x y : ob} → Hom[ x , y ] → Hom[ y , x ]) : Type (ℓ-max ℓ ℓ') where
field
†-invol : (f : Hom[ x , y ]) → f † † ≡ f
†-id : id † ≡ id {x}
†-seq : (f : Hom[ x , y ]) (g : Hom[ y , z ]) → (f ⋆ g) † ≡ g † ⋆ f †

open IsDagger

makeIsDagger : {_† : {x y : ob} → Hom[ x , y ] → Hom[ y , x ]}
→ (∀ {x y} (f : Hom[ x , y ]) → f † † ≡ f)
→ (∀ {x y z} (f : Hom[ x , y ]) (g : Hom[ y , z ]) → (f ⋆ g) † ≡ g † ⋆ f †)
→ IsDagger _†
makeIsDagger {_†} †-invol †-seq .†-invol = †-invol
makeIsDagger {_†} †-invol †-seq .†-seq = †-seq
makeIsDagger {_†} †-invol †-seq .†-id = -- this actually follows from the other axioms
id † ≡⟨ sym (⋆IdR _) ⟩
id † ⋆ id ≡⟨ congR _⋆_ (sym (†-invol id)) ⟩
id † ⋆ id † † ≡⟨ sym (†-seq (id †) id) ⟩
(id † ⋆ id) † ≡⟨ cong _† (⋆IdR _) ⟩
id † † ≡⟨ †-invol id ⟩
id ∎

record DaggerStr : Type (ℓ-max ℓ ℓ') where
field
_† : Hom[ x , y ] → Hom[ y , x ]
is-dag : IsDagger _†

open IsDagger is-dag public


record †Category (ℓ ℓ' : Level) : Type (ℓ-suc (ℓ-max ℓ ℓ')) where
no-eta-equality

field
cat : Category ℓ ℓ'
dagstr : DaggerStr cat

open DaggerStr dagstr public
open Category cat public

open IsDagger
open DaggerStr
open †Category

dag : ∀ (C : †Category ℓ ℓ') {x y} → C .cat [ x , y ] → C .cat [ y , x ]
dag C f = C ._† f

syntax dag C f = f †[ C ]
73 changes: 73 additions & 0 deletions Cubical/Categories/Dagger/Functor.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,73 @@
{-# OPTIONS --safe --hidden-argument-puns #-}

module Cubical.Categories.Dagger.Functor where

open import Cubical.Foundations.Prelude
open import Cubical.Data.Sigma

open import Cubical.Categories.Category
open import Cubical.Categories.Functor
open import Cubical.Categories.NaturalTransformation
open import Cubical.Categories.Functors.Constant

open import Cubical.Categories.Dagger.Base

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

open †Category

module _ (C : †Category ℓC ℓC') (D : †Category ℓD ℓD') where
record IsDagFunctor (F : Functor (C .cat) (D .cat)) : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) where
no-eta-equality
field
F-† : ∀ {x y} (f : C .cat [ x , y ]) → F ⟪ f †[ C ] ⟫ ≡ F ⟪ f ⟫ †[ D ]

open IsDagFunctor public

isPropIsDagFunctor : ∀ F → isProp (IsDagFunctor F)
isPropIsDagFunctor F a b i .F-† f = D .isSetHom _ _ (a .F-† f) (b .F-† f) i

DagFunctor : Type (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD'))
DagFunctor = Σ[ F ∈ Functor (C .cat) (D .cat) ] IsDagFunctor F

open Functor

†Id : DagFunctor C C
†Id .fst = Id
†Id .snd .F-† f = refl

†FuncComp : DagFunctor C D → DagFunctor D E → DagFunctor C E
†FuncComp F G .fst = G .fst ∘F F .fst
†FuncComp {C = C} {D = D} {E = E} F G .snd .F-† f =
G .fst ⟪ F .fst ⟪ f †[ C ] ⟫ ⟫ ≡⟨ cong (G .fst .F-hom) (F .snd .F-† f) ⟩
G .fst ⟪ F .fst ⟪ f ⟫ †[ D ] ⟫ ≡⟨ G .snd .F-† (F .fst .F-hom f) ⟩
G .fst ⟪ F .fst ⟪ f ⟫ ⟫ †[ E ] ∎

_∘†F_ : DagFunctor D E → DagFunctor C D → DagFunctor C E
G ∘†F F = †FuncComp F G

†Func≡ : (F G : DagFunctor C D) → F .fst ≡ G .fst → F ≡ G
†Func≡ {C} {D} F G = Σ≡Prop (isPropIsDagFunctor C D)

open †Category

†Constant : ob D → DagFunctor C D
†Constant {D} d .fst = Constant _ _ d
†Constant {D} d .snd .F-† _ = sym (D .†-id)

open NatTrans

NT† : (F G : DagFunctor C D) → NatTrans (F .fst) (G .fst) → NatTrans (G .fst) (F .fst)
NT† {C} {D} F G n .N-ob x = n ⟦ x ⟧ †[ D ]
NT† {C} {D} F G n .N-hom {x} {y} f =
G .fst ⟪ f ⟫ D.⋆ n ⟦ y ⟧ D.† ≡⟨ congL D._⋆_ (sym (D .†-invol (G .fst ⟪ f ⟫))) ⟩
G .fst ⟪ f ⟫ D.† D.† D.⋆ n ⟦ y ⟧ D.† ≡⟨ sym (D .†-seq (n ⟦ y ⟧) (G .fst ⟪ f ⟫ D.†)) ⟩
(n ⟦ y ⟧ D.⋆ G .fst ⟪ f ⟫ D.†) D.† ≡⟨ cong D._† (congR D._⋆_ (sym (G .snd .F-† f))) ⟩
(n ⟦ y ⟧ D.⋆ G .fst ⟪ f C.† ⟫) D.† ≡⟨ cong D._† (sym (n .N-hom (f C.†))) ⟩
(F .fst ⟪ f C.† ⟫ D.⋆ n ⟦ x ⟧) D.† ≡⟨ cong D._† (congL D._⋆_ (F .snd .F-† f)) ⟩
(F .fst ⟪ f ⟫ D.† D.⋆ n ⟦ x ⟧) D.† ≡⟨ D .†-seq (F .fst ⟪ f ⟫ D.†) (n ⟦ x ⟧) ⟩
n ⟦ x ⟧ D.† D.⋆ F .fst ⟪ f ⟫ D.† D.† ≡⟨ congR D._⋆_ (D .†-invol (F .fst ⟪ f ⟫)) ⟩
n ⟦ x ⟧ D.† D.⋆ F .fst ⟪ f ⟫ ∎
where module D = †Category D; module C = †Category C
74 changes: 74 additions & 0 deletions Cubical/Categories/Dagger/Instances/BinProduct.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,74 @@
{-# OPTIONS --safe #-}

module Cubical.Categories.Dagger.Instances.BinProduct where

open import Cubical.Foundations.Prelude
open import Cubical.Data.Sigma

open import Cubical.Categories.Category
open import Cubical.Categories.Functor
open import Cubical.Categories.Constructions.BinProduct
open import Cubical.Categories.Morphism

open import Cubical.Categories.Dagger.Base
open import Cubical.Categories.Dagger.Properties
open import Cubical.Categories.Dagger.Functor

private variable
ℓ ℓ' ℓ'' ℓ''' : Level

open DaggerStr
open IsDagger
open †Category

BinProdDaggerStr : {C : Category ℓ ℓ'} {D : Category ℓ'' ℓ'''} → DaggerStr C → DaggerStr D → DaggerStr (C ×C D)
BinProdDaggerStr dagC dagD ._† (f , g) = dagC ._† f , dagD ._† g
BinProdDaggerStr dagC dagD .is-dag .†-invol (f , g) = ≡-× (dagC .†-invol f) (dagD .†-invol g)
BinProdDaggerStr dagC dagD .is-dag .†-id = ≡-× (dagC .†-id) (dagD .†-id)
BinProdDaggerStr dagC dagD .is-dag .†-seq (f , g) (f' , g') = ≡-× (dagC .†-seq f f') (dagD .†-seq g g')

DagBinProd _׆_ : †Category ℓ ℓ' → †Category ℓ'' ℓ''' → †Category (ℓ-max ℓ ℓ'') (ℓ-max ℓ' ℓ''')
DagBinProd C D .cat = C .cat ×C D .cat
DagBinProd C D .dagstr = BinProdDaggerStr (C .dagstr) (D .dagstr)
_׆_ = DagBinProd

module _ (C : †Category ℓ ℓ') (D : †Category ℓ'' ℓ''') where
†Fst : DagFunctor (C ׆ D) C
†Fst .fst = Fst (C .cat) (D .cat)
†Fst .snd .F-† (f , g) = refl

†Snd : DagFunctor (C ׆ D) D
†Snd .fst = Snd (C .cat) (D .cat)
†Snd .snd .F-† (f , g) = refl

module _ where
private variable
B C D E : †Category ℓ ℓ'

_,†F_ : DagFunctor C D → DagFunctor C E → DagFunctor C (D ׆ E)
(F ,†F G) .fst = F .fst ,F G .fst
(F ,†F G) .snd .F-† f = ≡-× (F .snd .F-† f) (G .snd .F-† f)

_׆F_ : DagFunctor B D → DagFunctor C E → DagFunctor (B ׆ C) (D ׆ E)
_׆F_ {B = B} {C = C} F G = (F ∘†F †Fst B C) ,†F (G ∘†F †Snd B C)

†Δ : DagFunctor C (C ׆ C)
†Δ = †Id ,†F †Id

module _ (C : †Category ℓ ℓ') (D : †Category ℓ'' ℓ''') where
†Swap : DagFunctor (C ׆ D) (D ׆ C)
†Swap = †Snd C D ,†F †Fst C D

†Linj : ob D → DagFunctor C (C ׆ D)
†Linj d = †Id ,†F †Constant d

†Rinj : ob C → DagFunctor D (C ׆ D)
†Rinj c = †Constant c ,†F †Id

open areInv
open †Morphisms

†CatIso× : ∀ {x y z w} → †CatIso C x y → †CatIso D z w → †CatIso (C ׆ D) (x , z) (y , w)
†CatIso× (f , fiso) (g , giso) .fst = f , g
†CatIso× (f , fiso) (g , giso) .snd .sec = ≡-× (fiso .sec) (giso .sec)
†CatIso× (f , fiso) (g , giso) .snd .ret = ≡-× (fiso .ret) (giso .ret)
34 changes: 34 additions & 0 deletions Cubical/Categories/Dagger/Instances/Functors.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,34 @@
{-# OPTIONS --safe #-}

module Cubical.Categories.Dagger.Instances.Functors where

open import Cubical.Foundations.Prelude
open import Cubical.Data.Sigma

open import Cubical.Categories.Category
open import Cubical.Categories.Functor
open import Cubical.Categories.NaturalTransformation
open import Cubical.Categories.Constructions.FullSubcategory
open import Cubical.Categories.Instances.Functors

open import Cubical.Categories.Dagger.Base
open import Cubical.Categories.Dagger.Properties
open import Cubical.Categories.Dagger.Functor

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

module _ (C : †Category ℓC ℓC') (D : †Category ℓD ℓD') where

open Category
open †Category
open DaggerStr
open IsDagger
open NatTrans

†FUNCTOR : †Category (ℓ-max (ℓ-max ℓC ℓC') (ℓ-max ℓD ℓD')) (ℓ-max (ℓ-max ℓC ℓC') ℓD')
†FUNCTOR .cat = FullSubcategory (FUNCTOR (C .cat) (D .cat)) (IsDagFunctor C D)
†FUNCTOR .dagstr ._† {x = F} {y = G} = NT† F G
†FUNCTOR .dagstr .is-dag .†-invol n = makeNatTransPath (funExt λ x → D .†-invol (n ⟦ x ⟧))
†FUNCTOR .dagstr .is-dag .†-id = makeNatTransPath (funExt λ x → D .†-id)
†FUNCTOR .dagstr .is-dag .†-seq m n = makeNatTransPath (funExt λ x → D .†-seq (m ⟦ x ⟧) (n ⟦ x ⟧))
Loading