Merge branch 'Saizan-master' into dev

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-02 15:33:54 +01:00
commit e5f1fa018a
7 changed files with 132 additions and 92 deletions

View file

@ -8,7 +8,7 @@ import Cat.Category.Bij
import Cat.Category.Free import Cat.Category.Free
import Cat.Category.Properties import Cat.Category.Properties
import Cat.Categories.Sets import Cat.Categories.Sets
import Cat.Categories.Cat -- import Cat.Categories.Cat
import Cat.Categories.Rel import Cat.Categories.Rel
import Cat.Categories.Fun import Cat.Categories.Fun
import Cat.Categories.Cube import Cat.Categories.Cube

View file

@ -1,3 +1,4 @@
-- There is no category of categories in our interpretation
{-# OPTIONS --cubical --allow-unsolved-metas #-} {-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Categories.Cat where module Cat.Categories.Cat where

View file

@ -1,4 +1,4 @@
{-# OPTIONS --cubical #-} {-# OPTIONS --cubical --allow-unsolved-metas #-}
module Cat.Categories.Rel where module Cat.Categories.Rel where
open import Cubical open import Cubical
@ -160,5 +160,10 @@ Rel = record
; Arrow = λ S R Subset (S × R) ; Arrow = λ S R Subset (S × R)
; 𝟙 = λ {S} Diag S ; 𝟙 = λ {S} Diag S
; _∘_ = λ {A B C} S R λ {( a , c ) Σ[ b B ] ( (a , b) R × (b , c) S )} ; _∘_ = λ {A B C} S R λ {( a , c ) Σ[ b B ] ( (a , b) R × (b , c) S )}
; isCategory = record { assoc = funExt is-assoc ; ident = funExt ident-l , funExt ident-r } ; isCategory = record
{ assoc = funExt is-assoc
; ident = funExt ident-l , funExt ident-r
; arrow-is-set = {!!}
; univalent = {!!}
}
} }

View file

@ -1,3 +1,4 @@
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Categories.Sets where module Cat.Categories.Sets where
open import Cubical open import Cubical

View file

@ -1,4 +1,4 @@
{-# OPTIONS --cubical #-} {-# OPTIONS --allow-unsolved-metas --cubical #-}
module Cat.Category where module Cat.Category where
@ -22,23 +22,65 @@ open import Cubical
syntax ∃!-syntax (λ x B) = ∃![ x ] B syntax ∃!-syntax (λ x B) = ∃![ x ] B
-- All projections must be `isProp`'s -- Thierry: All projections must be `isProp`'s
-- According to definitions 9.1.1 and 9.1.6 in the HoTT book the
-- arrows of a category form a set (arrow-is-set), and there is an
-- equivalence between the equality of objects and isomorphisms
-- (univalent).
record IsCategory { ' : Level} record IsCategory { ' : Level}
(Object : Set ) (Object : Set )
(Arrow : Object Object Set ') (Arrow : Object Object Set ')
(𝟙 : {o : Object} Arrow o o) (𝟙 : {o : Object} Arrow o o)
(_⊕_ : { a b c : Object } Arrow b c Arrow a b Arrow a c) (__ : { a b c : Object } Arrow b c Arrow a b Arrow a c)
: Set (lsuc (' )) where : Set (lsuc (' )) where
field field
assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D } assoc : {A B C D : Object} { f : Arrow A B } { g : Arrow B C } { h : Arrow C D }
h (g f) (h g) f h (g f) (h g) f
ident : {A B : Object} {f : Arrow A B} ident : {A B : Object} {f : Arrow A B}
f 𝟙 f × 𝟙 f f f 𝟙 f × 𝟙 f f
arrow-is-set : {A B : Object} isSet (Arrow A B)
Isomorphism : {A B} (f : Arrow A B) Set '
Isomorphism {A} {B} f = Σ[ g Arrow B A ] g f 𝟙 × f g 𝟙
_≅_ : (A B : Object) Set '
_≅_ A B = Σ[ f Arrow A B ] (Isomorphism f)
idIso : (A : Object) A A
idIso A = 𝟙 , (𝟙 , ident)
id-to-iso : (A B : Object) A B A B
id-to-iso A B eq = transp (\ i A eq i) (idIso A)
-- TODO: might want to implement isEquiv differently, there are 3
-- equivalent formulations in the book.
field
univalent : {A B : Object} isEquiv (A B) (A B) (id-to-iso A B)
module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set '
Epimorphism {X} f = ( g₀ g₁ : Arrow B X ) g₀ f g₁ f g₀ g₁
Monomorphism : {X : Object} (f : Arrow A B) Set '
Monomorphism {X} f = ( g₀ g₁ : Arrow X A ) f g₀ f g₁ g₀ g₁
module _ {} {'} {Object : Set }
{Arrow : Object Object Set '}
{𝟙 : {o : Object} Arrow o o}
{_⊕_ : { a b c : Object } Arrow b c Arrow a b Arrow a c}
where
-- TODO, provable by using arrow-is-set and that isProp (isEquiv _ _ _)
-- This lemma will be useful to prove the equality of two categories.
IsCategory-is-prop : isProp (IsCategory Object Arrow 𝟙 _⊕_)
IsCategory-is-prop = {!!}
-- open IsCategory public
record Category ( ' : Level) : Set (lsuc (' )) where record Category ( ' : Level) : Set (lsuc (' )) where
-- adding no-eta-equality can speed up type-checking. -- adding no-eta-equality can speed up type-checking.
-- ONLY IF you define your categories with copatterns though.
no-eta-equality no-eta-equality
field field
-- Need something like: -- Need something like:
@ -64,23 +106,6 @@ _[_,_] = Arrow
_[_∘_] : { '} ( : Category ') {A B C : .Object} (g : [ B , C ]) (f : [ A , B ]) [ A , C ] _[_∘_] : { '} ( : Category ') {A B C : .Object} (g : [ B , C ]) (f : [ A , B ]) [ A , C ]
_[_∘_] = _∘_ _[_∘_] = _∘_
module _ { ' : Level} { : Category '} where
module _ { A B : .Object } where
Isomorphism : (f : .Arrow A B) Set '
Isomorphism f = Σ[ g .Arrow B A ] [ g f ] .𝟙 × [ f g ] .𝟙
Epimorphism : {X : .Object } (f : .Arrow A B) Set '
Epimorphism {X} f = ( g₀ g₁ : .Arrow B X ) [ g₀ f ] [ g₁ f ] g₀ g₁
Monomorphism : {X : .Object} (f : .Arrow A B) Set '
Monomorphism {X} f = ( g₀ g₁ : .Arrow X A ) [ f g₀ ] [ f g₁ ] g₀ g₁
-- Isomorphism of objects
_≅_ : (A B : Object ) Set '
_≅_ A B = Σ[ f .Arrow A B ] (Isomorphism f)
module _ { ' : Level} ( : Category ') {A B obj : Object } where module _ { ' : Level} ( : Category ') {A B obj : Object } where
IsProduct : (π₁ : Arrow obj A) (π₂ : Arrow obj B) Set ( ') IsProduct : (π₁ : Arrow obj A) (π₂ : Arrow obj B) Set ( ')
IsProduct π₁ π₂ IsProduct π₁ π₂
@ -130,7 +155,9 @@ module _ { ' : Level} ( : Category ') where
; Arrow = Function.flip ( .Arrow) ; Arrow = Function.flip ( .Arrow)
; 𝟙 = .𝟙 ; 𝟙 = .𝟙
; _∘_ = Function.flip ( ._∘_) ; _∘_ = Function.flip ( ._∘_)
; isCategory = record { assoc = sym assoc ; ident = swap ident } ; isCategory = record { assoc = sym assoc ; ident = swap ident
; arrow-is-set = {!!}
; univalent = {!!} }
} }
where where
open IsCategory ( .isCategory) open IsCategory ( .isCategory)

View file

@ -1,3 +1,4 @@
{-# OPTIONS --allow-unsolved-metas #-}
module Cat.Category.Free where module Cat.Category.Free where
open import Agda.Primitive open import Agda.Primitive
@ -32,5 +33,10 @@ module _ { ' : Level} ( : Category ') where
; Arrow = Path ; Arrow = Path
; 𝟙 = λ {o} emptyPath o ; 𝟙 = λ {o} emptyPath o
; _∘_ = λ {a b c} concatenate {a} {b} {c} ; _∘_ = λ {a b c} concatenate {a} {b} {c}
; isCategory = record { assoc = p-assoc ; ident = ident-r , ident-l } ; isCategory = record
{ assoc = p-assoc
; ident = ident-r , ident-l
; arrow-is-set = {!!}
; univalent = {!!}
}
} }

View file

@ -16,7 +16,7 @@ module _ { ' : Level} { : Category '} { A B : .Category.Obje
open Category open Category
open IsCategory (isCategory) open IsCategory (isCategory)
iso-is-epi : Isomorphism { = } f Epimorphism { = } {X = X} f iso-is-epi : Isomorphism f Epimorphism {X = X} f
iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq = begin iso-is-epi (f- , left-inv , right-inv) g₀ g₁ eq = begin
g₀ ≡⟨ sym (proj₁ ident) g₀ ≡⟨ sym (proj₁ ident)
g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv) g₀ 𝟙 ≡⟨ cong (_∘_ g₀) (sym right-inv)
@ -27,7 +27,7 @@ module _ { ' : Level} { : Category '} { A B : .Category.Obje
g₁ 𝟙 ≡⟨ proj₁ ident g₁ 𝟙 ≡⟨ proj₁ ident
g₁ g₁
iso-is-mono : Isomorphism { = } f Monomorphism { = } {X = X} f iso-is-mono : Isomorphism f Monomorphism {X = X} f
iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq = iso-is-mono (f- , (left-inv , right-inv)) g₀ g₁ eq =
begin begin
g₀ ≡⟨ sym (proj₂ ident) g₀ ≡⟨ sym (proj₂ ident)
@ -39,7 +39,7 @@ module _ { ' : Level} { : Category '} { A B : .Category.Obje
𝟙 g₁ ≡⟨ proj₂ ident 𝟙 g₁ ≡⟨ proj₂ ident
g₁ g₁
iso-is-epi-mono : Isomorphism { = } f Epimorphism { = } {X = X} f × Monomorphism { = } {X = X} f iso-is-epi-mono : Isomorphism f Epimorphism {X = X} f × Monomorphism {X = X} f
iso-is-epi-mono iso = iso-is-epi iso , iso-is-mono iso iso-is-epi-mono iso = iso-is-epi iso , iso-is-mono iso
{- {-
@ -54,71 +54,71 @@ open Category
open import Cat.Functor open import Cat.Functor
open Functor open Functor
module _ { : Level} { : Category } -- module _ { : Level} { : Category }
{isSObj : isSet ( .Object)} -- {isSObj : isSet ( .Object)}
{isz2 : {} {A B : Set } isSet (Sets [ A , B ])} where -- {isz2 : ∀ {} → {A B : Set } → isSet (Sets [ A , B ])} where
open import Cat.Categories.Cat using (Cat) -- -- open import Cat.Categories.Cat using (Cat)
open import Cat.Categories.Fun -- open import Cat.Categories.Fun
open import Cat.Categories.Sets -- open import Cat.Categories.Sets
module Cat = Cat.Categories.Cat -- -- module Cat = Cat.Categories.Cat
open Exponential -- open Exponential
private -- private
Cat = Cat -- Cat = Cat
prshf = presheaf { = } -- prshf = presheaf { = }
module = IsCategory ( .isCategory) -- module = IsCategory ( .isCategory)
-- Exp : Set (lsuc (lsuc )) -- -- Exp : Set (lsuc (lsuc ))
-- Exp = Exponential (Cat (lsuc ) ) -- -- Exp = Exponential (Cat (lsuc ) )
-- Sets (Opposite ) -- -- Sets (Opposite )
_⇑_ : (A B : Cat .Object) Cat .Object -- _⇑_ : (A B : Cat .Object) → Cat .Object
A B = (exponent A B) .obj -- A ⇑ B = (exponent A B) .obj
where -- where
open HasExponentials (Cat.hasExponentials ) -- open HasExponentials (Cat.hasExponentials )
module _ {A B : .Object} (f : .Arrow A B) where -- module _ {A B : .Object} (f : .Arrow A B) where
:func→: : NaturalTransformation (prshf A) (prshf B) -- :func→: : NaturalTransformation (prshf A) (prshf B)
:func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .assoc -- :func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .assoc
module _ {c : .Object} where -- module _ {c : .Object} where
eqTrans : (λ _ Transformation (prshf c) (prshf c)) -- eqTrans : (λ _ → Transformation (prshf c) (prshf c))
[ (λ _ x [ .𝟙 x ]) identityTrans (prshf c) ] -- [ (λ _ x → [ .𝟙 ∘ x ]) ≡ identityTrans (prshf c) ]
eqTrans = funExt λ x funExt λ x .ident .proj₂ -- eqTrans = funExt λ x → funExt λ x → .ident .proj₂
eqNat : (λ i Natural (prshf c) (prshf c) (eqTrans i)) -- eqNat : (λ i → Natural (prshf c) (prshf c) (eqTrans i))
[(λ _ funExt (λ _ .assoc)) identityNatural (prshf c)] -- [(λ _ → funExt (λ _ → .assoc)) ≡ identityNatural (prshf c)]
eqNat = λ i {A} {B} f -- eqNat = λ i {A} {B} f
let -- let
open IsCategory (Sets .isCategory) -- open IsCategory (Sets .isCategory)
lemm : (Sets [ eqTrans i B prshf c .func→ f ]) -- lemm : (Sets [ eqTrans i B ∘ prshf c .func→ f ]) ≡
(Sets [ prshf c .func→ f eqTrans i A ]) -- (Sets [ prshf c .func→ f ∘ eqTrans i A ])
lemm = {!!} -- lemm = {!!}
lem : (λ _ Sets [ Functor.func* (prshf c) A , prshf c .func* B ]) -- lem : (λ _ → Sets [ Functor.func* (prshf c) A , prshf c .func* B ])
[ Sets [ eqTrans i B prshf c .func→ f ] -- [ Sets [ eqTrans i B ∘ prshf c .func→ f ]
Sets [ prshf c .func→ f eqTrans i A ] ] -- ≡ Sets [ prshf c .func→ f ∘ eqTrans i A ] ]
lem -- lem
= isz2 _ _ lemm _ i -- = isz2 _ _ lemm _ i
-- (Sets [ eqTrans i B ∘ prshf c .func→ f ]) -- -- (Sets [ eqTrans i B ∘ prshf c .func→ f ])
-- (Sets [ prshf c .func→ f ∘ eqTrans i A ]) -- -- (Sets [ prshf c .func→ f ∘ eqTrans i A ])
-- lemm -- -- lemm
-- _ i -- -- _ i
in -- in
lem -- lem
-- eqNat = λ {A} {B} i [B,A] i' [A,c] → -- -- eqNat = λ {A} {B} i [B,A] i' [A,c] →
-- let -- -- let
-- k : [ {!!} , {!!} ] -- -- k : [ {!!} , {!!} ]
-- k = [A,c] -- -- k = [A,c]
-- in {! [ ? ∘ ? ]!} -- -- in {! [ ? ∘ ? ]!}
:ident: : (:func→: ( .𝟙 {c})) (Fun .𝟙 {o = prshf c}) -- :ident: : (:func→: ( .𝟙 {c})) (Fun .𝟙 {o = prshf c})
:ident: = Σ≡ eqTrans eqNat -- :ident: = Σ≡ eqTrans eqNat
yoneda : Functor (Fun { = Opposite } {𝔻 = Sets {}}) -- yoneda : Functor (Fun { = Opposite } {𝔻 = Sets {}})
yoneda = record -- yoneda = record
{ func* = prshf -- { func* = prshf
; func→ = :func→: -- ; func→ = :func→:
; isFunctor = record -- ; isFunctor = record
{ ident = :ident: -- { ident = :ident:
; distrib = {!!} -- ; distrib = {!!}
} -- }
} -- }