remove an index from Tm
This commit is contained in:
@@ -21,9 +21,9 @@ extend ctx ty = { env := ty :: ctx.env } ctx
|
|||||||
-- cribbed this, it avoids MonadError String m => everywhere
|
-- cribbed this, it avoids MonadError String m => everywhere
|
||||||
parameters {0 m : Type -> Type} {auto _ : MonadError String m}
|
parameters {0 m : Type -> Type} {auto _ : MonadError String m}
|
||||||
|
|
||||||
infer : {f : Nat} -> Context n f -> Raw -> m (Tm f n, Val f)
|
infer : {f : Nat} -> Context n f -> Raw -> m (Tm n, Val f)
|
||||||
-- I think I'm hand-waving n here, probably need it in Context
|
-- I think I'm hand-waving n here, probably need it in Context
|
||||||
check : {f : Nat} -> Context n f -> Raw -> Val f -> m (Tm f n)
|
check : {f : Nat} -> Context n f -> Raw -> Val f -> m (Tm n)
|
||||||
|
|
||||||
check ctx (RLam _ _ _) ty = ?ch_rhs
|
check ctx (RLam _ _ _) ty = ?ch_rhs
|
||||||
check ctx tm ty = do
|
check ctx tm ty = do
|
||||||
@@ -35,7 +35,7 @@ parameters {0 m : Type -> Type} {auto _ : MonadError String m}
|
|||||||
|
|
||||||
infer ctx (RVar nm) = go 0 ctx.types
|
infer ctx (RVar nm) = go 0 ctx.types
|
||||||
where
|
where
|
||||||
go : Nat -> List (String, Val f) -> m (Tm f n, Val f)
|
go : Nat -> List (String, Val f) -> m (Tm n, Val f)
|
||||||
go i [] = throwError "\{show nm} not in scope"
|
go i [] = throwError "\{show nm} not in scope"
|
||||||
-- REVIEW Local or Bnd (ezoo does not have both)
|
-- REVIEW Local or Bnd (ezoo does not have both)
|
||||||
go i ((x, ty) :: xs) = if x == nm then pure $ (Bnd ?i_not_fin, ty)
|
go i ((x, ty) :: xs) = if x == nm then pure $ (Bnd ?i_not_fin, ty)
|
||||||
|
|||||||
@@ -16,22 +16,21 @@ data Icit = Implicit | Explicit
|
|||||||
%name Icit icit
|
%name Icit icit
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data Tm : Nat -> Nat -> Type where
|
data Tm : Nat -> Type where
|
||||||
Local : Fin k -> Tm k n
|
Bnd : Fin n -> Tm n
|
||||||
Bnd : Fin n -> Tm k n
|
Ref : String -> Tm n
|
||||||
Ref : String -> Tm k n
|
Lam : Name -> Icit -> Tm (S n) -> Tm n
|
||||||
Lam : Name -> Icit -> Tm k (S n) -> Tm k n
|
App : Tm n -> Tm n -> Tm n
|
||||||
App : Tm k n -> Tm k n -> Tm k n
|
U : Tm n
|
||||||
U : Tm k n
|
Pi : Name -> Icit -> Tm n -> Tm (S n) -> Tm n
|
||||||
Pi : Name -> Icit -> Tm k n -> Tm k (S n) -> Tm k n
|
Let : Name -> Icit -> Tm n -> Tm n -> Tm (S n) -> Tm n
|
||||||
Let : Name -> Icit -> Tm k n -> Tm k n -> Tm k (S n) -> Tm k n
|
|
||||||
|
|
||||||
%name Tm t, u, v
|
%name Tm t, u, v
|
||||||
|
|
||||||
-- TODO derive
|
-- TODO derive
|
||||||
export
|
export
|
||||||
Eq (Tm k n) where
|
Eq (Tm n) where
|
||||||
(Local x) == (Local y) = x == y
|
-- (Local x) == (Local y) = x == y
|
||||||
(Bnd x) == (Bnd y) = x == y
|
(Bnd x) == (Bnd y) = x == y
|
||||||
(Ref x) == (Ref y) = x == y
|
(Ref x) == (Ref y) = x == y
|
||||||
(Lam str icit t) == y = ?rhs_3
|
(Lam str icit t) == y = ?rhs_3
|
||||||
@@ -71,7 +70,7 @@ Env : Nat -> Nat -> Type
|
|||||||
Env k n = Vect n (Val k)
|
Env k n = Vect n (Val k)
|
||||||
|
|
||||||
export
|
export
|
||||||
eval : Env k n -> Tm k n -> Val k
|
eval : Env k n -> Tm n -> Val k
|
||||||
|
|
||||||
export
|
export
|
||||||
vapp : Val k -> Val k -> Val k
|
vapp : Val k -> Val k -> Val k
|
||||||
@@ -93,24 +92,12 @@ thinVal VU = VU
|
|||||||
bind : (e : Nat) -> Val (plus e k) -> Env k n -> Env (e + k) (S n)
|
bind : (e : Nat) -> Val (plus e k) -> Env k n -> Env (e + k) (S n)
|
||||||
bind e v env = v :: map thinVal env
|
bind e v env = v :: map thinVal env
|
||||||
|
|
||||||
-- 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 (Ref x) = VRef x
|
||||||
eval env (Bnd n) = index n env
|
eval env (Bnd n) = index n env
|
||||||
eval env (Lam x icit t) = VLam x icit (\e, u => eval (bind e u env) (thin t))-- (MkClosure env t)
|
eval env (Lam x icit t) = VLam x icit (\e, u => eval (bind e u env) t)-- (MkClosure env t)
|
||||||
eval env (App t u) = vapp (eval env t) (eval env u)
|
eval env (App t u) = vapp (eval env t) (eval env u)
|
||||||
eval env U = VU
|
eval env U = VU
|
||||||
eval env (Pi x icit a b) = VPi x icit (eval env a) (\e, u => eval (bind e u env) (thin b))
|
eval env (Pi x icit a b) = VPi x icit (eval env a) (\e, u => eval (bind e u env) b)
|
||||||
-- This one we need to make
|
-- This one we need to make
|
||||||
eval env (Let x icit ty t u) = eval (eval env t :: env) u
|
eval env (Let x icit ty t u) = eval (eval env t :: env) u
|
||||||
|
|
||||||
@@ -118,7 +105,7 @@ vfresh : (l : Nat) -> Val (S l)
|
|||||||
vfresh l = VVar last
|
vfresh l = VVar last
|
||||||
|
|
||||||
export
|
export
|
||||||
quote : (k : Nat) -> Val k -> Tm 0 k
|
quote : (k : Nat) -> Val k -> Tm k
|
||||||
quote l (VVar k) = Bnd (complement k) -- level to index
|
quote l (VVar k) = Bnd (complement k) -- level to index
|
||||||
quote l (VApp t u) = App (quote l t) (quote l u)
|
quote l (VApp t u) = App (quote l t) (quote l u)
|
||||||
-- so this one is calling the kripke on [x] and a fresh var
|
-- so this one is calling the kripke on [x] and a fresh var
|
||||||
@@ -128,7 +115,7 @@ quote l VU = U
|
|||||||
quote _ (VRef n) = Ref n
|
quote _ (VRef n) = Ref n
|
||||||
|
|
||||||
export
|
export
|
||||||
nf : {n : Nat} -> Env 0 n -> Tm 0 n -> Tm 0 0
|
nf : {n : Nat} -> Env 0 n -> Tm n -> Tm 0
|
||||||
nf env t = quote 0 (eval env t)
|
nf env t = quote 0 (eval env t)
|
||||||
|
|
||||||
public export
|
public export
|
||||||
|
|||||||
Reference in New Issue
Block a user