Add notion of pre-category
This commit is contained in:
parent
8276deb4aa
commit
6c5b68a8ac
|
@ -72,16 +72,20 @@ module CatProduct {ℓ ℓ' : Level} (ℂ 𝔻 : Category ℓ ℓ') where
|
|||
= Σ≡ (fst ℂ.isIdentity) (fst 𝔻.isIdentity)
|
||||
, Σ≡ (snd ℂ.isIdentity) (snd 𝔻.isIdentity)
|
||||
|
||||
isPreCategory : IsPreCategory rawProduct
|
||||
IsPreCategory.isAssociative isPreCategory = Σ≡ ℂ.isAssociative 𝔻.isAssociative
|
||||
IsPreCategory.isIdentity isPreCategory = isIdentity
|
||||
IsPreCategory.arrowsAreSets isPreCategory = arrowsAreSets
|
||||
|
||||
postulate univalent : Univalence.Univalent isIdentity
|
||||
instance
|
||||
|
||||
isCategory : IsCategory rawProduct
|
||||
IsCategory.isAssociative isCategory = Σ≡ ℂ.isAssociative 𝔻.isAssociative
|
||||
IsCategory.isIdentity isCategory = isIdentity
|
||||
IsCategory.arrowsAreSets isCategory = arrowsAreSets
|
||||
IsCategory.isPreCategory isCategory = isPreCategory
|
||||
IsCategory.univalent isCategory = univalent
|
||||
|
||||
object : Category ℓ ℓ'
|
||||
Category.raw object = rawProduct
|
||||
Category.isCategory object = isCategory
|
||||
|
||||
fstF : Functor object ℂ
|
||||
fstF = record
|
||||
|
|
|
@ -35,15 +35,15 @@ module _ (ℓa ℓb : Level) where
|
|||
|
||||
open import Cubical.NType.Properties
|
||||
open import Cubical.Sigma
|
||||
instance
|
||||
isCategory : IsCategory RawFam
|
||||
isCategory = record
|
||||
{ isAssociative = λ {A} {B} {C} {D} {f} {g} {h} → isAssociative {A} {B} {C} {D} {f} {g} {h}
|
||||
; isIdentity = λ {A} {B} {f} → isIdentity {A} {B} {f = f}
|
||||
; arrowsAreSets = λ {
|
||||
{((A , hA) , famA)}
|
||||
{((B , hB) , famB)}
|
||||
→ setSig
|
||||
|
||||
isPreCategory : IsPreCategory RawFam
|
||||
IsPreCategory.isAssociative isPreCategory
|
||||
{A} {B} {C} {D} {f} {g} {h} = isAssociative {A} {B} {C} {D} {f} {g} {h}
|
||||
IsPreCategory.isIdentity isPreCategory
|
||||
{A} {B} {f} = isIdentity {A} {B} {f = f}
|
||||
IsPreCategory.arrowsAreSets isPreCategory
|
||||
{(A , hA) , famA} {(B , hB) , famB}
|
||||
= setSig
|
||||
{sA = setPi λ _ → hB}
|
||||
{sB = λ f →
|
||||
let
|
||||
|
@ -55,9 +55,11 @@ module _ (ℓa ℓb : Level) where
|
|||
res = {!!}
|
||||
in res
|
||||
}
|
||||
}
|
||||
; univalent = {!!}
|
||||
}
|
||||
|
||||
isCategory : IsCategory RawFam
|
||||
IsCategory.isPreCategory isCategory = isPreCategory
|
||||
IsCategory.univalent isCategory = {!!}
|
||||
|
||||
Fam : Category (lsuc (ℓa ⊔ ℓb)) (ℓa ⊔ ℓb)
|
||||
Category.raw Fam = RawFam
|
||||
Category.isCategory Fam = isCategory
|
||||
|
|
|
@ -61,6 +61,12 @@ module _ {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where
|
|||
arrowsAreSets : isSet (Path ℂ.Arrow A B)
|
||||
arrowsAreSets a b p q = {!!}
|
||||
|
||||
isPreCategory : IsPreCategory RawFree
|
||||
IsPreCategory.isAssociative isPreCategory {f = f} {g} {h} = isAssociative {r = f} {g} {h}
|
||||
IsPreCategory.isIdentity isPreCategory = isIdentity
|
||||
IsPreCategory.arrowsAreSets isPreCategory = arrowsAreSets
|
||||
|
||||
module _ {A B : ℂ.Object} where
|
||||
eqv : isEquiv (A ≡ B) (A ≅ B) (Univalence.id-to-iso isIdentity A B)
|
||||
eqv = {!!}
|
||||
|
||||
|
@ -68,9 +74,7 @@ module _ {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where
|
|||
univalent = eqv
|
||||
|
||||
isCategory : IsCategory RawFree
|
||||
IsCategory.isAssociative isCategory {f = f} {g} {h} = isAssociative {r = f} {g} {h}
|
||||
IsCategory.isIdentity isCategory = isIdentity
|
||||
IsCategory.arrowsAreSets isCategory = arrowsAreSets
|
||||
IsCategory.isPreCategory isCategory = isPreCategory
|
||||
IsCategory.univalent isCategory = univalent
|
||||
|
||||
Free : Category _ _
|
||||
|
|
|
@ -10,11 +10,11 @@ import Cat.Category.NaturalTransformation
|
|||
|
||||
module Fun {ℓc ℓc' ℓd ℓd' : Level} (ℂ : Category ℓc ℓc') (𝔻 : Category ℓd ℓd') where
|
||||
open NaturalTransformation ℂ 𝔻 public hiding (module Properties)
|
||||
open NaturalTransformation.Properties ℂ 𝔻
|
||||
private
|
||||
module ℂ = Category ℂ
|
||||
module 𝔻 = Category 𝔻
|
||||
|
||||
module _ where
|
||||
-- Functor categories. Objects are functors, arrows are natural transformations.
|
||||
raw : RawCategory (ℓc ⊔ ℓc' ⊔ ℓd ⊔ ℓd') (ℓc ⊔ ℓc' ⊔ ℓd')
|
||||
RawCategory.Object raw = Functor ℂ 𝔻
|
||||
|
@ -22,8 +22,16 @@ module Fun {ℓc ℓc' ℓd ℓd' : Level} (ℂ : Category ℓc ℓc') (𝔻 : C
|
|||
RawCategory.identity raw {F} = identity F
|
||||
RawCategory._∘_ raw {F} {G} {H} = NT[_∘_] {F} {G} {H}
|
||||
|
||||
module _ where
|
||||
open RawCategory raw hiding (identity)
|
||||
open Univalence (λ {A} {B} {f} → isIdentity {F = A} {B} {f})
|
||||
open NaturalTransformation.Properties ℂ 𝔻
|
||||
|
||||
isPreCategory : IsPreCategory raw
|
||||
IsPreCategory.isAssociative isPreCategory {A} {B} {C} {D} = isAssociative {A} {B} {C} {D}
|
||||
IsPreCategory.isIdentity isPreCategory {A} {B} = isIdentity {A} {B}
|
||||
IsPreCategory.arrowsAreSets isPreCategory {F} {G} = naturalTransformationIsSet {F} {G}
|
||||
|
||||
open IsPreCategory isPreCategory hiding (identity)
|
||||
|
||||
module _ (F : Functor ℂ 𝔻) where
|
||||
center : Σ[ G ∈ Object ] (F ≅ G)
|
||||
|
@ -167,17 +175,15 @@ module Fun {ℓc ℓc' ℓd ℓd' : Level} (ℂ : Category ℓc ℓc') (𝔻 : C
|
|||
re-ve : (x : A ≡ B) → reverse (obverse x) ≡ x
|
||||
re-ve = {!!}
|
||||
|
||||
done : isEquiv (A ≡ B) (A ≅ B) (Univalence.id-to-iso (λ { {A} {B} → isIdentity {F = A} {B}}) A B)
|
||||
done : isEquiv (A ≡ B) (A ≅ B) (id-to-iso A B)
|
||||
done = {!gradLemma obverse reverse ve-re re-ve!}
|
||||
|
||||
-- univalent : Univalent
|
||||
-- univalent = done
|
||||
univalent : Univalent
|
||||
univalent = {!done!}
|
||||
|
||||
isCategory : IsCategory raw
|
||||
IsCategory.isAssociative isCategory {A} {B} {C} {D} = isAssociative {A} {B} {C} {D}
|
||||
IsCategory.isIdentity isCategory {A} {B} = isIdentity {A} {B}
|
||||
IsCategory.arrowsAreSets isCategory {F} {G} = naturalTransformationIsSet {F} {G}
|
||||
IsCategory.univalent isCategory = {!!}
|
||||
IsCategory.isPreCategory isCategory = isPreCategory
|
||||
IsCategory.univalent isCategory = univalent
|
||||
|
||||
Fun : Category (ℓc ⊔ ℓc' ⊔ ℓd ⊔ ℓd') (ℓc ⊔ ℓc' ⊔ ℓd')
|
||||
Category.raw Fun = raw
|
||||
|
@ -199,26 +205,26 @@ module _ {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') where
|
|||
; _∘_ = λ {F G H} → NT[_∘_] {F = F} {G = G} {H = H}
|
||||
}
|
||||
|
||||
isCategory : IsCategory raw
|
||||
isCategory = record
|
||||
{ isAssociative =
|
||||
λ{ {A} {B} {C} {D} {f} {g} {h}
|
||||
→ F.isAssociative {A} {B} {C} {D} {f} {g} {h}
|
||||
}
|
||||
; isIdentity =
|
||||
λ{ {A} {B} {f}
|
||||
→ F.isIdentity {A} {B} {f}
|
||||
}
|
||||
; arrowsAreSets =
|
||||
λ{ {A} {B}
|
||||
→ F.arrowsAreSets {A} {B}
|
||||
}
|
||||
; univalent =
|
||||
λ{ {A} {B}
|
||||
→ F.univalent {A} {B}
|
||||
}
|
||||
}
|
||||
-- isCategory : IsCategory raw
|
||||
-- isCategory = record
|
||||
-- { isAssociative =
|
||||
-- λ{ {A} {B} {C} {D} {f} {g} {h}
|
||||
-- → F.isAssociative {A} {B} {C} {D} {f} {g} {h}
|
||||
-- }
|
||||
-- ; isIdentity =
|
||||
-- λ{ {A} {B} {f}
|
||||
-- → F.isIdentity {A} {B} {f}
|
||||
-- }
|
||||
-- ; arrowsAreSets =
|
||||
-- λ{ {A} {B}
|
||||
-- → F.arrowsAreSets {A} {B}
|
||||
-- }
|
||||
-- ; univalent =
|
||||
-- λ{ {A} {B}
|
||||
-- → F.univalent {A} {B}
|
||||
-- }
|
||||
-- }
|
||||
|
||||
Presh : Category (ℓ ⊔ lsuc ℓ') (ℓ ⊔ ℓ')
|
||||
Category.raw Presh = raw
|
||||
Category.isCategory Presh = isCategory
|
||||
-- Presh : Category (ℓ ⊔ lsuc ℓ') (ℓ ⊔ ℓ')
|
||||
-- Category.raw Presh = raw
|
||||
-- Category.isCategory Presh = isCategory
|
||||
|
|
|
@ -157,10 +157,12 @@ RawRel = record
|
|||
; _∘_ = λ {A B C} S R → λ {( a , c ) → Σ[ b ∈ B ] ( (a , b) ∈ R × (b , c) ∈ S )}
|
||||
}
|
||||
|
||||
RawIsCategoryRel : IsCategory RawRel
|
||||
RawIsCategoryRel = record
|
||||
{ isAssociative = funExt is-isAssociative
|
||||
; isIdentity = funExt ident-l , funExt ident-r
|
||||
; arrowsAreSets = {!!}
|
||||
; univalent = {!!}
|
||||
}
|
||||
isPreCategory : IsPreCategory RawRel
|
||||
|
||||
IsPreCategory.isAssociative isPreCategory = funExt is-isAssociative
|
||||
IsPreCategory.isIdentity isPreCategory = funExt ident-l , funExt ident-r
|
||||
IsPreCategory.arrowsAreSets isPreCategory = {!!}
|
||||
|
||||
Rel : PreCategory _ _
|
||||
PreCategory.raw Rel = RawRel
|
||||
PreCategory.isPreCategory Rel = isPreCategory
|
||||
|
|
|
@ -34,17 +34,24 @@ module _ (ℓ : Level) where
|
|||
RawCategory.identity SetsRaw = Function.id
|
||||
RawCategory._∘_ SetsRaw = Function._∘′_
|
||||
|
||||
module _ where
|
||||
private
|
||||
open RawCategory SetsRaw hiding (_∘_)
|
||||
|
||||
isIdentity : IsIdentity Function.id
|
||||
fst isIdentity = funExt λ _ → refl
|
||||
snd isIdentity = funExt λ _ → refl
|
||||
|
||||
open Univalence (λ {A} {B} {f} → isIdentity {A} {B} {f})
|
||||
|
||||
arrowsAreSets : ArrowsAreSets
|
||||
arrowsAreSets {B = (_ , s)} = setPi λ _ → s
|
||||
|
||||
isPreCat : IsPreCategory SetsRaw
|
||||
IsPreCategory.isAssociative isPreCat = refl
|
||||
IsPreCategory.isIdentity isPreCat {A} {B} = isIdentity {A} {B}
|
||||
IsPreCategory.arrowsAreSets isPreCat {A} {B} = arrowsAreSets {A} {B}
|
||||
|
||||
open IsPreCategory isPreCat hiding (_∘_)
|
||||
|
||||
isIso = Eqv.Isomorphism
|
||||
module _ {hA hB : hSet ℓ} where
|
||||
open Σ hA renaming (fst to A ; snd to sA)
|
||||
|
@ -255,9 +262,7 @@ module _ (ℓ : Level) where
|
|||
univalent = from[Contr] univalent[Contr]
|
||||
|
||||
SetsIsCategory : IsCategory SetsRaw
|
||||
IsCategory.isAssociative SetsIsCategory = refl
|
||||
IsCategory.isIdentity SetsIsCategory {A} {B} = isIdentity {A} {B}
|
||||
IsCategory.arrowsAreSets SetsIsCategory {A} {B} = arrowsAreSets {A} {B}
|
||||
IsCategory.isPreCategory SetsIsCategory = isPreCat
|
||||
IsCategory.univalent SetsIsCategory = univalent
|
||||
|
||||
𝓢𝓮𝓽 Sets : Category (lsuc ℓ) ℓ
|
||||
|
|
|
@ -203,15 +203,13 @@ record RawCategory (ℓa ℓb : Level) : Set (lsuc (ℓa ⊔ ℓb)) where
|
|||
--
|
||||
-- Sans `univalent` this would be what is referred to as a pre-category in
|
||||
-- [HoTT].
|
||||
record IsCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc (ℓa ⊔ ℓb)) where
|
||||
record IsPreCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc (ℓa ⊔ ℓb)) where
|
||||
open RawCategory ℂ public
|
||||
field
|
||||
isAssociative : IsAssociative
|
||||
isIdentity : IsIdentity identity
|
||||
arrowsAreSets : ArrowsAreSets
|
||||
open Univalence isIdentity public
|
||||
field
|
||||
univalent : Univalent
|
||||
|
||||
leftIdentity : {A B : Object} {f : Arrow A B} → identity ∘ f ≡ f
|
||||
leftIdentity {A} {B} {f} = fst (isIdentity {A = A} {B} {f})
|
||||
|
@ -251,19 +249,6 @@ record IsCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc
|
|||
iso→epi×mono : Isomorphism f → Epimorphism {X = X} f × Monomorphism {X = X} f
|
||||
iso→epi×mono iso = iso→epi iso , iso→mono iso
|
||||
|
||||
-- | The formulation of univalence expressed with _≃_ is trivially admissable -
|
||||
-- just "forget" the equivalence.
|
||||
univalent≃ : Univalent≃
|
||||
univalent≃ = _ , univalent
|
||||
|
||||
module _ {A B : Object} where
|
||||
open import Cat.Equivalence using (module Equiv≃)
|
||||
|
||||
iso-to-id : (A ≅ B) → (A ≡ B)
|
||||
iso-to-id = fst (Equiv≃.toIso _ _ univalent)
|
||||
|
||||
-- | All projections are propositions.
|
||||
module Propositionality where
|
||||
propIsAssociative : isProp IsAssociative
|
||||
propIsAssociative x y i = arrowsAreSets _ _ x y i
|
||||
|
||||
|
@ -298,8 +283,20 @@ record IsCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc
|
|||
identity ∘ g' ≡⟨ leftIdentity ⟩
|
||||
g' ∎
|
||||
|
||||
propUnivalent : isProp Univalent
|
||||
propUnivalent a b i = propPi (λ iso → propIsContr) a b i
|
||||
propIsInitial : ∀ I → isProp (IsInitial I)
|
||||
propIsInitial I x y i {X} = res X i
|
||||
where
|
||||
module _ (X : Object) where
|
||||
open Σ (x {X}) renaming (fst to fx ; snd to cx)
|
||||
open Σ (y {X}) renaming (fst to fy ; snd to cy)
|
||||
fp : fx ≡ fy
|
||||
fp = cx fy
|
||||
prop : (x : Arrow I X) → isProp (∀ f → x ≡ f)
|
||||
prop x = propPi (λ y → arrowsAreSets x y)
|
||||
cp : (λ i → ∀ f → fp i ≡ f) [ cx ≡ cy ]
|
||||
cp = lemPropF prop fp
|
||||
res : (fx , cx) ≡ (fy , cy)
|
||||
res i = fp i , cp i
|
||||
|
||||
propIsTerminal : ∀ T → isProp (IsTerminal T)
|
||||
propIsTerminal T x y i {X} = res X i
|
||||
|
@ -316,6 +313,35 @@ record IsCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc
|
|||
res : (fx , cx) ≡ (fy , cy)
|
||||
res i = fp i , cp i
|
||||
|
||||
record PreCategory (ℓa ℓb : Level) : Set (lsuc (ℓa ⊔ ℓb)) where
|
||||
field
|
||||
raw : RawCategory ℓa ℓb
|
||||
isPreCategory : IsPreCategory raw
|
||||
open IsPreCategory isPreCategory public
|
||||
|
||||
record IsCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc (ℓa ⊔ ℓb)) where
|
||||
field
|
||||
isPreCategory : IsPreCategory ℂ
|
||||
open IsPreCategory isPreCategory public
|
||||
field
|
||||
univalent : Univalent
|
||||
|
||||
-- | The formulation of univalence expressed with _≃_ is trivially admissable -
|
||||
-- just "forget" the equivalence.
|
||||
univalent≃ : Univalent≃
|
||||
univalent≃ = _ , univalent
|
||||
|
||||
module _ {A B : Object} where
|
||||
open import Cat.Equivalence using (module Equiv≃)
|
||||
|
||||
iso-to-id : (A ≅ B) → (A ≡ B)
|
||||
iso-to-id = fst (Equiv≃.toIso _ _ univalent)
|
||||
|
||||
-- | All projections are propositions.
|
||||
module Propositionality where
|
||||
propUnivalent : isProp Univalent
|
||||
propUnivalent a b i = propPi (λ iso → propIsContr) a b i
|
||||
|
||||
-- | Terminal objects are propositional - a.k.a uniqueness of terminal
|
||||
-- | objects.
|
||||
--
|
||||
|
@ -353,20 +379,6 @@ record IsCategory {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) : Set (lsuc
|
|||
res i = p0 i , p1 i
|
||||
|
||||
-- Merely the dual of the above statement.
|
||||
propIsInitial : ∀ I → isProp (IsInitial I)
|
||||
propIsInitial I x y i {X} = res X i
|
||||
where
|
||||
module _ (X : Object) where
|
||||
open Σ (x {X}) renaming (fst to fx ; snd to cx)
|
||||
open Σ (y {X}) renaming (fst to fy ; snd to cy)
|
||||
fp : fx ≡ fy
|
||||
fp = cx fy
|
||||
prop : (x : Arrow I X) → isProp (∀ f → x ≡ f)
|
||||
prop x = propPi (λ y → arrowsAreSets x y)
|
||||
cp : (λ i → ∀ f → fp i ≡ f) [ cx ≡ cy ]
|
||||
cp = lemPropF prop fp
|
||||
res : (fx , cx) ≡ (fy , cy)
|
||||
res i = fp i , cp i
|
||||
|
||||
propInitial : isProp Initial
|
||||
propInitial Xi Yi = res
|
||||
|
@ -404,6 +416,23 @@ module _ {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) where
|
|||
open RawCategory ℂ
|
||||
open Univalence
|
||||
private
|
||||
module _ (x y : IsPreCategory ℂ) where
|
||||
module x = IsPreCategory x
|
||||
module y = IsPreCategory y
|
||||
-- In a few places I use the result of propositionality of the various
|
||||
-- projections of `IsCategory` - Here I arbitrarily chose to use this
|
||||
-- result from `x : IsCategory C`. I don't know which (if any) possibly
|
||||
-- adverse effects this may have.
|
||||
-- module Prop = X.Propositionality
|
||||
|
||||
propIsPreCategory : x ≡ y
|
||||
IsPreCategory.isAssociative (propIsPreCategory i)
|
||||
= x.propIsAssociative x.isAssociative y.isAssociative i
|
||||
IsPreCategory.isIdentity (propIsPreCategory i)
|
||||
= x.propIsIdentity x.isIdentity y.isIdentity i
|
||||
IsPreCategory.arrowsAreSets (propIsPreCategory i)
|
||||
= x.propArrowIsSet x.arrowsAreSets y.arrowsAreSets i
|
||||
|
||||
module _ (x y : IsCategory ℂ) where
|
||||
module X = IsCategory x
|
||||
module Y = IsCategory y
|
||||
|
@ -414,7 +443,7 @@ module _ {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) where
|
|||
module Prop = X.Propositionality
|
||||
|
||||
isIdentity : (λ _ → IsIdentity identity) [ X.isIdentity ≡ Y.isIdentity ]
|
||||
isIdentity = Prop.propIsIdentity X.isIdentity Y.isIdentity
|
||||
isIdentity = X.propIsIdentity X.isIdentity Y.isIdentity
|
||||
|
||||
U : ∀ {a : IsIdentity identity}
|
||||
→ (λ _ → IsIdentity identity) [ X.isIdentity ≡ a ]
|
||||
|
@ -434,10 +463,13 @@ module _ {ℓa ℓb : Level} (ℂ : RawCategory ℓa ℓb) where
|
|||
eqUni : U isIdentity Y.univalent
|
||||
eqUni = helper Y.univalent
|
||||
done : x ≡ y
|
||||
IsCategory.isAssociative (done i) = Prop.propIsAssociative X.isAssociative Y.isAssociative i
|
||||
IsCategory.isIdentity (done i) = isIdentity i
|
||||
IsCategory.arrowsAreSets (done i) = Prop.propArrowIsSet X.arrowsAreSets Y.arrowsAreSets i
|
||||
IsCategory.isPreCategory (done i)
|
||||
= propIsPreCategory X.isPreCategory Y.isPreCategory i
|
||||
IsCategory.univalent (done i) = eqUni i
|
||||
-- IsCategory.isAssociative (done i) = Prop.propIsAssociative X.isAssociative Y.isAssociative i
|
||||
-- IsCategory.isIdentity (done i) = isIdentity i
|
||||
-- IsCategory.arrowsAreSets (done i) = Prop.propArrowIsSet X.arrowsAreSets Y.arrowsAreSets i
|
||||
-- IsCategory.univalent (done i) = eqUni i
|
||||
|
||||
propIsCategory : isProp (IsCategory ℂ)
|
||||
propIsCategory = done
|
||||
|
@ -486,6 +518,7 @@ module _ {ℓa ℓb : Level} (ℂ : Category ℓa ℓb) where
|
|||
module Opposite {ℓa ℓb : Level} where
|
||||
module _ (ℂ : Category ℓa ℓb) where
|
||||
private
|
||||
module _ where
|
||||
module ℂ = Category ℂ
|
||||
opRaw : RawCategory ℓa ℓb
|
||||
RawCategory.Object opRaw = ℂ.Object
|
||||
|
@ -498,7 +531,12 @@ module Opposite {ℓa ℓb : Level} where
|
|||
isIdentity : IsIdentity identity
|
||||
isIdentity = swap ℂ.isIdentity
|
||||
|
||||
open Univalence isIdentity
|
||||
isPreCategory : IsPreCategory opRaw
|
||||
IsPreCategory.isAssociative isPreCategory = sym ℂ.isAssociative
|
||||
IsPreCategory.isIdentity isPreCategory = isIdentity
|
||||
IsPreCategory.arrowsAreSets isPreCategory = ℂ.arrowsAreSets
|
||||
|
||||
open IsPreCategory isPreCategory
|
||||
|
||||
module _ {A B : ℂ.Object} where
|
||||
open import Cat.Equivalence as Equivalence hiding (_≅_)
|
||||
|
@ -571,9 +609,7 @@ module Opposite {ℓa ℓb : Level} where
|
|||
univalent = Equiv≃.fromIso _ _ h
|
||||
|
||||
isCategory : IsCategory opRaw
|
||||
IsCategory.isAssociative isCategory = sym ℂ.isAssociative
|
||||
IsCategory.isIdentity isCategory = isIdentity
|
||||
IsCategory.arrowsAreSets isCategory = ℂ.arrowsAreSets
|
||||
IsCategory.isPreCategory isCategory = isPreCategory
|
||||
IsCategory.univalent isCategory = univalent
|
||||
|
||||
opposite : Category ℓa ℓb
|
||||
|
|
|
@ -122,14 +122,15 @@ module Try0 {ℓa ℓb : Level} {ℂ : Category ℓa ℓb}
|
|||
}
|
||||
}
|
||||
|
||||
module _ where
|
||||
open RawCategory raw
|
||||
|
||||
propEqs : ∀ {X' : Object}{Y' : Object} (let X , xa , xb = X') (let Y , ya , yb = Y')
|
||||
→ (xy : ℂ.Arrow X Y) → isProp (ℂ [ ya ∘ xy ] ≡ xa × ℂ [ yb ∘ xy ] ≡ xb)
|
||||
propEqs xs = propSig (ℂ.arrowsAreSets _ _) (\ _ → ℂ.arrowsAreSets _ _)
|
||||
|
||||
isAssocitaive : IsAssociative
|
||||
isAssocitaive {A'@(A , a0 , a1)} {B , _} {C , c0 , c1} {D'@(D , d0 , d1)} {ff@(f , f0 , f1)} {gg@(g , g0 , g1)} {hh@(h , h0 , h1)} i
|
||||
isAssociative : IsAssociative
|
||||
isAssociative {A'@(A , a0 , a1)} {B , _} {C , c0 , c1} {D'@(D , d0 , d1)} {ff@(f , f0 , f1)} {gg@(g , g0 , g1)} {hh@(h , h0 , h1)} i
|
||||
= s0 i , lemPropF propEqs s0 {P.snd l} {P.snd r} i
|
||||
where
|
||||
l = hh ∘ (gg ∘ ff)
|
||||
|
@ -163,7 +164,12 @@ module Try0 {ℓa ℓb : Level} {ℂ : Category ℓa ℓb}
|
|||
arrowsAreSets {X , x0 , x1} {Y , y0 , y1}
|
||||
= sigPresNType {n = ⟨0⟩} ℂ.arrowsAreSets λ a → propSet (propEqs _)
|
||||
|
||||
open Univalence isIdentity
|
||||
isPreCat : IsPreCategory raw
|
||||
IsPreCategory.isAssociative isPreCat = isAssociative
|
||||
IsPreCategory.isIdentity isPreCat = isIdentity
|
||||
IsPreCategory.arrowsAreSets isPreCat = arrowsAreSets
|
||||
|
||||
open IsPreCategory isPreCat
|
||||
|
||||
-- module _ (X : Object) where
|
||||
-- center : Σ Object (X ≅_)
|
||||
|
@ -327,12 +333,8 @@ module Try0 {ℓa ℓb : Level} {ℂ : Category ℓa ℓb}
|
|||
res = Equiv≃.fromIso _ _ iso
|
||||
|
||||
isCat : IsCategory raw
|
||||
isCat = record
|
||||
{ isAssociative = isAssocitaive
|
||||
; isIdentity = isIdentity
|
||||
; arrowsAreSets = arrowsAreSets
|
||||
; univalent = univalent
|
||||
}
|
||||
IsCategory.isPreCategory isCat = isPreCat
|
||||
IsCategory.univalent isCat = univalent
|
||||
|
||||
cat : Category _ _
|
||||
cat = record
|
||||
|
|
Loading…
Reference in a new issue