Restructure in free monad

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-20 14:58:27 +01:00
parent 66cb5b363d
commit b03bfb0c77
3 changed files with 55 additions and 39 deletions

View file

@ -2,61 +2,77 @@
module Cat.Categories.Free where module Cat.Categories.Free where
open import Agda.Primitive open import Agda.Primitive
open import Cubical hiding (Path ; isSet ; empty) open import Relation.Binary
open import Cubical hiding (Path ; empty)
open import Data.Product open import Data.Product
open import Cat.Category open import Cat.Category
data Path { ' : Level} {A : Set } (R : A A Set ') : (a b : A) Set ( ') where module _ { : Level} {A : Set } {r : Level} where
empty : {a : A} Path R a a data Path (R : Rel A r) : (a b : A) Set ( r) where
cons : {a b c : A} R b c Path R a b Path R a c 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 module _ {R : Rel A r} where
concatenate empty p = p concatenate : {a b c : A} Path R b c Path R a b Path R a c
concatenate (cons x q) p = cons x (concatenate q p) concatenate empty p = p
_++_ = concatenate concatenate (cons x q) p = cons x (concatenate q p)
_++_ : {a b c : A} Path R b c Path R a b Path R a c
a ++ b = concatenate a b
singleton : {} {𝓤 : Set } {r} {R : 𝓤 𝓤 Set r} {A B : 𝓤} R A B Path R A B singleton : {a b : A} R a b Path R a b
singleton f = cons f empty singleton f = cons f empty
module _ { ' : Level} ( : Category ') where module _ {a b : Level} ( : Category a b) where
private private
module = Category module = Category
open Category
p-isAssociative : {A B C D : Object} {r : Path Arrow A B} {q : Path Arrow B C} {p : Path Arrow C D} RawFree : RawCategory a (a b)
RawCategory.Object RawFree = .Object
RawCategory.Arrow RawFree = Path .Arrow
RawCategory.𝟙 RawFree = empty
RawCategory._∘_ RawFree = concatenate
open RawCategory RawFree
open Univalence RawFree
isAssociative : {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 ++ (q ++ r) (p ++ q) ++ r
p-isAssociative {r = r} {q} {empty} = refl isAssociative {r = r} {q} {empty} = refl
p-isAssociative {A} {B} {C} {D} {r = r} {q} {cons x p} = begin isAssociative {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) ≡⟨ cong (cons x) lem
cons x ((p ++ q) ++ r) ≡⟨⟩ cons x ((p ++ q) ++ r) ≡⟨⟩
(cons x p ++ q) ++ r (cons x p ++ q) ++ r
where where
lem : p ++ (q ++ r) ((p ++ q) ++ r) lem : p ++ (q ++ r) ((p ++ q) ++ r)
lem = p-isAssociative {r = r} {q} {p} lem = isAssociative {r = r} {q} {p}
ident-r : {A} {B} {p : Path Arrow A B} concatenate p empty p ident-r : {A} {B} {p : Path .Arrow A B} concatenate p empty p
ident-r {p = empty} = refl ident-r {p = empty} = refl
ident-r {p = cons x p} = cong (cons x) ident-r 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 : {A} {B} {p : Path .Arrow A B} concatenate empty p p
ident-l = refl ident-l = refl
module _ {A B : Object} where isIdentity : IsIdentity 𝟙
isSet : Cubical.isSet (Path Arrow A B) isIdentity = ident-r , ident-l
isSet a b p q = {!!}
RawFree : RawCategory ( ') module _ {A B : .Object} where
RawFree = record arrowsAreSets : Cubical.isSet (Path .Arrow A B)
{ Object = Object arrowsAreSets a b p q = {!!}
; Arrow = Path Arrow
; 𝟙 = empty eqv : isEquiv (A B) (A B) (id-to-iso isIdentity A B)
; _∘_ = concatenate eqv = {!!}
} univalent : Univalent isIdentity
RawIsCategoryFree : IsCategory RawFree univalent = eqv
RawIsCategoryFree = record
{ isAssociative = λ { {f = f} {g} {h} p-isAssociative {r = f} {g} {h}} isCategory : IsCategory RawFree
; isIdentity = ident-r , ident-l IsCategory.isAssociative isCategory {f = f} {g} {h} = isAssociative {r = f} {g} {h}
; arrowsAreSets = {!!} IsCategory.isIdentity isCategory = isIdentity
; univalent = {!!} IsCategory.arrowsAreSets isCategory = arrowsAreSets
} IsCategory.univalent isCategory = univalent
Free : Category _ _
Free = record { raw = RawFree ; isCategory = isCategory }

View file

@ -10,6 +10,7 @@ open import Function using (_∘_)
open import Cubical hiding (_≃_) open import Cubical hiding (_≃_)
open import Cubical.Univalence using (univalence ; con ; _≃_ ; idtoeqv ; ua) open import Cubical.Univalence using (univalence ; con ; _≃_ ; idtoeqv ; ua)
open import Cubical.GradLemma open import Cubical.GradLemma
open import Cubical.NType.Properties
open import Cat.Category open import Cat.Category
open import Cat.Category.Functor open import Cat.Category.Functor
@ -59,8 +60,7 @@ module _ { : Level} {A B : Set } {a : A} where
module _ ( : Level) where module _ ( : Level) where
private private
open import Cubical.NType.Properties open import Cubical.Universe using (hSet) public
open import Cubical.Universe
SetsRaw : RawCategory (lsuc ) SetsRaw : RawCategory (lsuc )
RawCategory.Object SetsRaw = hSet {} RawCategory.Object SetsRaw = hSet {}

View file

@ -19,7 +19,7 @@ module _ (a b : Level) where
-- --
-- Since it doesn't we'll make the following (definitionally equivalent) ad-hoc definition. -- Since it doesn't we'll make the following (definitionally equivalent) ad-hoc definition.
_×_ : {a b} Category a b Category a b Category a b _×_ : {a b} Category a b Category a b Category a b
× 𝔻 = Cat.CatProduct.obj 𝔻 × 𝔻 = Cat.CatProduct.object 𝔻
record RawMonoidalCategory : Set where record RawMonoidalCategory : Set where
field field