commented out code with holes

This commit is contained in:
Andrea Vezzosi 2018-05-15 17:21:57 +02:00
parent 513d91ae4f
commit c75a1d5d8b
4 changed files with 74 additions and 71 deletions

View file

@ -1,6 +1,7 @@
{-# OPTIONS --allow-unsolved-metas --cubical --caching #-} {-# OPTIONS --allow-unsolved-metas --cubical --caching #-}
module Cat.Categories.Fun where module Cat.Categories.Fun where
open import Cat.Prelude open import Cat.Prelude
open import Cat.Equivalence open import Cat.Equivalence
@ -52,14 +53,14 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
lem : coe (pp {C}) 𝔻.identity f→g lem : coe (pp {C}) 𝔻.identity f→g
lem = trans (𝔻.9-1-9-right {b = Functor.omap F C} 𝔻.identity p*) 𝔻.rightIdentity lem = trans (𝔻.9-1-9-right {b = Functor.omap F C} 𝔻.identity p*) 𝔻.rightIdentity
idToNatTrans : NaturalTransformation F G -- idToNatTrans : NaturalTransformation F G
idToNatTrans = (λ C coe pp 𝔻.identity) , λ f begin -- idToNatTrans = (λ C → coe pp 𝔻.identity) , λ f → begin
coe pp 𝔻.identity 𝔻.<<< F.fmap f ≡⟨ cong (𝔻._<<< F.fmap f) lem -- coe pp 𝔻.identity 𝔻.<<< F.fmap f ≡⟨ cong (𝔻._<<< F.fmap f) lem ⟩
-- Just need to show that f→g is a natural transformation -- -- Just need to show that f→g is a natural transformation
-- I know that it has an inverse; g→f -- -- I know that it has an inverse; g→f
f→g 𝔻.<<< F.fmap f ≡⟨ {!lem!} -- f→g 𝔻.<<< F.fmap f ≡⟨ {!lem!} ⟩
G.fmap f 𝔻.<<< f→g ≡⟨ cong (G.fmap f 𝔻.<<<_) (sym lem) -- G.fmap f 𝔻.<<< f→g ≡⟨ cong (G.fmap f 𝔻.<<<_) (sym lem) ⟩
G.fmap f 𝔻.<<< coe pp 𝔻.identity -- G.fmap f 𝔻.<<< coe pp 𝔻.identity ∎
module _ {A B : Functor 𝔻} where module _ {A B : Functor 𝔻} where
module A = Functor A module A = Functor A
@ -92,70 +93,70 @@ module Fun {c c' d d' : Level} ( : Category c c') (𝔻 : C
U : (F : .Object 𝔻.Object) Set _ U : (F : .Object 𝔻.Object) Set _
U F = {A B : .Object} [ A , B ] 𝔻 [ F A , F B ] U F = {A B : .Object} [ A , B ] 𝔻 [ F A , F B ]
module _ -- module _
(omap : .Object 𝔻.Object) -- (omap : .Object → 𝔻.Object)
(p : A.omap omap) -- (p : A.omap ≡ omap)
where -- where
D : Set _ -- D : Set _
D = ( fmap : U omap) -- D = ( fmap : U omap)
( let -- → ( let
raw-B : RawFunctor 𝔻 -- raw-B : RawFunctor 𝔻
raw-B = record { omap = omap ; fmap = fmap } -- raw-B = record { omap = omap ; fmap = fmap }
) -- )
(isF-B' : IsFunctor 𝔻 raw-B) -- → (isF-B' : IsFunctor 𝔻 raw-B)
( let -- → ( let
B' : Functor 𝔻 -- B' : Functor 𝔻
B' = record { raw = raw-B ; isFunctor = isF-B' } -- B' = record { raw = raw-B ; isFunctor = isF-B' }
) -- )
(iso' : A B') PathP (λ i U (p i)) A.fmap fmap -- → (iso' : A ≊ B') → PathP (λ i → U (p i)) A.fmap fmap
-- D : Set _ -- -- D : Set _
-- D = PathP (λ i → U (p i)) A.fmap fmap -- -- D = PathP (λ i → U (p i)) A.fmap fmap
-- eeq : (λ f → A.fmap f) ≡ fmap -- -- eeq : (λ f → A.fmap f) ≡ fmap
-- eeq = funExtImp (λ A → funExtImp (λ B → funExt (λ f → isofmap {A} {B} f))) -- -- eeq = funExtImp (λ A → funExtImp (λ B → funExt (λ f → isofmap {A} {B} f)))
-- where -- -- where
-- module _ {X : .Object} {Y : .Object} (f : [ X , Y ]) where -- -- module _ {X : .Object} {Y : .Object} (f : [ X , Y ]) where
-- isofmap : A.fmap f ≡ fmap f -- -- isofmap : A.fmap f ≡ fmap f
-- isofmap = {!ap!} -- -- isofmap = {!ap!}
d : D A.omap refl -- d : D A.omap refl
d = res -- d = res
where -- where
module _ -- module _
( fmap : U A.omap ) -- ( fmap : U A.omap )
( let -- ( let
raw-B : RawFunctor 𝔻 -- raw-B : RawFunctor 𝔻
raw-B = record { omap = A.omap ; fmap = fmap } -- raw-B = record { omap = A.omap ; fmap = fmap }
) -- )
( isF-B' : IsFunctor 𝔻 raw-B ) -- ( isF-B' : IsFunctor 𝔻 raw-B )
( let -- ( let
B' : Functor 𝔻 -- B' : Functor 𝔻
B' = record { raw = raw-B ; isFunctor = isF-B' } -- B' = record { raw = raw-B ; isFunctor = isF-B' }
) -- )
( iso' : A B' ) -- ( iso' : A ≊ B' )
where -- where
module _ {X Y : .Object} (f : [ X , Y ]) where -- module _ {X Y : .Object} (f : [ X , Y ]) where
step : {!!} 𝔻.≊ {!!} -- step : {!!} 𝔻.≊ {!!}
step = {!!} -- step = {!!}
resres : A.fmap {X} {Y} f fmap {X} {Y} f -- resres : A.fmap {X} {Y} f ≡ fmap {X} {Y} f
resres = {!!} -- resres = {!!}
res : PathP (λ i U A.omap) A.fmap fmap -- res : PathP (λ i → U A.omap) A.fmap fmap
res i {X} {Y} f = resres f i -- res i {X} {Y} f = resres f i
fmapEq : PathP (λ i U (omapEq i)) A.fmap B.fmap -- fmapEq : PathP (λ i → U (omapEq i)) A.fmap B.fmap
fmapEq = pathJ D d B.omap omapEq B.fmap B.isFunctor iso -- fmapEq = pathJ D d B.omap omapEq B.fmap B.isFunctor iso
rawEq : A.raw B.raw -- rawEq : A.raw ≡ B.raw
rawEq i = record { omap = omapEq i ; fmap = fmapEq i } -- rawEq i = record { omap = omapEq i ; fmap = fmapEq i }
private -- private
f : (A B) (A B) -- f : (A ≡ B) → (A ≊ B)
f p = idToNatTrans p , idToNatTrans (sym p) , NaturalTransformation≡ A A (funExt (λ C {!!})) , {!!} -- f p = idToNatTrans p , idToNatTrans (sym p) , NaturalTransformation≡ A A (funExt (λ C → {!!})) , {!!}
g : (A B) (A B) -- g : (A ≊ B) → (A ≡ B)
g = Functor≡ rawEq -- g = Functor≡ ∘ rawEq
inv : AreInverses f g -- inv : AreInverses f g
inv = {!funExt λ p → ?!} , {!!} -- inv = {!funExt λ p → ?!} , {!!}
postulate
iso : (A B) (A B) iso : (A B) (A B)
iso = f , g , inv -- iso = f , g , inv
univ : (A B) (A B) univ : (A B) (A B)
univ = fromIsomorphism _ _ iso univ = fromIsomorphism _ _ iso

View file

@ -44,7 +44,7 @@ open Cat.Equivalence
-- about these. The laws defined are the types the propositions - not the -- about these. The laws defined are the types the propositions - not the
-- witnesses to them! -- witnesses to them!
record RawCategory (a b : Level) : Set (lsuc (a b)) where record RawCategory (a b : Level) : Set (lsuc (a b)) where
no-eta-equality -- no-eta-equality
field field
Object : Set a Object : Set a
Arrow : Object Object Set b Arrow : Object Object Set b

View file

@ -4,6 +4,7 @@ This module provides construction 2.3 in [voe]
{-# OPTIONS --cubical --caching #-} {-# OPTIONS --cubical --caching #-}
module Cat.Category.Monad.Voevodsky where module Cat.Category.Monad.Voevodsky where
open import Cat.Prelude open import Cat.Prelude
open import Cat.Category open import Cat.Category

View file

@ -1,6 +1,7 @@
{-# OPTIONS --cubical --caching #-} {-# OPTIONS --cubical --caching #-}
module Cat.Category.Product where module Cat.Category.Product where
open import Cat.Prelude as P hiding (_×_ ; fst ; snd) open import Cat.Prelude as P hiding (_×_ ; fst ; snd)
open import Cat.Equivalence open import Cat.Equivalence
@ -11,7 +12,7 @@ module _ {a b : Level} ( : Category a b) where
module _ (A B : Object) where module _ (A B : Object) where
record RawProduct : Set (a b) where record RawProduct : Set (a b) where
no-eta-equality -- no-eta-equality
field field
object : Object object : Object
fst : [ object , A ] fst : [ object , A ]