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
2018-02-19 10:09:49 +00:00
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
2018-02-19 10:09:49 +00:00
private
module = Category
open Category
2018-02-23 11:43:49 +00:00
p-isAssociative : {A B C D : Object} {r : Path Arrow A B} {q : Path Arrow B C} {p : Path Arrow C D}
2018-02-19 10:09:49 +00:00
p ++ (q ++ r) (p ++ q) ++ r
2018-02-23 11:43:49 +00:00
p-isAssociative {r = r} {q} {empty} = refl
p-isAssociative {A} {B} {C} {D} {r = r} {q} {cons x p} = begin
2018-02-19 10:09:49 +00:00
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)
2018-02-23 11:43:49 +00:00
lem = p-isAssociative {r = r} {q} {p}
2018-02-06 10:27:22 +00:00
2018-02-19 10:09:49 +00:00
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
2017-11-10 15:00:00 +00:00
2018-02-19 10:09:49 +00:00
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 = {!!}
2017-11-10 15:00:00 +00:00
2018-02-06 10:27:22 +00:00
RawFree : RawCategory ( ')
RawFree = record
2018-02-19 10:09:49 +00:00
{ Object = Object
; Arrow = Path Arrow
; 𝟙 = empty
; _∘_ = concatenate
}
RawIsCategoryFree : IsCategory RawFree
RawIsCategoryFree = record
2018-02-23 11:43:49 +00:00
{ isAssociative = λ { {f = f} {g} {h} p-isAssociative {r = f} {g} {h}}
2018-02-23 11:49:41 +00:00
; isIdentity = ident-r , ident-l
2018-02-23 11:51:44 +00:00
; arrowsAreSets = {!!}
; univalent = {!!}
2017-11-10 15:00:00 +00:00
}