cat/src/Cat/Categories/CwF.agda

58 lines
1.7 KiB
Agda
Raw Normal View History

2018-02-25 14:24:44 +00:00
module Cat.Categories.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
2018-02-25 14:23:33 +00:00
T : Functor (opposite ) (Fam a b)
-- Empty context
[] : Terminal
private
module T = Functor T
Type : (Γ : Object ) Set a
Type Γ = proj₁ (proj₁ (T.omap Γ))
module _ {Γ : Object } {A : Type Γ} where
2018-02-23 13:13:55 +00:00
-- module _ {A B : Object } {γ : [ A , B ]} where
-- k : Σ (proj₁ (omap T B) → proj₁ (omap T A))
2018-02-23 13:13:55 +00:00
-- (λ f →
-- {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
-- 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
record ContextComprehension : Set (a b) where
field
Γ&A : Object
proj1 : [ Γ&A , Γ ]
-- proj2 : ????
-- if γ : [ A , B ]
-- then T .fmap γ (written T[γ]) interpret substitutions in types and terms respectively.
-- field
-- ump : {Δ : .Object} → (γ : [ Δ , Γ ])
-- → (a : {!!}) → {!!}