drop indexes
This commit is contained in:
143
src/Lib/TT.idr
143
src/Lib/TT.idr
@@ -7,29 +7,26 @@ module Lib.TT
|
|||||||
-- For SourcePos
|
-- For SourcePos
|
||||||
import Lib.Parser.Impl
|
import Lib.Parser.Impl
|
||||||
import Data.Fin
|
import Data.Fin
|
||||||
import Data.Vect
|
import Data.List
|
||||||
|
|
||||||
public export
|
public export
|
||||||
Name : Type
|
Name : Type
|
||||||
Name = String
|
Name = String
|
||||||
|
|
||||||
-- Trying to do well-scoped here, so the indices are proven.
|
|
||||||
-- RIP out the indices
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data Icit = Implicit | Explicit
|
data Icit = Implicit | Explicit
|
||||||
|
|
||||||
%name Icit icit
|
%name Icit icit
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data Tm : Nat -> Type where
|
data Tm : Type where
|
||||||
Bnd : Fin n -> Tm n
|
Bnd : Nat -> Tm
|
||||||
Ref : String -> Tm n
|
Ref : String -> Tm
|
||||||
Lam : Name -> Icit -> Tm (S n) -> Tm n
|
Lam : Name -> Icit -> Tm -> Tm
|
||||||
App : Tm n -> Tm n -> Tm n
|
App : Tm -> Tm -> Tm
|
||||||
U : Tm n
|
U : Tm
|
||||||
Pi : Name -> Icit -> Tm n -> Tm (S n) -> Tm n
|
Pi : Name -> Icit -> Tm -> Tm -> Tm
|
||||||
Let : Name -> Icit -> Tm n -> Tm n -> Tm (S n) -> Tm n
|
Let : Name -> Icit -> Tm -> Tm -> Tm -> Tm
|
||||||
|
|
||||||
%name Tm t, u, v
|
%name Tm t, u, v
|
||||||
|
|
||||||
@@ -42,7 +39,7 @@ Eq Icit where
|
|||||||
|
|
||||||
||| Eq on Tm. We've got deBruijn indices, so I'm not comparing names
|
||| Eq on Tm. We've got deBruijn indices, so I'm not comparing names
|
||||||
export
|
export
|
||||||
Eq (Tm n) where
|
Eq (Tm) 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
|
||||||
@@ -55,74 +52,53 @@ Eq (Tm n) where
|
|||||||
|
|
||||||
-- public export
|
-- public export
|
||||||
-- data Closure : Nat -> Type
|
-- data Closure : Nat -> Type
|
||||||
data Val : Nat -> Type
|
data Val : Type
|
||||||
|
|
||||||
public export
|
|
||||||
0 Closure : Nat -> Type
|
|
||||||
|
|
||||||
-- IS/TypeTheory.idr is calling this a Kripke function space
|
-- IS/TypeTheory.idr is calling this a Kripke function space
|
||||||
-- Yaffle, IS/TypeTheory use a function here.
|
-- Yaffle, IS/TypeTheory use a function here.
|
||||||
-- Kovacs, Idris use Env and Tm
|
-- Kovacs, Idris use Env and Tm
|
||||||
|
|
||||||
-- in cctt kovacs refers to this as defunctionalization vs HOAS
|
-- in cctt kovacs refers to this choice as defunctionalization vs HOAS
|
||||||
-- https://github.com/AndrasKovacs/cctt/blob/main/README.md#defunctionalization
|
-- https://github.com/AndrasKovacs/cctt/blob/main/README.md#defunctionalization
|
||||||
|
-- the tradeoff is access to internals of the function
|
||||||
|
|
||||||
-- Yaffle is vars -> vars here
|
-- Yaffle is vars -> vars here
|
||||||
|
|
||||||
|
|
||||||
Closure n = Val (S n) -> Val n
|
public export
|
||||||
|
0 Closure : Type
|
||||||
|
Closure = Val -> Val
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data Val : Nat -> Type where
|
data Val : Type where
|
||||||
-- This will be local / flex with spine.
|
-- This will be local / flex with spine.
|
||||||
VVar : Fin n -> Val n
|
VVar : Nat -> Val
|
||||||
VRef : String -> Val n
|
VRef : String -> Val
|
||||||
VApp : Val n -> Lazy (Val n) -> Val n
|
VApp : Val -> Lazy (Val) -> Val
|
||||||
VLam : Name -> Icit -> Closure n -> Val n
|
VLam : Name -> Icit -> Closure -> Val
|
||||||
VPi : Name -> Icit -> Lazy (Val n) -> Closure n -> Val n
|
VPi : Name -> Icit -> Lazy Val -> Closure -> Val
|
||||||
VU : Val n
|
VU : Val
|
||||||
|
|
||||||
||| LocalEnv free vars
|
||| LocalEnv free vars
|
||||||
public export
|
public export
|
||||||
LocalEnv : Nat -> Nat -> Type
|
LocalEnv : Type
|
||||||
LocalEnv k n = Vect k (Val n)
|
LocalEnv = List Val
|
||||||
|
|
||||||
public export
|
public export
|
||||||
Env : Nat -> Type
|
Env : Type
|
||||||
Env n = Vect n (Val n)
|
Env = List Val
|
||||||
|
|
||||||
export
|
export
|
||||||
eval : LocalEnv k n -> Env n -> Tm (n + k) -> Val n
|
eval : LocalEnv -> Env -> Tm -> Val
|
||||||
|
|
||||||
thinVal : Val k -> Val (S k)
|
|
||||||
thinVal (VVar x) = VVar (shift 1 x)
|
|
||||||
thinVal (VRef str) = VRef str
|
|
||||||
thinVal (VApp x y) = VApp (thinVal x) (thinVal y)
|
|
||||||
thinVal (VLam str icit f) = VLam str icit (believe_me f) -- FIXME
|
|
||||||
thinVal (VPi str icit x f) = VPi str icit (thinVal x) (believe_me f)
|
|
||||||
thinVal VU = VU
|
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
vapp : Val n -> Val n -> Val n
|
vapp : Val -> Val -> Val
|
||||||
vapp (VLam _ icit t) u = t (thinVal u)
|
vapp (VLam _ icit t) u = t u
|
||||||
vapp t u = VApp t u
|
vapp t u = VApp t u
|
||||||
|
|
||||||
-- thinVal : {e : Nat} -> Val k -> Val (e + k)
|
bind : Val -> Env -> Env
|
||||||
-- thinVal (VVar x) = VVar (shift _ x)
|
bind v env = v :: env
|
||||||
-- 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))
|
|
||||||
-- 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))
|
|
||||||
-- thinVal VU = VU
|
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
bind : Val n -> Env n -> Env (S n)
|
|
||||||
bind v env = thinVal v :: map thinVal env
|
|
||||||
|
|
||||||
-- so here we have LocalEnv free vars in Yaffle
|
-- so here we have LocalEnv free vars in Yaffle
|
||||||
|
|
||||||
@@ -130,51 +106,45 @@ eval locs env (Ref x) = VRef x
|
|||||||
eval locs env (App t u) = vapp (eval locs env t) (eval locs env u)
|
eval locs env (App t u) = vapp (eval locs env t) (eval locs env u)
|
||||||
eval locs env U = VU
|
eval locs env U = VU
|
||||||
-- yaffle breaks out binder
|
-- yaffle breaks out binder
|
||||||
eval locs env (Lam x icit t) = VLam x icit (\u => (u :: locs) env (rewrite sym $ plusSuccRightSucc n k in t))
|
eval locs env (Lam x icit t) = VLam x icit (\u => eval (u :: locs) env t)
|
||||||
eval locs env (Pi x icit a b) = VPi x icit (eval locs env a) (\u => eval (u :: locs) env (rewrite sym $ plusSuccRightSucc n k in b))
|
eval locs env (Pi x icit a b) = VPi x icit (eval locs env a) (\u => eval (u :: locs) env b)
|
||||||
-- This one we need to make
|
-- This one we need to make
|
||||||
eval locs env (Let x icit ty t u) = eval (eval locs env t :: locs) env (rewrite sym $ plusSuccRightSucc n k in u)
|
eval locs env (Let x icit ty t u) = eval (eval locs env t :: locs) env u
|
||||||
eval locs env (Bnd i) = index i ?hole -- env
|
eval locs env (Bnd i) = let Just rval = getAt i env | _ => ?out_of_index in rval
|
||||||
|
|
||||||
vfresh : (l : Nat) -> Val (S l)
|
vfresh : Nat -> Val
|
||||||
vfresh l = VVar last
|
vfresh l = VVar l
|
||||||
|
|
||||||
export
|
export
|
||||||
quote : (k : Nat) -> Val k -> Tm k
|
quote : (k : Nat) -> Val -> Tm
|
||||||
quote l (VVar k) = Bnd (complement k) -- level to index
|
quote l (VVar k) = Bnd (S l `minus` 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
|
||||||
quote l (VLam x icit t) = Lam x icit (quote (S l) (believe_me $ t (vfresh l))) -- that one is too big
|
quote l (VLam x icit t) = Lam x icit (quote (S l) (t (vfresh l))) -- that one is too big
|
||||||
quote l (VPi x icit a b) = Pi x icit (quote l a) (quote (S l) (believe_me $ b $ vfresh l))
|
quote l (VPi x icit a b) = Pi x icit (quote l a) (quote (S l) (b $ vfresh l))
|
||||||
quote l VU = U
|
quote l VU = U
|
||||||
quote _ (VRef n) = Ref n
|
quote _ (VRef n) = Ref n
|
||||||
|
|
||||||
-- vars -> vars -> vars in yaffle
|
-- vars -> vars -> vars in yaffle
|
||||||
export
|
export
|
||||||
nf : {n : Nat} -> Env n -> Tm n -> Tm n
|
nf : {n : Nat} -> Env -> Tm -> Tm
|
||||||
nf env t = quote n (eval [] env (rewrite plusZeroRightNeutral n in t))
|
nf env t = quote n (eval [] env t)
|
||||||
|
|
||||||
public export
|
public export
|
||||||
conv : (lvl : Nat) -> Val n -> Val n -> Bool
|
conv : (lvl : Nat) -> Val -> Val -> Bool
|
||||||
|
|
||||||
data BD = Bound | Defined
|
data BD = Bound | Defined
|
||||||
|
|
||||||
public export
|
public export
|
||||||
Types : Nat -> Type
|
Types : Type
|
||||||
Types n = Vect n (Name, Lazy (Val n))
|
Types = List (Name, Lazy Val)
|
||||||
|
|
||||||
|
|
||||||
|
|
||||||
-- REVIEW indices
|
-- REVIEW indices
|
||||||
public export
|
public export
|
||||||
record Context (n : Nat) where
|
record Context where
|
||||||
constructor MkCtx
|
constructor MkCtx
|
||||||
|
env : Env
|
||||||
-- These are values, they'll be the length of the environment
|
types : List (String, Val)
|
||||||
env : Env n -- Vect n (Val f)
|
|
||||||
|
|
||||||
-- fine for now, consider a map later
|
|
||||||
types : Vect n (String, Val n)
|
|
||||||
pos : SourcePos
|
pos : SourcePos
|
||||||
|
|
||||||
-- data Env : (tm : SnocList Name -> Type) -> SnocList Name -> Type where
|
-- data Env : (tm : SnocList Name -> Type) -> SnocList Name -> Type where
|
||||||
@@ -183,14 +153,9 @@ record Context (n : Nat) where
|
|||||||
-- Still need to sort out the indices - one or two on env?
|
-- Still need to sort out the indices - one or two on env?
|
||||||
|
|
||||||
||| add a binding to environment
|
||| add a binding to environment
|
||||||
extend : { n : Nat} -> Context n -> String -> Val n -> Context (S n)
|
extend : { n : Nat} -> Context -> String -> Val -> Context
|
||||||
extend (MkCtx env types pos) name ty with (length env)
|
extend (MkCtx env types pos) name ty =
|
||||||
_ | l =
|
MkCtx (VVar (length env) :: env) ((name, ty) :: types) pos
|
||||||
|
|
||||||
let types' = Data.Vect.(::) (name, thinVal ty) (map (map thinVal) types) in
|
|
||||||
let l' : Fin (S n) := last in
|
|
||||||
MkCtx {n=S n} (VVar l' :: map thinVal env) types' pos
|
|
||||||
-- ?hole_0 -- { env := (Val (length env.(Context env types pos)) :: (Context env types pos).env), types := (n, ty) :: (Context env types pos).types } (Context env types pos)
|
|
||||||
|
|
||||||
|
|
||||||
-- weirich has 'hints' to store the claims before the def is seen/checked
|
-- weirich has 'hints' to store the claims before the def is seen/checked
|
||||||
@@ -240,7 +205,7 @@ extend (MkCtx env types pos) name ty with (length env)
|
|||||||
-- MkCtx (VVar l :: env) ((x,a) :: map (map thinVal) types) (Bound :: bds) (l+1) pos
|
-- MkCtx (VVar l :: env) ((x,a) :: map (map thinVal) types) (Bound :: bds) (l+1) pos
|
||||||
|
|
||||||
-- public export
|
-- public export
|
||||||
-- define : Name -> Val n -> Lazy (Val n) -> Ctx n -> Ctx (S n)
|
-- define : Name -> Val -> Lazy (Val) -> Ctx n -> Ctx (S n)
|
||||||
-- define x v ty (MkCtx env types bds l pos) =
|
-- define x v ty (MkCtx env types bds l pos) =
|
||||||
-- MkCtx (v :: env) ((x,ty) :: map (map thinVal) types) (Defined :: bds) (l + 1) pos
|
-- MkCtx (v :: env) ((x,ty) :: map (map thinVal) types) (Defined :: bds) (l + 1) pos
|
||||||
|
|
||||||
|
|||||||
Reference in New Issue
Block a user