cat/src/Cat/Categories/Fam.agda

63 lines
2.1 KiB
Agda
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Categories.Fam where
open import Cat.Prelude
import Function
open import Cat.Category
module _ (a b : Level) where
private
Object = Σ[ hA hSet a ] (fst hA hSet b)
Arr : Object Object Set (a b)
Arr ((A , _) , B) ((A' , _) , B') = Σ[ f (A A') ] ({x : A} fst (B x) fst (B' (f x)))
identity : {A : Object} Arr A A
fst identity = λ x x
snd identity = λ b b
_∘_ : {a b c : Object} Arr b c Arr a b Arr a c
(g , g') (f , f') = g Function.∘ f , g' Function.∘ f'
RawFam : RawCategory (lsuc (a b)) (a b)
RawFam = record
{ Object = Object
; Arrow = Arr
; identity = λ { {A} identity {A = A}}
; _∘_ = λ {a b c} _∘_ {a} {b} {c}
}
open RawCategory RawFam hiding (Object ; identity)
isAssociative : IsAssociative
isAssociative = Σ≡ refl refl
isIdentity : IsIdentity λ { {A} identity {A} }
isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
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
}
isCategory : IsCategory RawFam
IsCategory.isPreCategory isCategory = isPreCategory
IsCategory.univalent isCategory = {!!}
Fam : Category (lsuc (a b)) (a b)
Category.raw Fam = RawFam
Category.isCategory Fam = isCategory