Finish equality principle for categories

This commit is contained in:
Frederik Hanghøj Iversen 2018-03-12 13:51:29 +01:00
parent 35390c02d3
commit a7214fcc66

View file

@ -38,7 +38,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.NType.Properties using ( propIsEquiv ) open import Cubical.NType.Properties using ( propIsEquiv ; lemPropF )
open import Cat.Wishlist open import Cat.Wishlist
@ -195,9 +195,9 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
-- --
-- Proves that all projections of `IsCategory` are mere propositions as well as -- Proves that all projections of `IsCategory` are mere propositions as well as
-- `IsCategory` itself being a mere proposition. -- `IsCategory` itself being a mere proposition.
module Propositionality {a b : Level} {C : RawCategory a b} where module Propositionality {a b : Level} ( : RawCategory a b) where
open RawCategory C open RawCategory
module _ ( : IsCategory C) where module _ ( : IsCategory ) where
open IsCategory using (isAssociative ; arrowsAreSets ; isIdentity ; Univalent) open IsCategory using (isAssociative ; arrowsAreSets ; isIdentity ; Univalent)
open import Cubical.NType open import Cubical.NType
open import Cubical.NType.Properties open import Cubical.NType.Properties
@ -241,11 +241,11 @@ module Propositionality {a b : Level} {C : RawCategory a b} where
propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i propUnivalent a b i = propPi (λ iso propHasLevel ⟨-2⟩) a b i
private private
module _ (x y : IsCategory C) where module _ (x y : IsCategory ) where
module IC = IsCategory module IC = IsCategory
module X = IsCategory x module X = IsCategory x
module Y = IsCategory y module Y = IsCategory y
open Univalence C open Univalence
-- In a few places I use the result of propositionality of the various -- In a few places I use the result of propositionality of the various
-- projections of `IsCategory` - I've arbitrarily chosed to use this -- projections of `IsCategory` - I've arbitrarily chosed to use this
-- result from `x : IsCategory C`. I don't know which (if any) possibly -- result from `x : IsCategory C`. I don't know which (if any) possibly
@ -275,7 +275,7 @@ module Propositionality {a b : Level} {C : RawCategory a b} where
IC.arrowsAreSets (done i) = propArrowIsSet x X.arrowsAreSets Y.arrowsAreSets i IC.arrowsAreSets (done i) = propArrowIsSet x X.arrowsAreSets Y.arrowsAreSets i
IC.univalent (done i) = eqUni i IC.univalent (done i) = eqUni i
propIsCategory : isProp (IsCategory C) propIsCategory : isProp (IsCategory )
propIsCategory = done propIsCategory = done
-- | Univalent categories -- | Univalent categories
@ -297,15 +297,8 @@ module _ {a b : Level} { 𝔻 : Category a b} where
module _ (rawEq : .raw 𝔻.raw) where module _ (rawEq : .raw 𝔻.raw) where
private private
P : (target : RawCategory a b) ({!!} target) Set _
P _ eq = isCategory' (λ i IsCategory (eq i)) [ .isCategory isCategory' ]
p : P .raw refl
p isCategory' = Propositionality.propIsCategory {!!} {!!}
-- TODO Make and use heterogeneous version of Category≡
isCategoryEq : (λ i IsCategory (rawEq i)) [ .isCategory 𝔻.isCategory ] isCategoryEq : (λ i IsCategory (rawEq i)) [ .isCategory 𝔻.isCategory ]
isCategoryEq = {!!} isCategoryEq = lemPropF Propositionality.propIsCategory rawEq
Category≡ : 𝔻 Category≡ : 𝔻
Category≡ i = record Category≡ i = record