more work on well scoped

This commit is contained in:
2023-07-18 22:28:08 -07:00
parent 59f726ab96
commit f221f09423
3 changed files with 153 additions and 57 deletions

View File

@@ -10,7 +10,7 @@ Name = String
-- Trying to do well-scoped here, so the indices are proven.
export
public export
data Icit = Implicit | Explicit
%name Icit icit
@@ -28,9 +28,24 @@ data Tm : Nat -> Nat -> Type where
%name Tm t, u, v
-- TODO derive
export
Eq (Tm k n) where
(Local x) == (Local y) = x == y
(Bnd x) == (Bnd y) = x == y
(Ref x) == (Ref y) = x == y
(Lam str icit t) == y = ?rhs_3
(App t u) == y = ?rhs_4
U == y = ?rhs_5
(Pi str icit t u) == y = ?rhs_6
(Let str icit t u v) == y = ?rhs_7
_ == _ = False
-- public export
-- data Closure : Nat -> Type
data Val : Nat -> Type
public export
0 Closure : Nat -> Type
-- IS/TypeTheory.idr is calling this a Kripke function space
@@ -58,49 +73,51 @@ Env k n = Vect n (Val k)
export
eval : Env k n -> Tm k n -> Val k
export
vapp : Val k -> Val k -> Val k
vapp (VLam _ icit t) u = t 0 u
vapp t u = VApp t u
-- weakenEnv : (l : Nat) -> Env k n -> Env (l + k) n
-- thinEnv : (l : Nat) -> Env k n -> Env (l + k) n
weakenVal : {e : Nat} -> Val k -> Val (e + k)
weakenVal (VVar x) = VVar (shift _ x)
weakenVal (VRef str) = VRef str
weakenVal (VApp x y) = VApp (weakenVal x) (weakenVal y)
weakenVal (VLam str icit f) = VLam str icit
thinVal : {e : Nat} -> Val k -> Val (e + k)
thinVal (VVar x) = VVar (shift _ x)
thinVal (VRef str) = VRef str
thinVal (VApp x y) = VApp (thinVal x) (thinVal y)
thinVal (VLam str icit f) = VLam str icit
(\g, v => rewrite plusAssociative g e k in f (g + e) (rewrite sym $ plusAssociative g e k in v))
weakenVal (VPi str icit x f) = VPi str icit (weakenVal {e} x)
thinVal (VPi str icit x f) = VPi str icit (thinVal {e} x)
(\g, v => rewrite plusAssociative g e k in f (g + e) (rewrite sym $ plusAssociative g e k in v))
weakenVal VU = VU
thinVal VU = VU
bind : (e : Nat) -> Val (plus e k) -> Env k n -> Env (e + k) (S n)
bind e v env = v :: map weakenVal env
bind e v env = v :: map thinVal env
-- is this weaken or thin?
weaken : {e : Nat} -> Tm k (S n) -> Tm (plus e k) (S n)
weaken (Local x) = Local (shift _ x)
weaken (Bnd x) = Bnd x
weaken (Ref str) = Ref str
weaken (Lam str x t) = Lam str x (weaken t)
weaken (App t u) = App (weaken t) (weaken u)
weaken U = U
weaken (Pi str x t u) = Pi str x (weaken t) (weaken u)
weaken (Let str x t u v) = Let str x (weaken t) (weaken u) (weaken v)
-- is this thin or thin?
thin : {e : Nat} -> Tm k (S n) -> Tm (plus e k) (S n)
thin (Local x) = Local (shift _ x)
thin (Bnd x) = Bnd x
thin (Ref str) = Ref str
thin (Lam str x t) = Lam str x (thin t)
thin (App t u) = App (thin t) (thin u)
thin U = U
thin (Pi str x t u) = Pi str x (thin t) (thin u)
thin (Let str x t u v) = Let str x (thin t) (thin u) (thin v)
eval env (Local x) = VVar x -- this is a hole in intrinsic, vfree x in the other
eval env (Ref x) = VRef x
eval env (Bnd n) = index n env
eval env (Lam x icit t) = VLam x icit (\e, u => eval (bind e u env) (weaken t))-- (MkClosure env t)
eval env (Lam x icit t) = VLam x icit (\e, u => eval (bind e u env) (thin t))-- (MkClosure env t)
eval env (App t u) = vapp (eval env t) (eval env u)
eval env U = VU
eval env (Pi x icit a b) = VPi x icit (eval env a) (\e, u => eval (bind e u env) (weaken b))
eval env (Pi x icit a b) = VPi x icit (eval env a) (\e, u => eval (bind e u env) (thin b))
-- This one we need to make
eval env (Let x icit ty t u) = eval (eval env t :: env) u
vfresh : (l : Nat) -> Val (S l)
vfresh l = VVar last
export
quote : (k : Nat) -> Val k -> Tm 0 k
quote l (VVar k) = Bnd (complement k) -- level to index
quote l (VApp t u) = App (quote l t) (quote l u)
@@ -110,6 +127,7 @@ quote l (VPi x icit a b) = Pi x icit (quote l a) (quote (S l) (b 1 $ vfresh l))
quote l VU = U
quote _ (VRef n) = Ref n
export
nf : {n : Nat} -> Env 0 n -> Tm 0 n -> Tm 0 0
nf env t = quote 0 (eval env t)
@@ -122,29 +140,29 @@ public export
Types : Nat -> Type
Types n = Vect n (Name, Lazy (Val n))
public export
record Ctx (n : Nat) where
constructor MkCtx
env : Env k n -- for eval
types : Types n -- name lookup, pp
bds : Vect n BD -- meta creation
lvl : Nat -- This is n, do we need it?
-- Kovacs and Weirich use a position node
pos : SourcePos
-- public export
-- record Ctx (n : Nat) where
-- constructor MkCtx
-- env : Env k n -- for eval
-- types : Types n -- name lookup, pp
-- bds : Vect n BD -- meta creation
-- lvl : Nat -- This is n, do we need it?
-- -- Kovacs and Weirich use a position node
-- pos : SourcePos
%name Ctx ctx
public export
emptyCtx : Ctx Z
emptyCtx = MkCtx {k=0} [] [] [] 0 (0,0)
-- %name Ctx ctx
-- public export
-- bind : Name -> Lazy (Val (k + n)) -> Ctx n -> Ctx (S n)
-- bind x a (MkCtx env types bds l pos) =
-- MkCtx (VVar l :: env) ((x,a) :: types) (Bound :: bds) (l+1) pos
-- emptyCtx : Ctx Z
-- emptyCtx = MkCtx {k=0} [] [] [] 0 (0,0)
-- public export
-- bindCtx : Name -> Lazy (Val (zz + n)) -> Ctx n -> Ctx (S n)
-- bindCtx x a (MkCtx env types bds l pos) =
-- MkCtx (VVar l :: env) ((x,a) :: map (map thinVal) types) (Bound :: bds) (l+1) pos
-- public export
-- define : Name -> Val n -> Lazy (Val n) -> Ctx n -> Ctx (S n)
-- define x v ty (MkCtx env types bds l pos) =
-- MkCtx (v :: env) ((x,ty) :: types) (Defined :: bds) (l + 1) pos
-- MkCtx (v :: env) ((x,ty) :: map (map thinVal) types) (Defined :: bds) (l + 1) pos