Use co-patterns

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-05 14:47:15 +01:00
parent 20dc9d26ac
commit 83ccde62e9
5 changed files with 77 additions and 64 deletions

View file

@ -65,15 +65,16 @@ module _ { ' : Level} (Ns : Set ) where
Hom = Σ Hom' rules Hom = Σ Hom' rules
module Raw = RawCategory
-- The category of names and substitutions -- The category of names and substitutions
Raw : RawCategory -- o (lsuc lzero ⊔ o) Raw : RawCategory -- o (lsuc lzero ⊔ o)
Raw = record Raw.Object Raw = FiniteDecidableSubset
{ Object = FiniteDecidableSubset Raw.Arrow Raw = Hom
-- { Object = Ns → Bool Raw.𝟙 Raw {o} = inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} }
; Arrow = Hom Raw._∘_ Raw = {!!}
; 𝟙 = λ { {o} inj₁ , λ { (i , ii) (j , jj) eq Σ≡ eq {!refl!} } }
; _∘_ = {!!} postulate IsCategory : IsCategory Raw
}
postulate RawIsCategory : IsCategory Raw
: Category : Category
= Raw , RawIsCategory raw = Raw
isCategory = IsCategory

View file

@ -46,9 +46,9 @@ module _ (a b : Level) where
isCategory = record isCategory = record
{ assoc = λ {A} {B} {C} {D} {f} {g} {h} assoc {D = D} {f} {g} {h} { assoc = λ {A} {B} {C} {D} {f} {g} {h} assoc {D = D} {f} {g} {h}
; ident = λ {A} {B} {f} ident {A} {B} {f = f} ; ident = λ {A} {B} {f} ident {A} {B} {f = f}
; arrow-is-set = ? ; arrow-is-set = {!!}
; univalent = ? ; univalent = {!!}
} }
Fam : Category (lsuc (a b)) (a b) Fam : Category (lsuc (a b)) (a b)
Fam = RawFam , isCategory Category.raw Fam = RawFam

View file

@ -110,12 +110,12 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
:isCategory: = record :isCategory: = record
{ assoc = λ {A B C D} :assoc: {A} {B} {C} {D} { assoc = λ {A B C D} :assoc: {A} {B} {C} {D}
; ident = λ {A B} :ident: {A} {B} ; ident = λ {A B} :ident: {A} {B}
; arrow-is-set = ? ; arrow-is-set = {!!}
; univalent = ? ; univalent = {!!}
} }
Fun : Category (c c' d d') (c c' d') Fun : Category (c c' d d') (c c' d')
Fun = RawFun , :isCategory: raw Fun = RawFun
module _ { ' : Level} ( : Category ') where module _ { ' : Level} ( : Category ') where
open import Cat.Categories.Sets open import Cat.Categories.Sets

View file

@ -1,4 +1,4 @@
{-# OPTIONS --allow-unsolved-metas #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Categories.Sets where module Cat.Categories.Sets where
open import Cubical open import Cubical
@ -13,23 +13,22 @@ open Category
module _ { : Level} where module _ { : Level} where
SetsRaw : RawCategory (lsuc ) SetsRaw : RawCategory (lsuc )
SetsRaw = record RawCategory.Object SetsRaw = Set
{ Object = Set RawCategory.Arrow SetsRaw = λ T U T U
; Arrow = λ T U T U RawCategory.𝟙 SetsRaw = Function.id
; 𝟙 = Function.id RawCategory._∘_ SetsRaw = Function._∘_
; _∘_ = Function._∘_
}
open IsCategory
SetsIsCategory : IsCategory SetsRaw SetsIsCategory : IsCategory SetsRaw
SetsIsCategory = record assoc SetsIsCategory = refl
{ assoc = refl proj₁ (ident SetsIsCategory) = funExt λ _ refl
; ident = funExt (λ _ refl) , funExt (λ _ refl) proj₂ (ident SetsIsCategory) = funExt λ _ refl
; arrow-is-set = {!!} arrow-is-set SetsIsCategory = {!!}
; univalent = {!!} univalent SetsIsCategory = {!!}
}
Sets : Category (lsuc ) Sets : Category (lsuc )
Sets = SetsRaw , SetsIsCategory raw Sets = SetsRaw
isCategory Sets = SetsIsCategory
private private
module _ {X A B : Set } (f : X A) (g : X B) where module _ {X A B : Set } (f : X A) (g : X B) where

View file

@ -109,12 +109,10 @@ module _ {a} {b} { : RawCategory a b} where
module x = IsCategory x module x = IsCategory x
module y = IsCategory y module y = IsCategory y
Category : (a b : Level) Set (lsuc (a b)) record Category (a b : Level) : Set (lsuc (a b)) where
Category a b = Σ (RawCategory a b) IsCategory field
raw : RawCategory a b
module Category {a b : Level} ( : Category a b) where {{isCategory}} : IsCategory raw
raw = fst
isCategory = snd
private private
module = RawCategory raw module = RawCategory raw
@ -134,42 +132,57 @@ module Category {a b : Level} ( : Category a b) where
_[_∘_] : {A B C : Object} (g : .Arrow B C) (f : .Arrow A B) .Arrow A C _[_∘_] : {A B C : Object} (g : .Arrow B C) (f : .Arrow A B) .Arrow A C
_[_∘_] = ._∘_ _[_∘_] = ._∘_
open Category using ( Object ; _[_,_] ; _[_∘_])
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
private private
open Category open Category
module = RawCategory ( .fst)
OpRaw : RawCategory a b
OpRaw = record
{ Object = .Object
; Arrow = Function.flip .Arrow
; 𝟙 = .𝟙
; _∘_ = Function.flip (._∘_)
}
open IsCategory isCategory
OpIsCategory : IsCategory OpRaw
OpIsCategory = record
{ assoc = sym assoc
; ident = swap ident
; arrow-is-set = {!!}
; univalent = {!!}
}
Opposite : Category a b
Opposite = OpRaw , OpIsCategory
-- A consequence of no-eta-equality; `Opposite-is-involution` is no longer OpRaw : RawCategory a b
-- definitional - i.e.; you must match on the fields: RawCategory.Object OpRaw = Object
-- RawCategory.Arrow OpRaw = Function.flip Arrow
-- Opposite-is-involution : ∀ { '} → {C : Category {} {'}} → Opposite (Opposite C) ≡ C RawCategory.𝟙 OpRaw = 𝟙
-- Object (Opposite-is-involution {C = C} i) = Object C RawCategory._∘_ OpRaw = Function.flip _∘_
-- Arrow (Opposite-is-involution i) = {!!}
-- 𝟙 (Opposite-is-involution i) = {!!} open IsCategory isCategory
-- _⊕_ (Opposite-is-involution i) = {!!}
-- assoc (Opposite-is-involution i) = {!!} OpIsCategory : IsCategory OpRaw
-- ident (Opposite-is-involution i) = {!!} IsCategory.assoc OpIsCategory = sym assoc
IsCategory.ident OpIsCategory = swap ident
IsCategory.arrow-is-set OpIsCategory = {!!}
IsCategory.univalent OpIsCategory = {!!}
Opposite : Category a b
raw Opposite = OpRaw
Category.isCategory Opposite = OpIsCategory
-- As demonstrated here a side-effect of having no-eta-equality on constructors
-- means that we need to pick things apart to show that things are indeed
-- definitionally equal. I.e; a thing that would normally be provable in one
-- line now takes more than 20!!
module _ {a b : Level} { : Category a b} where
private
open RawCategory
module C = Category
rawOp : Category.raw (Opposite (Opposite )) Category.raw
Object (rawOp _) = C.Object
Arrow (rawOp _) = C.Arrow
𝟙 (rawOp _) = C.𝟙
_∘_ (rawOp _) = C._∘_
open Category
open IsCategory
module IsCat = IsCategory ( .isCategory)
rawIsCat : (i : I) IsCategory (rawOp i)
assoc (rawIsCat i) = IsCat.assoc
ident (rawIsCat i) = IsCat.ident
arrow-is-set (rawIsCat i) = IsCat.arrow-is-set
univalent (rawIsCat i) = IsCat.univalent
Opposite-is-involution : Opposite (Opposite )
raw (Opposite-is-involution i) = rawOp i
isCategory (Opposite-is-involution i) = rawIsCat i
module _ {a b : Level} ( : Category a b) where module _ {a b : Level} ( : Category a b) where
open Category
unique = isContr unique = isContr
IsInitial : Object Set (a b) IsInitial : Object Set (a b)