Move various type-synonyms to RawCategory

This commit is contained in:
Frederik Hanghøj Iversen 2018-02-20 16:22:38 +01:00
parent 38ec53d5c2
commit 10df9511a4

View file

@ -25,19 +25,46 @@ open import Cat.Wishlist
syntax ∃!-syntax (λ x B) = ∃![ x ] B
record RawCategory ( ' : Level) : Set (lsuc (' )) where
record RawCategory (a b : Level) : Set (lsuc (a b)) where
no-eta-equality
field
Object : Set
Arrow : Object Object Set '
𝟙 : {o : Object} Arrow o o
Object : Set a
Arrow : Object Object Set b
𝟙 : {A : Object} Arrow A A
_∘_ : {A B C : Object} Arrow B C Arrow A B Arrow A C
infixl 10 _∘_
domain : { a b : Object } Arrow a b Object
domain {a = a} _ = a
codomain : { a b : Object } Arrow a b Object
codomain {b = b} _ = b
IsAssociative : Set (a b)
IsAssociative = {A B C D} {f : Arrow A B} {g : Arrow B C} {h : Arrow C D}
h (g f) (h g) f
IsIdentity : ({A : Object} Arrow A A) Set (a b)
IsIdentity id = {A B : Object} {f : Arrow A B}
f id f × id f f
IsInverseOf : {A B} (Arrow A B) (Arrow B A) Set b
IsInverseOf = λ f g g f 𝟙 × f g 𝟙
Isomorphism : {A B} (f : Arrow A B) Set b
Isomorphism {A} {B} f = Σ[ g Arrow B A ] IsInverseOf f g
_≅_ : (A B : Object) Set b
_≅_ A B = Σ[ f Arrow A B ] (Isomorphism f)
module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set b
Epimorphism {X} f = ( g₀ g₁ : Arrow B X ) g₀ f g₁ f g₀ g₁
Monomorphism : {X : Object} (f : Arrow A B) Set b
Monomorphism {X} f = ( g₀ g₁ : Arrow X A ) f g₀ f g₁ g₀ g₁
-- Thierry: All projections must be `isProp`'s
-- According to definitions 9.1.1 and 9.1.6 in the HoTT book the
@ -47,15 +74,6 @@ record RawCategory ( ' : Level) : Set (lsuc (' ⊔ )) where
record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc (a b)) where
open RawCategory
module Raw = RawCategory
IsAssociative : Set (a b)
IsAssociative = {A B C D} {f : Arrow A B} {g : Arrow B C} {h : Arrow C D}
h (g f) (h g) f
IsIdentity : ({A : Object} Arrow A A) Set (a b)
IsIdentity id = {A B : Object} {f : Arrow A B}
f id f × id f f
field
assoc : IsAssociative
ident : IsIdentity 𝟙
@ -72,9 +90,6 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
propArrowIsSet : isProp ( {A B} isSet (Arrow A B))
propArrowIsSet a b i = isSetIsProp a b i
IsInverseOf : {A B} (Arrow A B) (Arrow B A) Set b
IsInverseOf = λ f g g f 𝟙 × f g 𝟙
propIsInverseOf : {A B f g} isProp (IsInverseOf {A} {B} f g)
propIsInverseOf x y = λ i
let
@ -84,15 +99,6 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
hh = arrowIsSet _ _ (snd x) (snd y)
in h i , hh i
Isomorphism : {A B} (f : Arrow A B) Set b
Isomorphism {A} {B} f = Σ[ g Arrow B A ] IsInverseOf f g
inverse : {A B} {f : Arrow A B} Isomorphism f Arrow B A
inverse iso = fst iso
_≅_ : (A B : Object) Set b
_≅_ A B = Σ[ f Arrow A B ] (Isomorphism f)
idIso : (A : Object) A A
idIso A = 𝟙 , (𝟙 , ident)
@ -107,13 +113,6 @@ record IsCategory {a b : Level} ( : RawCategory a b) : Set (lsuc
field
univalent : Univalent
module _ {A B : Object} where
Epimorphism : {X : Object } (f : Arrow A B) Set b
Epimorphism {X} f = ( g₀ g₁ : Arrow B X ) g₀ f g₁ f g₀ g₁
Monomorphism : {X : Object} (f : Arrow A B) Set b
Monomorphism {X} f = ( g₀ g₁ : Arrow X A ) f g₀ f g₁ g₀ g₁
module _ {A B : Object} {f : Arrow A B} where
isoIsProp : isProp (Isomorphism f)
isoIsProp a@(g , η , ε) a'@(g' , η' , ε') =
@ -161,7 +160,7 @@ module _ {a} {b} { : RawCategory a b} where
open RawCategory
Pp : (x.ident y.ident) I Set (a b)
Pp eqIdent i = {A B : Object}
isEquiv (A B) (A x. B)
isEquiv (A B) (A B)
(λ A≡B
transp
(λ j
@ -191,9 +190,12 @@ record Category (a b : Level) : Set (lsuc (a ⊔ b)) where
raw : RawCategory a b
{{isCategory}} : IsCategory raw
-- What happens if we just open this up to the public?
private
module = RawCategory raw
open RawCategory raw public using ( Isomorphism ; Epimorphism ; Monomorphism )
Object : Set a
Object = .Object