Make level-parameters to Category explicit
This commit is contained in:
parent
0990a3778f
commit
07e4269399
|
@ -20,7 +20,7 @@ snd (lift-eq a b i) = b i
|
||||||
|
|
||||||
open Functor
|
open Functor
|
||||||
open Category
|
open Category
|
||||||
module _ {ℓ ℓ' : Level} {A B : Category {ℓ} {ℓ'}} where
|
module _ {ℓ ℓ' : Level} {A B : Category ℓ ℓ'} where
|
||||||
lift-eq-functors : {f g : Functor A B}
|
lift-eq-functors : {f g : Functor A B}
|
||||||
→ (eq* : Functor.func* f ≡ Functor.func* g)
|
→ (eq* : Functor.func* f ≡ Functor.func* g)
|
||||||
→ (eq→ : PathP (λ i → ∀ {x y} → Arrow A x y → Arrow B (eq* i x) (eq* i y))
|
→ (eq→ : PathP (λ i → ∀ {x y} → Arrow A x y → Arrow B (eq* i x) (eq* i y))
|
||||||
|
@ -41,11 +41,11 @@ module _ {ℓ ℓ' : Level} {A B : Category {ℓ} {ℓ'}} where
|
||||||
module _ {ℓ ℓ' : Level} where
|
module _ {ℓ ℓ' : Level} where
|
||||||
private
|
private
|
||||||
_⊛_ = functor-comp
|
_⊛_ = functor-comp
|
||||||
module _ {A B C D : Category {ℓ} {ℓ'}} {f : Functor A B} {g : Functor B C} {h : Functor C D} where
|
module _ {A B C D : Category ℓ ℓ'} {f : Functor A B} {g : Functor B C} {h : Functor C D} where
|
||||||
postulate assc : h ⊛ (g ⊛ f) ≡ (h ⊛ g) ⊛ f
|
postulate assc : h ⊛ (g ⊛ f) ≡ (h ⊛ g) ⊛ f
|
||||||
-- assc = lift-eq-functors refl refl {!refl!} λ i j → {!!}
|
-- assc = lift-eq-functors refl refl {!refl!} λ i j → {!!}
|
||||||
|
|
||||||
module _ {A B : Category {ℓ} {ℓ'}} {f : Functor A B} where
|
module _ {A B : Category ℓ ℓ'} {f : Functor A B} where
|
||||||
lem : (func* f) ∘ (func* (identity {C = A})) ≡ func* f
|
lem : (func* f) ∘ (func* (identity {C = A})) ≡ func* f
|
||||||
lem = refl
|
lem = refl
|
||||||
-- lemmm : func→ {C = A} {D = B} (f ⊛ identity) ≡ func→ f
|
-- lemmm : func→ {C = A} {D = B} (f ⊛ identity) ≡ func→ f
|
||||||
|
@ -62,10 +62,10 @@ module _ {ℓ ℓ' : Level} where
|
||||||
postulate ident-l : identity ⊛ f ≡ f
|
postulate ident-l : identity ⊛ f ≡ f
|
||||||
-- ident-l = lift-eq-functors lem lemmm {!refl!} {!!}
|
-- ident-l = lift-eq-functors lem lemmm {!refl!} {!!}
|
||||||
|
|
||||||
CatCat : Category {lsuc (ℓ ⊔ ℓ')} {ℓ ⊔ ℓ'}
|
CatCat : Category (lsuc (ℓ ⊔ ℓ')) (ℓ ⊔ ℓ')
|
||||||
CatCat =
|
CatCat =
|
||||||
record
|
record
|
||||||
{ Object = Category {ℓ} {ℓ'}
|
{ Object = Category ℓ ℓ'
|
||||||
; Arrow = Functor
|
; Arrow = Functor
|
||||||
; 𝟙 = identity
|
; 𝟙 = identity
|
||||||
; _⊕_ = functor-comp
|
; _⊕_ = functor-comp
|
||||||
|
@ -74,7 +74,7 @@ module _ {ℓ ℓ' : Level} where
|
||||||
; ident = ident-r , ident-l
|
; ident = ident-r , ident-l
|
||||||
}
|
}
|
||||||
|
|
||||||
module _ {ℓ : Level} (C D : Category {ℓ} {ℓ}) where
|
module _ {ℓ : Level} (C D : Category ℓ ℓ) where
|
||||||
private
|
private
|
||||||
proj₁ : Arrow CatCat (catProduct C D) C
|
proj₁ : Arrow CatCat (catProduct C D) C
|
||||||
proj₁ = record { func* = fst ; func→ = fst ; ident = refl ; distrib = refl }
|
proj₁ = record { func* = fst ; func→ = fst ; ident = refl ; distrib = refl }
|
||||||
|
|
|
@ -154,7 +154,7 @@ module _ {A B C D : Set} {S : Subset (A × B)} {R : Subset (B × C)} {Q : Subset
|
||||||
≡ (Σ[ b ∈ B ] (a , b) ∈ S × (Σ[ c ∈ C ] (b , c) ∈ R × (c , d) ∈ Q))
|
≡ (Σ[ b ∈ B ] (a , b) ∈ S × (Σ[ c ∈ C ] (b , c) ∈ R × (c , d) ∈ Q))
|
||||||
is-assoc = equivToPath equi
|
is-assoc = equivToPath equi
|
||||||
|
|
||||||
Rel : Category
|
Rel : Category (lsuc lzero) (lsuc lzero)
|
||||||
Rel = record
|
Rel = record
|
||||||
{ Object = Set
|
{ Object = Set
|
||||||
; Arrow = λ S R → Subset (S × R)
|
; Arrow = λ S R → Subset (S × R)
|
||||||
|
|
|
@ -15,7 +15,7 @@ open import Cat.Functor
|
||||||
Fun : {ℓ : Level} → ( T U : Set ℓ ) → Set ℓ
|
Fun : {ℓ : Level} → ( T U : Set ℓ ) → Set ℓ
|
||||||
Fun T U = T → U
|
Fun T U = T → U
|
||||||
|
|
||||||
Sets : {ℓ : Level} → Category {lsuc ℓ} {ℓ}
|
Sets : {ℓ : Level} → Category (lsuc ℓ) ℓ
|
||||||
Sets {ℓ} = record
|
Sets {ℓ} = record
|
||||||
{ Object = Set ℓ
|
{ Object = Set ℓ
|
||||||
; Arrow = λ T U → Fun {ℓ} T U
|
; Arrow = λ T U → Fun {ℓ} T U
|
||||||
|
@ -26,11 +26,11 @@ Sets {ℓ} = record
|
||||||
}
|
}
|
||||||
|
|
||||||
-- Covariant Presheaf
|
-- Covariant Presheaf
|
||||||
Representable : {ℓ ℓ' : Level} → (ℂ : Category {ℓ} {ℓ'}) → Set (ℓ ⊔ lsuc ℓ')
|
Representable : {ℓ ℓ' : Level} → (ℂ : Category ℓ ℓ') → Set (ℓ ⊔ lsuc ℓ')
|
||||||
Representable {ℓ' = ℓ'} ℂ = Functor ℂ (Sets {ℓ'})
|
Representable {ℓ' = ℓ'} ℂ = Functor ℂ (Sets {ℓ'})
|
||||||
|
|
||||||
-- The "co-yoneda" embedding.
|
-- The "co-yoneda" embedding.
|
||||||
representable : {ℓ ℓ' : Level} → {ℂ : Category {ℓ} {ℓ'}} → Category.Object ℂ → Representable ℂ
|
representable : ∀ {ℓ ℓ'} {ℂ : Category ℓ ℓ'} → Category.Object ℂ → Representable ℂ
|
||||||
representable {ℂ = ℂ} A = record
|
representable {ℂ = ℂ} A = record
|
||||||
{ func* = λ B → ℂ.Arrow A B
|
{ func* = λ B → ℂ.Arrow A B
|
||||||
; func→ = λ f g → f ℂ.⊕ g
|
; func→ = λ f g → f ℂ.⊕ g
|
||||||
|
@ -41,11 +41,11 @@ representable {ℂ = ℂ} A = record
|
||||||
open module ℂ = Category ℂ
|
open module ℂ = Category ℂ
|
||||||
|
|
||||||
-- Contravariant Presheaf
|
-- Contravariant Presheaf
|
||||||
Presheaf : {ℓ ℓ' : Level} → (ℂ : Category {ℓ} {ℓ'}) → Set (ℓ ⊔ lsuc ℓ')
|
Presheaf : ∀ {ℓ ℓ'} (ℂ : Category ℓ ℓ') → Set (ℓ ⊔ lsuc ℓ')
|
||||||
Presheaf {ℓ' = ℓ'} ℂ = Functor (Opposite ℂ) (Sets {ℓ'})
|
Presheaf {ℓ' = ℓ'} ℂ = Functor (Opposite ℂ) (Sets {ℓ'})
|
||||||
|
|
||||||
-- Alternate name: `yoneda`
|
-- Alternate name: `yoneda`
|
||||||
presheaf : {ℓ ℓ' : Level} → {ℂ : Category {ℓ} {ℓ'}} → Category.Object (Opposite ℂ) → Presheaf ℂ
|
presheaf : {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} → Category.Object (Opposite ℂ) → Presheaf ℂ
|
||||||
presheaf {ℂ = ℂ} B = record
|
presheaf {ℂ = ℂ} B = record
|
||||||
{ func* = λ A → ℂ.Arrow A B
|
{ func* = λ A → ℂ.Arrow A B
|
||||||
; func→ = λ f g → g ℂ.⊕ f
|
; func→ = λ f g → g ℂ.⊕ f
|
||||||
|
|
|
@ -24,7 +24,7 @@ syntax ∃!-syntax (λ x → B) = ∃![ x ] B
|
||||||
|
|
||||||
postulate undefined : {ℓ : Level} → {A : Set ℓ} → A
|
postulate undefined : {ℓ : Level} → {A : Set ℓ} → A
|
||||||
|
|
||||||
record Category {ℓ ℓ'} : 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.
|
||||||
no-eta-equality
|
no-eta-equality
|
||||||
field
|
field
|
||||||
|
@ -44,7 +44,7 @@ record Category {ℓ ℓ'} : Set (lsuc (ℓ' ⊔ ℓ)) where
|
||||||
|
|
||||||
open Category public
|
open Category public
|
||||||
|
|
||||||
module _ {ℓ ℓ' : Level} {ℂ : Category {ℓ} {ℓ'}} { A B : ℂ .Object } where
|
module _ {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} { A B : ℂ .Object } where
|
||||||
private
|
private
|
||||||
open module ℂ = Category ℂ
|
open module ℂ = Category ℂ
|
||||||
_+_ = ℂ._⊕_
|
_+_ = ℂ._⊕_
|
||||||
|
@ -93,10 +93,10 @@ epi-mono-is-not-iso f =
|
||||||
-}
|
-}
|
||||||
|
|
||||||
-- Isomorphism of objects
|
-- Isomorphism of objects
|
||||||
_≅_ : { ℓ ℓ' : Level } → { ℂ : Category {ℓ} {ℓ'} } → ( A B : Object ℂ ) → Set ℓ'
|
_≅_ : ∀ {ℓ ℓ'} {ℂ : Category ℓ ℓ'} (A B : Object ℂ) → Set ℓ'
|
||||||
_≅_ {ℂ = ℂ} A B = Σ[ f ∈ ℂ .Arrow A B ] (Isomorphism {ℂ = ℂ} f)
|
_≅_ {ℂ = ℂ} A B = Σ[ f ∈ ℂ .Arrow A B ] (Isomorphism {ℂ = ℂ} f)
|
||||||
|
|
||||||
IsProduct : ∀ {ℓ ℓ'} (ℂ : Category {ℓ} {ℓ'}) {A B obj : Object ℂ} (π₁ : Arrow ℂ obj A) (π₂ : Arrow ℂ obj B) → Set (ℓ ⊔ ℓ')
|
IsProduct : ∀ {ℓ ℓ'} (ℂ : Category ℓ ℓ') {A B obj : Object ℂ} (π₁ : Arrow ℂ obj A) (π₂ : Arrow ℂ obj B) → Set (ℓ ⊔ ℓ')
|
||||||
IsProduct ℂ {A = A} {B = B} π₁ π₂
|
IsProduct ℂ {A = A} {B = B} π₁ π₂
|
||||||
= ∀ {X : ℂ.Object} (x₁ : ℂ.Arrow X A) (x₂ : ℂ.Arrow X B)
|
= ∀ {X : ℂ.Object} (x₁ : ℂ.Arrow X A) (x₂ : ℂ.Arrow X B)
|
||||||
→ ∃![ x ] (π₁ ℂ.⊕ x ≡ x₁ × π₂ ℂ.⊕ x ≡ x₂)
|
→ ∃![ x ] (π₁ ℂ.⊕ x ≡ x₁ × π₂ ℂ.⊕ x ≡ x₂)
|
||||||
|
@ -110,7 +110,7 @@ IsProduct ℂ {A = A} {B = B} π₁ π₂
|
||||||
-- isProduct : ∀ {X : ℂ .Object} (x₁ : ℂ .Arrow X A) (x₂ : ℂ .Arrow X B)
|
-- isProduct : ∀ {X : ℂ .Object} (x₁ : ℂ .Arrow X A) (x₂ : ℂ .Arrow X B)
|
||||||
-- → ∃![ x ] (ℂ ._⊕_ π₁ x ≡ x₁ × ℂ. _⊕_ π₂ x ≡ x₂)
|
-- → ∃![ x ] (ℂ ._⊕_ π₁ x ≡ x₁ × ℂ. _⊕_ π₂ x ≡ x₂)
|
||||||
|
|
||||||
record Product {ℓ ℓ' : Level} {ℂ : Category {ℓ} {ℓ'}} (A B : ℂ .Object) : Set (ℓ ⊔ ℓ') where
|
record Product {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} (A B : ℂ .Object) : Set (ℓ ⊔ ℓ') where
|
||||||
no-eta-equality
|
no-eta-equality
|
||||||
field
|
field
|
||||||
obj : ℂ .Object
|
obj : ℂ .Object
|
||||||
|
@ -119,7 +119,7 @@ record Product {ℓ ℓ' : Level} {ℂ : Category {ℓ} {ℓ'}} (A B : ℂ .Obje
|
||||||
{{isProduct}} : IsProduct ℂ proj₁ proj₂
|
{{isProduct}} : IsProduct ℂ proj₁ proj₂
|
||||||
|
|
||||||
mutual
|
mutual
|
||||||
catProduct : {ℓ : Level} → (C D : Category {ℓ} {ℓ}) → Category {ℓ} {ℓ}
|
catProduct : ∀ {ℓ} (C D : Category ℓ ℓ) → Category ℓ ℓ
|
||||||
catProduct C D =
|
catProduct C D =
|
||||||
record
|
record
|
||||||
{ Object = C.Object × D.Object
|
{ Object = C.Object × D.Object
|
||||||
|
@ -146,10 +146,10 @@ mutual
|
||||||
-- arrowProduct = {!!}
|
-- arrowProduct = {!!}
|
||||||
|
|
||||||
-- Arrows in the product-category
|
-- Arrows in the product-category
|
||||||
arrowProduct : ∀ {ℓ} {C D : Category {ℓ} {ℓ}} (c d : Object (catProduct C D)) → Set ℓ
|
arrowProduct : ∀ {ℓ} {C D : Category ℓ ℓ} (c d : Object (catProduct C D)) → Set ℓ
|
||||||
arrowProduct {C = C} {D = D} (c , d) (c' , d') = Arrow C c c' × Arrow D d d'
|
arrowProduct {C = C} {D = D} (c , d) (c' , d') = Arrow C c c' × Arrow D d d'
|
||||||
|
|
||||||
Opposite : ∀ {ℓ ℓ'} → Category {ℓ} {ℓ'} → Category {ℓ} {ℓ'}
|
Opposite : ∀ {ℓ ℓ'} → Category ℓ ℓ' → Category ℓ ℓ'
|
||||||
Opposite ℂ =
|
Opposite ℂ =
|
||||||
record
|
record
|
||||||
{ Object = ℂ.Object
|
{ Object = ℂ.Object
|
||||||
|
@ -173,10 +173,10 @@ Opposite ℂ =
|
||||||
-- assoc (Opposite-is-involution i) = {!!}
|
-- assoc (Opposite-is-involution i) = {!!}
|
||||||
-- ident (Opposite-is-involution i) = {!!}
|
-- ident (Opposite-is-involution i) = {!!}
|
||||||
|
|
||||||
Hom : {ℓ ℓ' : Level} → (ℂ : Category {ℓ} {ℓ'}) → (A B : Object ℂ) → Set ℓ'
|
Hom : {ℓ ℓ' : Level} → (ℂ : Category ℓ ℓ') → (A B : Object ℂ) → Set ℓ'
|
||||||
Hom ℂ A B = Arrow ℂ A B
|
Hom ℂ A B = Arrow ℂ A B
|
||||||
|
|
||||||
module _ {ℓ ℓ' : Level} {ℂ : Category {ℓ} {ℓ'}} where
|
module _ {ℓ ℓ' : Level} {ℂ : Category ℓ ℓ'} where
|
||||||
HomFromArrow : (A : ℂ .Object) → {B B' : ℂ .Object} → (g : ℂ .Arrow B B')
|
HomFromArrow : (A : ℂ .Object) → {B B' : ℂ .Object} → (g : ℂ .Arrow B B')
|
||||||
→ Hom ℂ A B → Hom ℂ A B'
|
→ Hom ℂ A B → Hom ℂ A B'
|
||||||
HomFromArrow _A = _⊕_ ℂ
|
HomFromArrow _A = _⊕_ ℂ
|
||||||
|
|
|
@ -6,12 +6,12 @@ open import Data.Product
|
||||||
|
|
||||||
open import Cat.Category as C
|
open import Cat.Category as C
|
||||||
|
|
||||||
module _ {ℓ ℓ' : Level} (ℂ : Category {ℓ} {ℓ'}) where
|
module _ {ℓ ℓ' : Level} (ℂ : Category ℓ ℓ') where
|
||||||
private
|
private
|
||||||
open module ℂ = Category ℂ
|
open module ℂ = Category ℂ
|
||||||
Obj = ℂ.Object
|
Obj = ℂ.Object
|
||||||
|
|
||||||
Path : ( a b : Obj ) → Set
|
Path : ( a b : Obj ) → Set ℓ'
|
||||||
Path a b = undefined
|
Path a b = undefined
|
||||||
|
|
||||||
postulate emptyPath : (o : Obj) → Path o o
|
postulate emptyPath : (o : Obj) → Path o o
|
||||||
|
@ -28,7 +28,7 @@ module _ {ℓ ℓ' : Level} (ℂ : Category {ℓ} {ℓ'}) where
|
||||||
ident-r : concatenate {A} {A} {B} p (emptyPath A) ≡ p
|
ident-r : concatenate {A} {A} {B} p (emptyPath A) ≡ p
|
||||||
ident-l : concatenate {A} {B} {B} (emptyPath B) p ≡ p
|
ident-l : concatenate {A} {B} {B} (emptyPath B) p ≡ p
|
||||||
|
|
||||||
Free : Category
|
Free : Category ℓ ℓ'
|
||||||
Free = record
|
Free = record
|
||||||
{ Object = Obj
|
{ Object = Obj
|
||||||
; Arrow = Path
|
; Arrow = Path
|
||||||
|
|
|
@ -7,7 +7,7 @@ open import Cat.Functor
|
||||||
open import Cat.Categories.Sets
|
open import Cat.Categories.Sets
|
||||||
|
|
||||||
module _ {ℓa ℓa' ℓb ℓb'} where
|
module _ {ℓa ℓa' ℓb ℓb'} where
|
||||||
Exponential : Category {ℓa} {ℓa'} → Category {ℓb} {ℓb'} → Category {{!!}} {{!!}}
|
Exponential : Category ℓa ℓa' → Category ℓb ℓb' → Category ? ?
|
||||||
Exponential A B = record
|
Exponential A B = record
|
||||||
{ Object = {!!}
|
{ Object = {!!}
|
||||||
; Arrow = {!!}
|
; Arrow = {!!}
|
||||||
|
@ -19,5 +19,5 @@ module _ {ℓa ℓa' ℓb ℓb'} where
|
||||||
|
|
||||||
_⇑_ = Exponential
|
_⇑_ = Exponential
|
||||||
|
|
||||||
yoneda : ∀ {ℓ ℓ'} → {ℂ : Category {ℓ} {ℓ'}} → Functor ℂ (Sets ⇑ (Opposite ℂ))
|
yoneda : ∀ {ℓ ℓ'} → {ℂ : Category ℓ ℓ'} → Functor ℂ (Sets ⇑ (Opposite ℂ))
|
||||||
yoneda = {!!}
|
yoneda = {!!}
|
||||||
|
|
|
@ -42,7 +42,7 @@ module _ {ℓ ℓ' : Level} (Ns : Set ℓ) where
|
||||||
Mor = Σ themap rules
|
Mor = Σ themap rules
|
||||||
|
|
||||||
-- The category of names and substitutions
|
-- The category of names and substitutions
|
||||||
ℂ : Category -- {ℓo} {lsuc lzero ⊔ ℓo}
|
ℂ : Category ℓ ℓ -- ℓo (lsuc lzero ⊔ ℓo)
|
||||||
ℂ = record
|
ℂ = record
|
||||||
-- { Object = FiniteDecidableSubset
|
-- { Object = FiniteDecidableSubset
|
||||||
{ Object = Ns → Bool
|
{ Object = Ns → Bool
|
||||||
|
|
|
@ -6,7 +6,7 @@ open import Function
|
||||||
|
|
||||||
open import Cat.Category
|
open import Cat.Category
|
||||||
|
|
||||||
record Functor {ℓc ℓc' ℓd ℓd'} (C : Category {ℓc} {ℓc'}) (D : Category {ℓd} {ℓd'})
|
record Functor {ℓc ℓc' ℓd ℓd'} (C : Category ℓc ℓc') (D : Category ℓd ℓd')
|
||||||
: Set (ℓc ⊔ ℓc' ⊔ ℓd ⊔ ℓd') where
|
: Set (ℓc ⊔ ℓc' ⊔ ℓd ⊔ ℓd') where
|
||||||
private
|
private
|
||||||
open module C = Category C
|
open module C = Category C
|
||||||
|
@ -21,7 +21,7 @@ record Functor {ℓc ℓc' ℓd ℓd'} (C : Category {ℓc} {ℓc'}) (D : Catego
|
||||||
distrib : { c c' c'' : C.Object} {a : C.Arrow c c'} {a' : C.Arrow c' c''}
|
distrib : { c c' c'' : C.Object} {a : C.Arrow c c'} {a' : C.Arrow c' c''}
|
||||||
→ func→ (a' C.⊕ a) ≡ func→ a' D.⊕ func→ a
|
→ func→ (a' C.⊕ a) ≡ func→ a' D.⊕ func→ a
|
||||||
|
|
||||||
module _ {ℓ ℓ' : Level} {A B C : Category {ℓ} {ℓ'}} (F : Functor B C) (G : Functor A B) where
|
module _ {ℓ ℓ' : Level} {A B C : Category ℓ ℓ'} (F : Functor B C) (G : Functor A B) where
|
||||||
private
|
private
|
||||||
open module F = Functor F
|
open module F = Functor F
|
||||||
open module G = Functor G
|
open module G = Functor G
|
||||||
|
@ -56,7 +56,7 @@ module _ {ℓ ℓ' : Level} {A B C : Category {ℓ} {ℓ'}} (F : Functor B C) (G
|
||||||
}
|
}
|
||||||
|
|
||||||
-- The identity functor
|
-- The identity functor
|
||||||
identity : {ℓ ℓ' : Level} → {C : Category {ℓ} {ℓ'}} → Functor C C
|
identity : ∀ {ℓ ℓ'} → {C : Category ℓ ℓ'} → Functor C C
|
||||||
-- Identity = record { F* = λ x → x ; F→ = λ x → x ; ident = refl ; distrib = refl }
|
-- Identity = record { F* = λ x → x ; F→ = λ x → x ; ident = refl ; distrib = refl }
|
||||||
identity = record
|
identity = record
|
||||||
{ func* = λ x → x
|
{ func* = λ x → x
|
||||||
|
|
Loading…
Reference in a new issue