Move functor-equality to functor module
This commit is contained in:
parent
a480fca956
commit
7a77ba230c
|
@ -23,23 +23,6 @@ eqpair eqa eqb i = eqa i , eqb i
|
|||
open Functor
|
||||
open Category
|
||||
|
||||
module _ {ℓ ℓ' : Level} {A B : Category ℓ ℓ'} where
|
||||
lift-eq-functors : {f g : Functor A B}
|
||||
→ (eq* : f .func* ≡ g .func*)
|
||||
→ (eq→ : PathP (λ i → ∀ {x y} → A .Arrow x y → B .Arrow (eq* i x) (eq* i y))
|
||||
(f .func→) (g .func→))
|
||||
-- → (eq→ : Functor.func→ f ≡ {!!}) -- Functor.func→ g)
|
||||
-- Use PathP
|
||||
-- directly to show heterogeneous equalities by using previous
|
||||
-- equalities (i.e. continuous paths) to create new continuous paths.
|
||||
→ (eqI : PathP (λ i → ∀ {c : A .Object} → eq→ i (A .𝟙 {c}) ≡ B .𝟙 {eq* i c})
|
||||
(ident f) (ident g))
|
||||
→ (eqD : PathP (λ i → { c c' c'' : A .Object} {a : A .Arrow c c'} {a' : A .Arrow c' c''}
|
||||
→ eq→ i (A ._⊕_ a' a) ≡ B ._⊕_ (eq→ i a') (eq→ i a))
|
||||
(distrib f) (distrib g))
|
||||
→ f ≡ g
|
||||
lift-eq-functors eq* eq→ eqI eqD i = record { func* = eq* i ; func→ = eq→ i ; ident = eqI i ; distrib = eqD i }
|
||||
|
||||
-- The category of categories
|
||||
module _ (ℓ ℓ' : Level) where
|
||||
private
|
||||
|
@ -59,10 +42,9 @@ module _ (ℓ ℓ' : Level) where
|
|||
postulate eqD : PathP (λ i → { c c' c'' : A .Object} {a : A .Arrow c c'} {a' : A .Arrow c' c''}
|
||||
→ eq→ i (A ._⊕_ a' a) ≡ D ._⊕_ (eq→ i a') (eq→ i a))
|
||||
(distrib (h ∘f (g ∘f f))) (distrib ((h ∘f g) ∘f f))
|
||||
-- eqD = {!!}
|
||||
|
||||
assc : h ∘f (g ∘f f) ≡ (h ∘f g) ∘f f
|
||||
assc = lift-eq-functors eq* eq→ eqI eqD
|
||||
assc = Functor≡ eq* eq→ eqI eqD
|
||||
|
||||
module _ {A B : Category ℓ ℓ'} {f : Functor A B} where
|
||||
lem : (func* f) ∘ (func* (identity {C = A})) ≡ func* f
|
||||
|
|
|
@ -19,9 +19,29 @@ record Functor {ℓc ℓc' ℓd ℓd'} (C : Category ℓc ℓc') (D : Category
|
|||
distrib : { c c' c'' : C .Object} {a : C .Arrow c c'} {a' : C .Arrow c' c''}
|
||||
→ func→ (C ._⊕_ a' a) ≡ D ._⊕_ (func→ a') (func→ a)
|
||||
|
||||
open Functor
|
||||
open Category
|
||||
|
||||
module _ {ℓ ℓ' : Level} {ℂ 𝔻 : Category ℓ ℓ'} where
|
||||
private
|
||||
_ℂ⊕_ = ℂ ._⊕_
|
||||
Functor≡ : {F G : Functor ℂ 𝔻}
|
||||
→ (eq* : F .func* ≡ G .func*)
|
||||
→ (eq→ : PathP (λ i → ∀ {x y} → ℂ .Arrow x y → 𝔻 .Arrow (eq* i x) (eq* i y))
|
||||
(F .func→) (G .func→))
|
||||
-- → (eq→ : Functor.func→ f ≡ {!!}) -- Functor.func→ g)
|
||||
-- Use PathP
|
||||
-- directly to show heterogeneous equalities by using previous
|
||||
-- equalities (i.e. continuous paths) to create new continuous paths.
|
||||
→ (eqI : PathP (λ i → ∀ {A : ℂ .Object} → eq→ i (ℂ .𝟙 {A}) ≡ 𝔻 .𝟙 {eq* i A})
|
||||
(ident F) (ident G))
|
||||
→ (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))
|
||||
(distrib F) (distrib G))
|
||||
→ F ≡ G
|
||||
Functor≡ eq* eq→ eqI eqD i = record { func* = eq* i ; func→ = eq→ i ; ident = eqI i ; distrib = eqD i }
|
||||
|
||||
module _ {ℓ ℓ' : Level} {A B C : Category ℓ ℓ'} (F : Functor B C) (G : Functor A B) where
|
||||
open Functor
|
||||
open Category
|
||||
private
|
||||
F* = F .func*
|
||||
F→ = F .func→
|
||||
|
|
Loading…
Reference in a new issue