Use equality construction principle

Also update submodules
This commit is contained in:
Frederik Hanghøj Iversen 2018-01-30 22:41:18 +01:00
parent 255b0236f9
commit 86d3d7368e
7 changed files with 68 additions and 39 deletions

@ -1 +1 @@
Subproject commit de23244a73d6dab55715fd5a107a5de805c55764
Subproject commit b5bfbc3c170b0bd0c9aaac1b4d4b3f9b06832bf6

@ -1 +1 @@
Subproject commit a83f5f4c63e5dfd8143ac03163868c63a56802de
Subproject commit 1d6730c4999daa2b04e9dd39faa0791d0b5c3b48

View File

@ -44,7 +44,7 @@ module _ ( ' : Level) where
((h ∘f (g ∘f f)) .isFunctor .distrib) (((h ∘f g) ∘f f) .isFunctor .distrib)
assc : h ∘f (g ∘f f) (h ∘f g) ∘f f
assc = Functor≡ eq* eq→ eqI eqD
assc = Functor≡ eq* eq→ (IsFunctor≡ eqI eqD)
module _ { 𝔻 : Category '} {F : Functor 𝔻} where
module _ where
@ -58,16 +58,17 @@ module _ ( ' : Level) where
(func→ (F ∘f identity)) (func→ F)
eq→ = refl
postulate
eqI-r : PathP (λ i {c : .Object}
PathP (λ _ Arrow 𝔻 (func* F c) (func* F c)) (func→ F ( .𝟙)) (𝔻 .𝟙))
((F ∘f identity) .isFunctor .ident) (F .isFunctor .ident)
eqI-r
: (λ i {c : .Object} (λ _ 𝔻 [ func* F c , func* F c ])
[ func→ F ( .𝟙) 𝔻 .𝟙 ])
[(F ∘f identity) .isFunctor .ident F .isFunctor .ident ]
eqD-r : PathP
(λ i
{A B C : .Object} {f : .Arrow A B} {g : .Arrow B C}
eq→ i ( [ g f ]) 𝔻 [ eq→ i g eq→ i f ])
((F ∘f identity) .isFunctor .distrib) (F .isFunctor .distrib)
ident-r : F ∘f identity F
ident-r = Functor≡ eq* eq→ eqI-r eqD-r
ident-r = Functor≡ eq* eq→ (IsFunctor≡ eqI-r eqD-r)
module _ where
private
postulate
@ -75,13 +76,14 @@ module _ ( ' : Level) where
eq→ : PathP
(λ i {x y : Object } .Arrow x y 𝔻 .Arrow (eq* i x) (eq* i y))
((identity ∘f F) .func→) (F .func→)
eqI : PathP (λ i {A : .Object} eq→ i ( .𝟙 {A}) 𝔻 .𝟙 {eq* i A})
((identity ∘f F) .isFunctor .ident) (F .isFunctor .ident)
eqI : (λ i {A : .Object} eq→ i ( .𝟙 {A}) 𝔻 .𝟙 {eq* i A})
[ ((identity ∘f F) .isFunctor .ident) (F .isFunctor .ident) ]
eqD : PathP (λ i {A B C : .Object} {f : .Arrow A B} {g : .Arrow B C}
eq→ i ( [ g f ]) 𝔻 [ eq→ i g eq→ i f ])
((identity ∘f F) .isFunctor .distrib) (F .isFunctor .distrib)
-- (λ z → eq* i z) (eq→ i)
ident-l : identity ∘f F F
ident-l = Functor≡ eq* eq→ eqI eqD
ident-l = Functor≡ eq* eq→ λ i record { ident = eqI i ; distrib = eqD i }
Cat : Category (lsuc ( ')) ( ')
Cat =

View File

@ -53,7 +53,7 @@ module _ {c c' d d' : Level} { : Category c c'} {𝔻 : Cat
𝔻 [ F→ f identityTrans F A ]
where
F→ = F .func→
open module 𝔻 = IsCategory (𝔻 .isCategory)
module 𝔻 = IsCategory (𝔻 .isCategory)
identityNat : (F : Functor 𝔻) NaturalTransformation F F
identityNat F = identityTrans F , identityNatural F

View File

@ -9,6 +9,8 @@ open import Cubical
open import Cat.Category
open import Cat.Functor
open import Cat.Categories.Sets
open import Cat.Equality
open Equality.Data.Product
module _ { ' : Level} { : Category '} { A B : .Category.Object } {X : .Category.Object} (f : .Category.Arrow A B) where
open Category
@ -58,6 +60,7 @@ module _ { : Level} { : Category } where
private
Cat = Cat
prshf = presheaf { = }
module = IsCategory ( .isCategory)
-- Exp : Set (lsuc (lsuc ))
-- Exp = Exponential (Cat (lsuc ) )
@ -70,22 +73,24 @@ module _ { : Level} { : Category } where
module _ {A B : .Object} (f : .Arrow A B) where
:func→: : NaturalTransformation (prshf A) (prshf B)
:func→: = (λ C x [ f x ]) , λ f₁ funExt λ x lem
where
lem = ( .isCategory) .IsCategory.assoc
:func→: = (λ C x [ f x ]) , λ f₁ funExt λ _ .assoc
module _ {c : .Object} where
eqTrans : (:func→: ( .𝟙 {c})) .proj₁ (Fun .𝟙 {o = prshf c}) .proj₁
eqTrans = funExt λ x funExt λ x .isCategory .IsCategory.ident .proj₂
eqNat
: PathP (λ i {A B : .Object} (f : Opposite .Arrow A B)
Sets [ eqTrans i B prshf c .Functor.func→ f ]
Sets [ prshf c .Functor.func→ f eqTrans i A ])
((:func→: ( .𝟙 {c})) .proj₂) ((Fun .𝟙 {o = prshf c}) .proj₂)
eqNat = λ i f i' x₁ {! ._⊕_ ? ?!}
-- eqNat i f = {!!}
-- Sets ._⊕_ (eq₁ i B) (prshf A .func→ f) ≡ Sets ._⊕_ (prshf B .func→ f) (eq₁ i A)
eqTrans : (λ _ Transformation (prshf c) (prshf c))
[ (λ _ x [ .𝟙 x ]) identityTrans (prshf c) ]
eqTrans = funExt λ x funExt λ x .ident .proj₂
eqNat : (λ i Natural (prshf c) (prshf c) (eqTrans i))
[(λ _ funExt (λ _ .assoc)) identityNatural (prshf c)]
eqNat = {!!}
-- eqNat = λ {A} {B} i [B,A] i' [A,c] →
-- let
-- k : [ {!!} , {!!} ]
-- k = [A,c]
-- in {! [ ? ∘ ? ]!}
:ident: : (:func→: ( .𝟙 {c})) (Fun .𝟙 {o = prshf c})
:ident: = NaturalTransformation≡ eqTrans eqNat
:ident: = Σ eqTrans eqNat
yoneda : Functor (Fun { = Opposite } {𝔻 = Sets {}})
yoneda = record

21
src/Cat/Equality.agda Normal file
View File

@ -0,0 +1,21 @@
-- Defines equality-principles for data-types from the standard library.
module Cat.Equality where
open import Level
open import Cubical
-- _[_≡_] = PathP
module Equality where
module Data where
module Product where
open import Data.Product public
module _ {a b : Level} {A : Set a} {B : A Set b} {a b : Σ A B}
(proj₁≡ : (λ _ A) [ proj₁ a proj₁ b ])
(proj₂≡ : (λ i B (proj₁≡ i)) [ proj₂ a proj₂ b ]) where
Σ≡ : a b
proj₁ (Σ≡ i) = proj₁≡ i
proj₂ (Σ≡ i) = proj₂≡ i

View File

@ -32,27 +32,28 @@ open Functor
module _ { ' : Level} { 𝔻 : Category '} where
-- IsFunctor≡ : ∀ {A B : .Object} {func* : .Object → 𝔻 .Object} {func→ : {A B : .Object} → .Arrow A B → 𝔻 .Arrow (func* A) (func* B)} {F G : IsFunctor 𝔻 func* func→}
-- → (eqI : PathP (λ i → ∀ {A : .Object} → func→ ( .𝟙 {A}) ≡ 𝔻 .𝟙 {func* A})
-- (F .ident) (G .ident))
-- → (eqD : PathP (λ i → {A B C : .Object} {f : .Arrow A B} {g : .Arrow B C}
-- → func→ ( ._⊕_ g f) ≡ 𝔻 ._⊕_ (func→ g) (func→ f))
-- (F .distrib) (G .distrib))
-- → F ≡ G
-- IsFunctor≡ eqI eqD i = record { ident = eqI i ; distrib = eqD i }
IsFunctor≡
: {func* : .Object 𝔻 .Object}
{func→ : {A B : .Object} .Arrow A B 𝔻 .Arrow (func* A) (func* B)}
{F G : IsFunctor 𝔻 func* func→}
(eqI
: (λ i {A} func→ ( .𝟙 {A}) 𝔻 .𝟙 {func* A})
[ F .ident G .ident ])
(eqD :
(λ i {A B C} {f : [ A , B ]} {g : [ B , C ]}
func→ ( [ g f ]) 𝔻 [ func→ g func→ f ])
[ F .distrib G .distrib ])
(λ _ IsFunctor 𝔻 (λ i func* i) func→) [ F G ]
IsFunctor≡ eqI eqD i = record { ident = eqI i ; distrib = eqD i }
Functor≡ : {F G : Functor 𝔻}
(eq* : F .func* G .func*)
(eq→ : PathP (λ i {x y} [ x , y ] 𝔻 [ eq* i x , eq* i y ])
(F .func→) (G .func→))
-- → (eqIsF : PathP (λ i → IsFunctor 𝔻 (eq* i) (eq→ i)) (F .isFunctor) (G .isFunctor))
(eqI : PathP (λ i {A : .Object} eq→ i ( .𝟙 {A}) 𝔻 .𝟙 {eq* i A})
(F .isFunctor .ident) (G .isFunctor .ident))
(eqD : PathP (λ i {A B C : .Object} {f : [ A , B ]} {g : [ B , C ]}
eq→ i ( [ g f ]) 𝔻 [ eq→ i g eq→ i f ])
(F .isFunctor .distrib) (G .isFunctor .distrib))
(eqIsFunctor : (λ i IsFunctor 𝔻 (eq* i) (eq→ i)) [ F .isFunctor G .isFunctor ])
F G
Functor≡ eq* eq→ eqI eqD i = record { func* = eq* i ; func→ = eq→ i ; isFunctor = record { ident = eqI i ; distrib = eqD i } }
Functor≡ eq* eq→ eqIsFunctor i = record { func* = eq* i ; func→ = eq→ i ; isFunctor = eqIsFunctor i }
module _ { ' : Level} {A B C : Category '} (F : Functor B C) (G : Functor A B) where
private