Towards IsCategory-is-prop
This commit is contained in:
parent
6ea3b5f2b2
commit
fecb4dc1ce
|
@ -12,6 +12,7 @@ open import Data.Product renaming
|
||||||
open import Data.Empty
|
open import Data.Empty
|
||||||
import Function
|
import Function
|
||||||
open import Cubical
|
open import Cubical
|
||||||
|
open import Cubical.GradLemma using ( propIsEquiv )
|
||||||
|
|
||||||
∃! : ∀ {a b} {A : Set a}
|
∃! : ∀ {a b} {A : Set a}
|
||||||
→ (A → Set b) → Set (a ⊔ b)
|
→ (A → Set b) → Set (a ⊔ b)
|
||||||
|
@ -71,12 +72,27 @@ module _ {ℓ} {ℓ'} {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}
|
||||||
where
|
where
|
||||||
|
|
||||||
-- TODO, provable by using arrow-is-set and that isProp (isEquiv _ _ _)
|
-- TODO, provable by using arrow-is-set and that isProp (isEquiv _ _ _)
|
||||||
-- This lemma will be useful to prove the equality of two categories.
|
-- This lemma will be useful to prove the equality of two categories.
|
||||||
IsCategory-is-prop : isProp (IsCategory Object Arrow 𝟙 _⊕_)
|
IsCategory-is-prop : isProp (IsCategory Object Arrow 𝟙 _⊕_)
|
||||||
IsCategory-is-prop = {!!}
|
IsCategory-is-prop x y i = record
|
||||||
|
{ assoc = x.arrow-is-set _ _ x.assoc y.assoc i
|
||||||
|
; ident =
|
||||||
|
( x.arrow-is-set _ _ (fst x.ident) (fst y.ident) i
|
||||||
|
, x.arrow-is-set _ _ (snd x.ident) (snd y.ident) i
|
||||||
|
)
|
||||||
|
-- ; arrow-is-set = {!λ x₁ y₁ p q → x.arrow-is-set _ _ p q!}
|
||||||
|
; arrow-is-set = λ _ _ p q →
|
||||||
|
let
|
||||||
|
golden : x.arrow-is-set _ _ p q ≡ y.arrow-is-set _ _ p q
|
||||||
|
golden = λ j k l → {!!}
|
||||||
|
in
|
||||||
|
golden i
|
||||||
|
; univalent = λ y₁ → {!!}
|
||||||
|
}
|
||||||
|
where
|
||||||
|
module x = IsCategory x
|
||||||
|
module y = IsCategory y
|
||||||
|
|
||||||
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.
|
||||||
|
|
Loading…
Reference in a new issue