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

Categories of propositions and sieves #1163

Draft
wants to merge 2 commits into
base: master
Choose a base branch
from
Draft
Show file tree
Hide file tree
Changes from all 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/Instances/Functors.agda
Original file line number Diff line number Diff line change
Expand Up @@ -23,6 +23,7 @@ open import Cubical.Categories.Functor.Base
open import Cubical.Categories.Morphism
open import Cubical.Categories.NaturalTransformation.Base
open import Cubical.Categories.NaturalTransformation.Properties
open import Cubical.Categories.Thin

private
variable
Expand Down Expand Up @@ -138,3 +139,5 @@ module _ (C : Category ℓC ℓC') (D : Category ℓD ℓD') where
≡⟨ D .⋆Assoc _ _ _ ⟩
(β .N-ob e ∘⟨ D ⟩ G .F-hom g) ∘⟨ D ⟩ (α .N-ob d ∘⟨ D ⟩ F .F-hom f) ∎

isThinFUNCTOR : isThin D → isThin FUNCTOR
isThinFUNCTOR thinD F G = isThin→isPropNatTrans thinD
28 changes: 28 additions & 0 deletions Cubical/Categories/Instances/Propositions.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,28 @@
{-# OPTIONS --safe #-}

module Cubical.Categories.Instances.Propositions where

open import Cubical.Foundations.Prelude
open import Cubical.Foundations.HLevels
open import Cubical.Foundations.Function

open import Cubical.Categories.Category
open Category
open import Cubical.Categories.Thin

private
variable
ℓ : Level

PROP : ∀ (ℓ : Level) → Category (ℓ-suc ℓ) ℓ
ob (PROP ℓ) = hProp ℓ
Hom[_,_] (PROP ℓ) (P , _) (Q , _) = P → Q
id (PROP ℓ) = idfun _
_⋆_ (PROP ℓ) f g x = g (f x)
⋆IdL (PROP ℓ) {P , _} {Q , propQ} f = isProp→ propQ _ _
⋆IdR (PROP ℓ) {P , _} {Q , propQ} f = isProp→ propQ _ _
⋆Assoc (PROP ℓ) {P , _} {Q , _} {R , _} {S , propS} f g h = isPropΠ (λ _ → propS) _ _
isSetHom (PROP ℓ) {P , _} {Q , propQ} = isProp→isSet (isProp→ propQ)

isThinPROP : isThin (PROP ℓ)
isThinPROP (P , propP) (Q , propQ) = isProp→ propQ
37 changes: 37 additions & 0 deletions Cubical/Categories/Instances/SievesAsFunctors.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,37 @@
{-# OPTIONS --safe #-}

module Cubical.Categories.Instances.SievesAsFunctors where

-- There are also sieves in Cubical.Site.Sieve
Copy link
Collaborator

Choose a reason for hiding this comment

The reason will be displayed to describe this comment to others. Learn more.

I think it would be good to also have a pointer there.
Maybe it would also be good to mention that the definitions are equivalent.


open import Cubical.Foundations.Prelude
open import Cubical.Foundations.HLevels

open import Cubical.Categories.Category
open Category
open import Cubical.Categories.Instances.Functors
open import Cubical.Categories.Instances.Propositions
open import Cubical.Categories.Constructions.Slice
open import Cubical.Categories.Thin

private
variable
ℓ ℓ' ℓP : Level
C : Category ℓ ℓ'

SIEVE : (C : Category ℓ ℓ') → (ℓP : Level)
→ Category (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓP)) (ℓ-max (ℓ-max ℓ ℓ') ℓP)
SIEVE C ℓP = FUNCTOR (C ^op) (PROP ℓP)

Sieve : (C : Category ℓ ℓ') → (ℓP : Level) → Type (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓP))
Sieve C ℓP = SIEVE C ℓP .ob

isThinSIEVE : isThin (SIEVE C ℓP)
isThinSIEVE = isThinFUNCTOR _ _ isThinPROP

SIEVE_ON : (C : Category ℓ ℓ') → (c : C .ob) → (ℓP : Level)
→ Category (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓP)) (ℓ-max (ℓ-max ℓ ℓ') ℓP)
SIEVE_ON C c ℓP = SIEVE (SliceCat C c) ℓP

SieveOn : (C : Category ℓ ℓ') → (c : C .ob) → (ℓP : Level) → Type (ℓ-max (ℓ-max ℓ ℓ') (ℓ-suc ℓP))
SieveOn C c ℓP = SIEVE_ON C c ℓP .ob
3 changes: 3 additions & 0 deletions Cubical/Categories/NaturalTransformation/Properties.agda
Original file line number Diff line number Diff line change
Expand Up @@ -16,6 +16,7 @@ open import Cubical.Categories.Functor.Properties
open import Cubical.Categories.Morphism
open import Cubical.Categories.Isomorphism
open import Cubical.Categories.NaturalTransformation.Base
open import Cubical.Categories.Thin

private
variable
Expand Down Expand Up @@ -193,3 +194,5 @@ inv congNatIso^opFiso = _
rightInv congNatIso^opFiso _ = refl
leftInv congNatIso^opFiso _ = refl

isThin→isPropNatTrans : ∀ {F G : Functor C D} → isThin D → isProp (NatTrans F G)
isThin→isPropNatTrans thinD μ ν = makeNatTransPath (isPropΠ (λ x → thinD _ _) _ _)
12 changes: 12 additions & 0 deletions Cubical/Categories/Thin.agda
Original file line number Diff line number Diff line change
@@ -0,0 +1,12 @@
{-# OPTIONS --safe #-}
module Cubical.Categories.Thin where

open import Cubical.Foundations.Prelude
open import Cubical.Categories.Category
open Category

private
variable ℓ ℓ' : Level

isThin : (C : Category ℓ ℓ') → Type (ℓ-max ℓ ℓ')
isThin C = (x y : C .ob) → isProp (C [ x , y ])
Loading