cat/src/Cat/Categories/Cube.agda

81 lines
2.1 KiB
Agda
Raw Normal View History

{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Categories.Cube where
open import Level
open import Data.Bool hiding (T)
open import Data.Sum hiding ([_,_])
open import Data.Unit
open import Data.Empty
open import Data.Product
open import Cubical
open import Function
open import Relation.Nullary
open import Relation.Nullary.Decidable
open import Cat.Category
2018-02-05 13:59:53 +00:00
open import Cat.Category.Functor
open import Cat.Equality
open Equality.Data.Product
-- See chapter 1 for a discussion on how presheaf categories are CwF's.
-- See section 6.8 in Huber's thesis for details on how to implement the
-- categorical version of CTT
open Category hiding (_∘_)
open Functor
module _ { ' : Level} (Ns : Set ) where
-- Ns is the "namespace"
o = (suc zero )
FiniteDecidableSubset : Set
FiniteDecidableSubset = Ns Dec
isTrue : Bool Set
isTrue false =
isTrue true =
elmsof : FiniteDecidableSubset Set
elmsof P = Σ Ns (λ σ True (P σ)) -- (σ : Ns) → isTrue (P σ)
𝟚 : Set
𝟚 = Bool
module _ (I J : FiniteDecidableSubset) where
private
Hom' : Set
Hom' = elmsof I elmsof J 𝟚
isInl : {a b : Level} {A : Set a} {B : Set b} A B Set
isInl (inj₁ _) =
isInl (inj₂ _) =
Def : Set
Def = (f : Hom') Σ (elmsof I) (λ i isInl (f i))
rules : Hom' Set
rules f = (i j : elmsof I)
case (f i) of λ
{ (inj₁ (fi , _)) case (f j) of λ
{ (inj₁ (fj , _)) fi fj i j
; (inj₂ _) Lift
}
; (inj₂ _) Lift
}
Hom = Σ Hom' rules
2018-02-05 13:47:15 +00:00
module Raw = RawCategory
-- The category of names and substitutions
Raw : RawCategory -- o (lsuc lzero ⊔ o)
2018-02-05 13:47:15 +00:00
Raw.Object Raw = FiniteDecidableSubset
Raw.Arrow Raw = Hom
Raw.𝟙 Raw {o} = inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} }
Raw._∘_ Raw = {!!}
postulate IsCategory : IsCategory Raw
: Category
2018-02-05 13:47:15 +00:00
raw = Raw
isCategory = IsCategory