2018-02-02 14:33:54 +00:00
|
|
|
|
{-# OPTIONS --allow-unsolved-metas #-}
|
2018-01-17 22:00:27 +00:00
|
|
|
|
module Cat.Category.Free where
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
|
|
|
|
open import Agda.Primitive
|
2018-01-30 10:19:48 +00:00
|
|
|
|
open import Cubical hiding (Path)
|
2018-01-17 22:00:27 +00:00
|
|
|
|
open import Data.Product
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
2018-01-17 22:00:27 +00:00
|
|
|
|
open import Cat.Category as C
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') where
|
2017-11-10 15:00:00 +00:00
|
|
|
|
private
|
|
|
|
|
open module ℂ = Category ℂ
|
|
|
|
|
|
2018-01-21 14:03:00 +00:00
|
|
|
|
postulate
|
2018-02-05 11:21:39 +00:00
|
|
|
|
Path : (a b : ℂ.Object) → Set ℓ'
|
|
|
|
|
emptyPath : (o : ℂ.Object) → Path o o
|
|
|
|
|
concatenate : {a b c : ℂ.Object} → Path b c → Path a b → Path a c
|
2017-11-10 15:00:00 +00:00
|
|
|
|
|
|
|
|
|
private
|
2018-02-05 11:21:39 +00:00
|
|
|
|
module _ {A B C D : ℂ.Object} {r : Path A B} {q : Path B C} {p : Path C D} where
|
2017-11-10 15:00:00 +00:00
|
|
|
|
postulate
|
|
|
|
|
p-assoc : concatenate {A} {C} {D} p (concatenate {A} {B} {C} q r)
|
|
|
|
|
≡ concatenate {A} {B} {D} (concatenate {B} {C} {D} p q) r
|
2018-02-05 11:21:39 +00:00
|
|
|
|
module _ {A B : ℂ.Object} {p : Path A B} where
|
2017-11-10 15:00:00 +00:00
|
|
|
|
postulate
|
|
|
|
|
ident-r : concatenate {A} {A} {B} p (emptyPath A) ≡ p
|
|
|
|
|
ident-l : concatenate {A} {B} {B} (emptyPath B) p ≡ p
|
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
RawFree : RawCategory ℓ ℓ'
|
|
|
|
|
RawFree = record
|
2018-02-05 11:21:39 +00:00
|
|
|
|
{ Object = ℂ.Object
|
2017-11-10 15:00:00 +00:00
|
|
|
|
; Arrow = Path
|
|
|
|
|
; 𝟙 = λ {o} → emptyPath o
|
2018-01-30 18:19:16 +00:00
|
|
|
|
; _∘_ = λ {a b c} → concatenate {a} {b} {c}
|
2018-02-05 10:43:38 +00:00
|
|
|
|
}
|
|
|
|
|
RawIsCategoryFree : IsCategory RawFree
|
|
|
|
|
RawIsCategoryFree = record
|
|
|
|
|
{ assoc = p-assoc
|
|
|
|
|
; ident = ident-r , ident-l
|
2018-02-06 09:34:43 +00:00
|
|
|
|
; arrowIsSet = {!!}
|
2018-02-05 10:43:38 +00:00
|
|
|
|
; univalent = {!!}
|
2017-11-10 15:00:00 +00:00
|
|
|
|
}
|