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)
|
2018-01-17 22:00:27 +00:00
|
|
|
|
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
|
|
|
|
|
|
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
|
|
|
|
|
|
2018-01-21 00:11:08 +00:00
|
|
|
|
module _ {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') where
|
2018-02-06 10:27:22 +00:00
|
|
|
|
module ℂ = Category ℂ
|
2018-02-19 10:09:49 +00:00
|
|
|
|
open Category ℂ
|
2018-02-06 10:27:22 +00:00
|
|
|
|
|
2018-02-19 10:09:49 +00:00
|
|
|
|
private
|
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 ℓ (ℓ ⊔ ℓ')
|
2018-02-05 10:43:38 +00:00
|
|
|
|
RawFree = record
|
2018-02-19 10:09:49 +00:00
|
|
|
|
{ Object = Object
|
|
|
|
|
; Arrow = Path Arrow
|
|
|
|
|
; 𝟙 = empty
|
|
|
|
|
; _∘_ = concatenate
|
2018-02-05 10:43:38 +00:00
|
|
|
|
}
|
|
|
|
|
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-06 09:34:43 +00:00
|
|
|
|
; arrowIsSet = {!!}
|
2018-02-05 10:43:38 +00:00
|
|
|
|
; univalent = {!!}
|
2017-11-10 15:00:00 +00:00
|
|
|
|
}
|