cat/src/Cat/Categories/Free.agda

63 lines
2.1 KiB
Agda
Raw Normal View History

2018-02-02 14:33:54 +00:00
{-# OPTIONS --allow-unsolved-metas #-}
2018-02-06 09:35:52 +00:00
module Cat.Categories.Free where
2017-11-10 15:00:00 +00:00
open import Agda.Primitive
2018-02-06 10:27:22 +00:00
open import Cubical hiding (Path ; isSet ; empty)
open import Data.Product
2017-11-10 15:00:00 +00:00
2018-02-06 10:27:22 +00:00
open import Cat.Category
open IsCategory
open Category
-- data Path { : Level} {A : Set } : (a b : A) → Set where
-- emptyPath : {a : A} → Path a a
-- concatenate : {a b c : A} → Path a b → Path b c → Path a b
2017-11-10 15:00:00 +00:00
module _ { ' : Level} ( : Category ') where
2018-02-06 10:27:22 +00:00
module = Category
-- import Data.List
-- P : (a b : Object ) → Set (')
-- P = {!Data.List.List ?!}
-- Generalized paths:
-- data P { : Level} {A : Set } (R : A → A → Set ) : (a b : A) → Set where
-- e : {a : A} → P R a a
-- c : {a b c : A} → R a b → P R b c → P R a c
-- Path's are like lists with directions.
-- This implementation is specialized to categories.
data Path : (a b : Object ) Set ( ') where
empty : {A : Object } Path A A
cons : {A B C} [ B , C ] Path A B Path A C
2017-11-10 15:00:00 +00:00
2018-02-06 10:27:22 +00:00
concatenate : {A B C : Object } Path B C Path A B Path A C
concatenate empty p = p
concatenate (cons x q) p = cons x (concatenate q p)
2017-11-10 15:00:00 +00:00
private
2018-02-06 10:27:22 +00:00
module _ {A B C D : Object } where
p-assoc : {r : Path A B} {q : Path B C} {p : Path C D} concatenate p (concatenate q r) concatenate (concatenate p q) r
p-assoc {r} {q} {p} = {!!}
module _ {A B : Object } {p : Path A B} where
-- postulate
-- ident-r : concatenate {A} {A} {B} p (lift 𝟙) ≡ p
-- ident-l : concatenate {A} {B} {B} (lift 𝟙) p ≡ p
module _ {A B : Object } where
isSet : Cubical.isSet (Path A B)
2018-02-06 10:27:22 +00:00
isSet = {!!}
RawFree : RawCategory ( ')
RawFree = record
2018-02-06 10:27:22 +00:00
{ Object = Object
2017-11-10 15:00:00 +00:00
; Arrow = Path
2018-02-06 10:27:22 +00:00
; 𝟙 = λ {o} {!lift 𝟙!}
; _∘_ = λ {a b c} {!concatenate {a} {b} {c}!}
}
RawIsCategoryFree : IsCategory RawFree
RawIsCategoryFree = record
2018-02-06 10:27:22 +00:00
{ assoc = {!p-assoc!}
; ident = {!ident-r , ident-l!}
2018-02-06 09:34:43 +00:00
; arrowIsSet = {!!}
; univalent = {!!}
2017-11-10 15:00:00 +00:00
}