cat/src/Cat/Categories/Fam.agda

71 lines
2.3 KiB
Agda
Raw Normal View History

2018-01-31 14:15:00 +00:00
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Categories.Fam where
open import Agda.Primitive
open import Data.Product
import Function
2018-02-23 12:36:54 +00:00
open import Cubical
open import Cubical.Universe
2018-01-31 14:15:00 +00:00
open import Cat.Category
open import Cat.Equality
open Equality.Data.Product
module _ (a b : Level) where
2018-01-31 14:15:00 +00:00
private
2018-02-23 12:36:54 +00:00
Object = Σ[ hA hSet {a} ] (proj₁ hA hSet {b})
Arr : Object Object Set (a b)
Arr ((A , _) , B) ((A' , _) , B') = Σ[ f (A A') ] ({x : A} proj₁ (B x) proj₁ (B' (f x)))
𝟙 : {A : Object} Arr A A
proj₁ 𝟙 = λ x x
proj₂ 𝟙 = λ b b
_∘_ : {a b c : Object} Arr b c Arr a b Arr a c
2018-01-31 14:15:00 +00:00
(g , g') (f , f') = g Function.∘ f , g' Function.∘ f'
RawFam : RawCategory (lsuc (a b)) (a b)
RawFam = record
2018-02-23 12:36:54 +00:00
{ Object = Object
; Arrow = Arr
2018-02-23 12:36:54 +00:00
; 𝟙 = λ { {A} 𝟙 {A = A}}
; _∘_ = λ {a b c} _∘_ {a} {b} {c}
}
2018-02-23 12:36:54 +00:00
open RawCategory RawFam hiding (Object ; 𝟙)
isAssociative : IsAssociative
isAssociative = Σ≡ refl refl
isIdentity : IsIdentity λ { {A} 𝟙 {A} }
isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
open import Cubical.NType.Properties
open import Cubical.Sigma
2018-01-31 14:15:00 +00:00
instance
isCategory : IsCategory RawFam
2018-01-31 14:15:00 +00:00
isCategory = record
2018-02-23 12:36:54 +00:00
{ isAssociative = λ {A} {B} {C} {D} {f} {g} {h} isAssociative {A} {B} {C} {D} {f} {g} {h}
2018-02-23 11:49:41 +00:00
; isIdentity = λ {A} {B} {f} isIdentity {A} {B} {f = f}
; arrowsAreSets = λ {
{((A , hA) , famA)}
{((B , hB) , famB)}
setSig
{sA = setPi λ _ hB}
{sB = λ f
let
helpr : isSet ((a : A) proj₁ (famA a) proj₁ (famB (f a)))
helpr = setPi λ a setPi λ _ proj₂ (famB (f a))
-- It's almost like above, but where the first argument is
-- implicit.
res : isSet ({a : A} proj₁ (famA a) proj₁ (famB (f a)))
res = {!!}
in res
}
}
2018-02-05 13:47:15 +00:00
; univalent = {!!}
2018-01-31 14:15:00 +00:00
}
Fam : Category (lsuc (a b)) (a b)
2018-02-05 13:47:15 +00:00
Category.raw Fam = RawFam