Use other equality principle

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-05 17:10:41 +01:00
parent 3151fb3e46
commit 9ec6ce9eba
3 changed files with 19 additions and 36 deletions

View file

@ -25,14 +25,14 @@ module _ ( ' : Level) where
private
module _ {𝔸 𝔹 𝔻 : Category '} {F : Functor 𝔸 𝔹} {G : Functor 𝔹 } {H : Functor 𝔻} where
assc : F[ H F[ G F ] ] F[ F[ H G ] F ]
assc = Functor≡ refl refl
assc = Functor≡ refl
module _ { 𝔻 : Category '} {F : Functor 𝔻} where
ident-r : F[ F identity ] F
ident-r = Functor≡ refl refl
ident-r = Functor≡ refl
ident-l : F[ identity F ] F
ident-l = Functor≡ refl refl
ident-l = Functor≡ refl
RawCat : RawCategory (lsuc ( ')) ( ')
RawCat =
@ -133,16 +133,10 @@ module CatProduct { ' : Level} ( 𝔻 : Category ') where
open module x = Functor x₂
isUniqL : F[ proj₁ x ] x₁
isUniqL = Functor≡ eq* eq→
where
eq* : (F[ proj₁ x ]) .func* x₁ .func*
eq* = refl
eq→ : (λ i {A : Object X} {B : Object X} X [ A , B ] [ eq* i A , eq* i B ])
[ (F[ proj₁ x ]) .func→ x₁ .func→ ]
eq→ = refl
isUniqL = Functor≡ refl
isUniqR : F[ proj₂ x ] x₂
isUniqR = Functor≡ refl refl
isUniqR = Functor≡ refl
isUniq : F[ proj₁ x ] x₁ × F[ proj₂ x ] x₂
isUniq = isUniqL , isUniqR

View file

@ -112,25 +112,10 @@ module _
module _ { ' : Level} { 𝔻 : Category '} where
Functor≡ : {F G : Functor 𝔻}
(eq* : func* F func* G)
(eq→ : (λ i {x y} [ x , y ] 𝔻 [ eq* i x , eq* i y ])
[ func→ F func→ G ])
F G
Functor≡ {F} {G} eq* eq→ i = record
{ raw = eqR i
; isFunctor = eqIsF i
}
where
eqR : raw F raw G
eqR i = record { func* = eq* i ; func→ = eq→ i }
eqIsF : (λ i IsFunctor 𝔻 (eqR i)) [ isFunctor F isFunctor G ]
eqIsF = IsFunctorIsProp' (isFunctor F) (isFunctor G)
FunctorEq : {F G : Functor 𝔻}
raw F raw G
F G
raw (FunctorEq eq i) = eq i
isFunctor (FunctorEq {F} {G} eq i)
raw (Functor≡ eq i) = eq i
isFunctor (Functor≡ {F} {G} eq i)
= res i
where
res : (λ i IsFunctor 𝔻 (eq i)) [ isFunctor F isFunctor G ]

View file

@ -6,7 +6,7 @@ open import Agda.Primitive
open import Data.Product
open import Cubical
open import Cubical.NType.Properties using (lemPropF)
open import Cubical.NType.Properties using (lemPropF ; lemSig)
open import Cat.Category
open import Cat.Category.Functor as F
@ -537,8 +537,12 @@ module _ {a b : Level} { : Category a b} where
rawEq→ : P (RawFunctor.func* right) refl (RawFunctor.func→ right)
-- rawEq→ : (fmap' : Fmap {!!}) → RawFunctor.func→ left ≡ fmap'
rawEq→ = begin
(λ {A} {B} RawFunctor.func→ left) ≡⟨ {!!}
(λ {A} {B} RawFunctor.func→ right)
(λ f RawFunctor.func→ left f) ≡⟨⟩
(λ f KM.fmap f) ≡⟨⟩
(λ f KM.bind (f >>> KM.pure)) ≡⟨ {!!}
(λ f RawFunctor.func→ right f)
where
module KM = K.Monad (forth m)
-- destfmap =
source = (Functor.raw (K.Monad.R (forth m)))
-- p : (fmap' : Fmap (RawFunctor.func* source)) → (λ i → Fmap (refl i)) [ func→ source ≡ fmap' ]
@ -546,16 +550,16 @@ module _ {a b : Level} { : Category a b} where
rawEq : Functor.raw (K.Monad.R (forth m)) Functor.raw R
rawEq = RawFunctor≡ {x = left} {right} refl λ fmap' {!rawEq→!}
Req : M.RawMonad.R (backRaw (forth m)) R
Req = FunctorEq rawEq
Req = Functor rawEq
ηeq : M.RawMonad.η (backRaw (forth m)) η
ηeq = {!!}
postulate ηNatTransEq : {!!} [ M.RawMonad.ηNatTrans (backRaw (forth m)) ηNatTrans ]
open NaturalTransformation
postulate
ηNatTransEq : (λ i NaturalTransformation F.identity (Req i))
[ M.RawMonad.ηNatTrans (backRaw (forth m)) ηNatTrans ]
backRawEq : backRaw (forth m) M.Monad.raw m
-- stuck
M.RawMonad.R (backRawEq i) = Req i
M.RawMonad.ηNatTrans (backRawEq i) = let t = NaturalTransformation≡ F.identity R ηeq in {!t i!}
M.RawMonad.ηNatTrans (backRawEq i) = {!!} -- ηNatTransEq i
M.RawMonad.μNatTrans (backRawEq i) = {!!}
backeq : (m : M.Monad) back (forth m) m