cat/src/Cat/CwF.agda

56 lines
1.7 KiB
Agda
Raw Normal View History

module Cat.CwF where
open import Agda.Primitive
open import Data.Product
open import Cat.Category
2018-02-05 13:59:53 +00:00
open import Cat.Category.Functor
open import Cat.Categories.Fam
open Category
open Functor
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
-- It's objects are called contexts
Contexts = Object
-- It's arrows are called substitutions
Substitutions = Arrow
field
-- A functor T
T : Functor (Opposite ) (Fam a b)
-- Empty context
[] : Terminal
Type : (Γ : Object ) Set a
Type Γ = proj₁ (T .func* Γ)
module _ {Γ : Object } {A : Type Γ} where
module _ {A B : Object } {γ : [ A , B ]} where
k : Σ (proj₁ (func* T B) proj₁ (func* T A))
(λ f
{x : proj₁ (func* T B)}
proj₂ (func* T B) x proj₂ (func* T A) (f x))
k = T .func→ γ
k₁ : proj₁ (func* T B) proj₁ (func* T A)
k₁ = proj₁ k
k₂ : ({x : proj₁ (func* T B)}
proj₂ (func* T B) x proj₂ (func* T A) (k₁ x))
k₂ = proj₂ k
record ContextComprehension : Set (a b) where
field
Γ&A : Object
proj1 : [ Γ&A , Γ ]
-- proj2 : ????
-- if γ : [ A , B ]
-- then T .func→ γ (written T[γ]) interpret substitutions in types and terms respectively.
-- field
-- ump : {Δ : .Object} → (γ : [ Δ , Γ ])
-- → (a : {!!}) → {!!}