cat/src/Cat/Categories/Fam.agda

63 lines
2.1 KiB
Agda
Raw Normal View History

2018-01-31 14:15:00 +00:00
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Categories.Fam where
2018-03-21 13:47:01 +00:00
open import Cat.Prelude
2018-01-31 14:15:00 +00:00
import Function
open import Cat.Category
module _ (a b : Level) where
2018-01-31 14:15:00 +00:00
private
2018-04-05 08:41:56 +00:00
Object = Σ[ hA hSet a ] (fst hA hSet b)
2018-02-23 12:36:54 +00:00
Arr : Object Object Set (a b)
2018-04-05 08:41:56 +00:00
Arr ((A , _) , B) ((A' , _) , B') = Σ[ f (A A') ] ({x : A} fst (B x) fst (B' (f x)))
identity : {A : Object} Arr A A
2018-04-05 08:41:56 +00:00
fst identity = λ x x
snd identity = λ b b
2018-02-23 12:36:54 +00:00
_∘_ : {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
; identity = λ { {A} identity {A = A}}
; _∘_ = λ {a b c} _∘_ {a} {b} {c}
}
open RawCategory RawFam hiding (Object ; identity)
2018-02-23 12:36:54 +00:00
isAssociative : IsAssociative
isAssociative = Σ≡ refl refl
isIdentity : IsIdentity λ { {A} identity {A} }
2018-02-23 12:36:54 +00:00
isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
2018-04-05 12:37:25 +00:00
isPreCategory : IsPreCategory RawFam
IsPreCategory.isAssociative isPreCategory
{A} {B} {C} {D} {f} {g} {h} = isAssociative {A} {B} {C} {D} {f} {g} {h}
IsPreCategory.isIdentity isPreCategory
{A} {B} {f} = isIdentity {A} {B} {f = f}
IsPreCategory.arrowsAreSets isPreCategory
{(A , hA) , famA} {(B , hB) , famB}
= setSig
{sA = setPi λ _ hB}
{sB = λ f
let
helpr : isSet ((a : A) fst (famA a) fst (famB (f a)))
helpr = setPi λ a setPi λ _ snd (famB (f a))
-- It's almost like above, but where the first argument is
-- implicit.
res : isSet ({a : A} fst (famA a) fst (famB (f a)))
res = {!!}
in res
2018-01-31 14:15:00 +00:00
}
2018-04-05 12:37:25 +00:00
isCategory : IsCategory RawFam
IsCategory.isPreCategory isCategory = isPreCategory
IsCategory.univalent isCategory = {!!}
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
2018-04-05 12:37:25 +00:00
Category.isCategory Fam = isCategory