2018-01-31 14:15:00 +00:00
|
|
|
|
{-# OPTIONS --allow-unsolved-metas #-}
|
|
|
|
|
module Cat.Categories.Fam where
|
|
|
|
|
|
|
|
|
|
open import Agda.Primitive
|
|
|
|
|
open import Data.Product
|
|
|
|
|
open import Cubical
|
|
|
|
|
import Function
|
|
|
|
|
|
|
|
|
|
open import Cat.Category
|
|
|
|
|
open import Cat.Equality
|
|
|
|
|
|
|
|
|
|
open Equality.Data.Product
|
|
|
|
|
|
2018-02-02 13:47:33 +00:00
|
|
|
|
module _ (ℓa ℓb : Level) where
|
2018-01-31 14:15:00 +00:00
|
|
|
|
private
|
2018-02-05 10:43:38 +00:00
|
|
|
|
Obj' = Σ[ A ∈ Set ℓa ] (A → Set ℓb)
|
|
|
|
|
Arr : Obj' → Obj' → Set (ℓa ⊔ ℓb)
|
2018-01-31 14:15:00 +00:00
|
|
|
|
Arr (A , B) (A' , B') = Σ[ f ∈ (A → A') ] ({x : A} → B x → B' (f x))
|
2018-02-05 10:43:38 +00:00
|
|
|
|
one : {o : Obj'} → Arr o o
|
2018-01-31 14:15:00 +00:00
|
|
|
|
proj₁ one = λ x → x
|
|
|
|
|
proj₂ one = λ b → b
|
2018-02-05 10:43:38 +00:00
|
|
|
|
_∘_ : {a b c : Obj'} → Arr b c → Arr a b → Arr a c
|
2018-01-31 14:15:00 +00:00
|
|
|
|
(g , g') ∘ (f , f') = g Function.∘ f , g' Function.∘ f'
|
2018-02-05 10:43:38 +00:00
|
|
|
|
_⟨_∘_⟩ : {a b : Obj'} → (c : Obj') → Arr b c → Arr a b → Arr a c
|
2018-01-31 14:15:00 +00:00
|
|
|
|
c ⟨ g ∘ f ⟩ = _∘_ {c = c} g f
|
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
module _ {A B C D : Obj'} {f : Arr A B} {g : Arr B C} {h : Arr C D} where
|
2018-02-23 11:43:49 +00:00
|
|
|
|
isAssociative : (D ⟨ h ∘ C ⟨ g ∘ f ⟩ ⟩) ≡ D ⟨ D ⟨ h ∘ g ⟩ ∘ f ⟩
|
|
|
|
|
isAssociative = Σ≡ refl refl
|
2018-01-31 14:15:00 +00:00
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
module _ {A B : Obj'} {f : Arr A B} where
|
2018-02-23 11:49:41 +00:00
|
|
|
|
isIdentity : B ⟨ f ∘ one ⟩ ≡ f × B ⟨ one {B} ∘ f ⟩ ≡ f
|
|
|
|
|
isIdentity = (Σ≡ refl refl) , Σ≡ refl refl
|
2018-01-31 14:15:00 +00:00
|
|
|
|
|
2018-02-05 10:43:38 +00:00
|
|
|
|
|
|
|
|
|
RawFam : RawCategory (lsuc (ℓa ⊔ ℓb)) (ℓa ⊔ ℓb)
|
|
|
|
|
RawFam = record
|
|
|
|
|
{ Object = Obj'
|
|
|
|
|
; Arrow = Arr
|
|
|
|
|
; 𝟙 = one
|
|
|
|
|
; _∘_ = λ {a b c} → _∘_ {a} {b} {c}
|
|
|
|
|
}
|
|
|
|
|
|
2018-01-31 14:15:00 +00:00
|
|
|
|
instance
|
2018-02-05 10:43:38 +00:00
|
|
|
|
isCategory : IsCategory RawFam
|
2018-01-31 14:15:00 +00:00
|
|
|
|
isCategory = record
|
2018-02-23 11:43:49 +00:00
|
|
|
|
{ isAssociative = λ {A} {B} {C} {D} {f} {g} {h} → isAssociative {D = D} {f} {g} {h}
|
2018-02-23 11:49:41 +00:00
|
|
|
|
; isIdentity = λ {A} {B} {f} → isIdentity {A} {B} {f = f}
|
2018-02-06 09:34:43 +00:00
|
|
|
|
; arrowIsSet = {!!}
|
2018-02-05 13:47:15 +00:00
|
|
|
|
; univalent = {!!}
|
2018-01-31 14:15:00 +00:00
|
|
|
|
}
|
|
|
|
|
|
|
|
|
|
Fam : Category (lsuc (ℓa ⊔ ℓb)) (ℓa ⊔ ℓb)
|
2018-02-05 13:47:15 +00:00
|
|
|
|
Category.raw Fam = RawFam
|