2018-02-25 14:24:44 +00:00
|
|
|
|
module Cat.Categories.CwF where
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
2018-03-21 13:56:43 +00:00
|
|
|
|
open import Cat.Prelude
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
|
|
|
|
open import Cat.Category
|
2018-02-05 13:59:53 +00:00
|
|
|
|
open import Cat.Category.Functor
|
2018-02-02 13:47:33 +00:00
|
|
|
|
open import Cat.Categories.Fam
|
|
|
|
|
|
|
|
|
|
module _ {ℓa ℓb : Level} where
|
|
|
|
|
record CwF : Set (lsuc (ℓa ⊔ ℓb)) where
|
|
|
|
|
-- "A category with families consists of"
|
|
|
|
|
field
|
|
|
|
|
-- "A base category"
|
|
|
|
|
ℂ : Category ℓa ℓb
|
2018-03-21 13:56:43 +00:00
|
|
|
|
module ℂ = Category ℂ
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- It's objects are called contexts
|
2018-03-21 13:56:43 +00:00
|
|
|
|
Contexts = ℂ.Object
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- It's arrows are called substitutions
|
2018-03-21 13:56:43 +00:00
|
|
|
|
Substitutions = ℂ.Arrow
|
2018-02-02 13:47:33 +00:00
|
|
|
|
field
|
|
|
|
|
-- A functor T
|
2018-02-25 14:23:33 +00:00
|
|
|
|
T : Functor (opposite ℂ) (Fam ℓa ℓb)
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- Empty context
|
2018-03-21 13:56:43 +00:00
|
|
|
|
[] : ℂ.Terminal
|
2018-02-06 13:24:34 +00:00
|
|
|
|
private
|
|
|
|
|
module T = Functor T
|
2018-03-21 13:56:43 +00:00
|
|
|
|
Type : (Γ : ℂ.Object) → Set ℓa
|
2018-03-08 10:03:56 +00:00
|
|
|
|
Type Γ = proj₁ (proj₁ (T.omap Γ))
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
2018-03-21 13:56:43 +00:00
|
|
|
|
module _ {Γ : ℂ.Object} {A : Type Γ} where
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
2018-02-23 13:13:55 +00:00
|
|
|
|
-- module _ {A B : Object ℂ} {γ : ℂ [ A , B ]} where
|
2018-03-08 10:03:56 +00:00
|
|
|
|
-- k : Σ (proj₁ (omap T B) → proj₁ (omap T A))
|
2018-02-23 13:13:55 +00:00
|
|
|
|
-- (λ f →
|
2018-03-08 10:03:56 +00:00
|
|
|
|
-- {x : proj₁ (omap T B)} →
|
|
|
|
|
-- proj₂ (omap T B) x → proj₂ (omap T A) (f x))
|
|
|
|
|
-- k = T.fmap γ
|
|
|
|
|
-- k₁ : proj₁ (omap T B) → proj₁ (omap T A)
|
2018-02-23 13:13:55 +00:00
|
|
|
|
-- k₁ = proj₁ k
|
2018-03-08 10:03:56 +00:00
|
|
|
|
-- k₂ : ({x : proj₁ (omap T B)} →
|
|
|
|
|
-- proj₂ (omap T B) x → proj₂ (omap T A) (k₁ x))
|
2018-02-23 13:13:55 +00:00
|
|
|
|
-- k₂ = proj₂ k
|
2018-02-02 13:47:33 +00:00
|
|
|
|
|
|
|
|
|
record ContextComprehension : Set (ℓa ⊔ ℓb) where
|
|
|
|
|
field
|
2018-03-21 13:56:43 +00:00
|
|
|
|
Γ&A : ℂ.Object
|
2018-02-05 11:21:39 +00:00
|
|
|
|
proj1 : ℂ [ Γ&A , Γ ]
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- proj2 : ????
|
|
|
|
|
|
|
|
|
|
-- if γ : ℂ [ A , B ]
|
2018-03-08 10:03:56 +00:00
|
|
|
|
-- then T .fmap γ (written T[γ]) interpret substitutions in types and terms respectively.
|
2018-02-02 13:47:33 +00:00
|
|
|
|
-- field
|
|
|
|
|
-- ump : {Δ : ℂ .Object} → (γ : ℂ [ Δ , Γ ])
|
|
|
|
|
-- → (a : {!!}) → {!!}
|