cat/src/Cat/Categories/Free.agda

73 lines
2.4 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.Free where
open import Agda.Primitive
open import Cubical hiding (Path ; isSet ; empty)
open import Data.Product
open import Cat.Category
open IsCategory
-- 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
-- import Data.List
-- P : (a b : Object ) → Set (')
-- P = {!Data.List.List ?!}
-- Generalized paths:
data Path { ' : Level} {A : Set } (R : A A Set ') : (a b : A) Set ( ') where
empty : {a : A} Path R a a
cons : {a b c : A} R b c Path R a b Path R a c
concatenate _++_ : { '} {A : Set } {a b c : A} {R : A A Set '} Path R b c Path R a b Path R a c
concatenate empty p = p
concatenate (cons x q) p = cons x (concatenate q p)
_++_ = concatenate
singleton : {} {𝓤 : Set } {r} {R : 𝓤 𝓤 Set r} {A B : 𝓤} R A B Path R A B
singleton f = cons f empty
module _ { ' : Level} ( : Category ') where
module = Category
open Category
private
p-assoc : {A B C D : Object} {r : Path Arrow A B} {q : Path Arrow B C} {p : Path Arrow C D}
p ++ (q ++ r) (p ++ q) ++ r
p-assoc {r = r} {q} {empty} = refl
p-assoc {A} {B} {C} {D} {r = r} {q} {cons x p} = begin
cons x p ++ (q ++ r) ≡⟨ cong (cons x) lem
cons x ((p ++ q) ++ r) ≡⟨⟩
(cons x p ++ q) ++ r
where
lem : p ++ (q ++ r) ((p ++ q) ++ r)
lem = p-assoc {r = r} {q} {p}
ident-r : {A} {B} {p : Path Arrow A B} concatenate p empty p
ident-r {p = empty} = refl
ident-r {p = cons x p} = cong (cons x) ident-r
ident-l : {A} {B} {p : Path Arrow A B} concatenate empty p p
ident-l = refl
module _ {A B : Object} where
isSet : Cubical.isSet (Path Arrow A B)
isSet a b p q = {!!}
RawFree : RawCategory ( ')
RawFree = record
{ Object = Object
; Arrow = Path Arrow
; 𝟙 = empty
; _∘_ = concatenate
}
RawIsCategoryFree : IsCategory RawFree
RawIsCategoryFree = record
{ assoc = λ { {f = f} {g} {h} p-assoc {r = f} {g} {h}}
; ident = ident-r , ident-l
; arrowIsSet = {!!}
; univalent = {!!}
}