implicits working, but _slow_

This commit is contained in:
2024-07-09 22:43:24 -07:00
parent edfa5ef443
commit a4d851b563
11 changed files with 624 additions and 454 deletions

View File

@@ -7,6 +7,7 @@ import Lib.Prettier
import Data.List
import Data.Vect
import Data.String
import Lib.Types
import Lib.TT
import Lib.TopContext
import Syntax
@@ -16,83 +17,80 @@ import Syntax
data PRen = PR Nat Nat (List Nat)
-- IORef for metas needs IO
parameters {0 m : Type -> Type} {auto _ : HasIO m} {auto _ : MonadError Error m} (top : TopContext)
-- unify : Nat -> Val -> Val -> m ()
-- unify l (VLam _ _ t) (VLam _ _ u) = unify (l + 1) (t $$ VVar l) (u $$ VVar l)
-- unify l t (VLam _ _ u) = unify (l + 1) (vapp t (VVar l)) (u $$ VVar l)
-- unify l (VVar k) u = ?unify_rhs_0
-- unify l (VRef str mt) u = ?unify_rhs_1
-- unify l (VMeta k) u = ?unify_rhs_2
-- unify l (VApp x y) u = ?unify_rhs_3
-- unify l (VPi str icit x y) u = ?unify_rhs_5
-- unify l VU u = ?unify_rhs_6
forceMeta : Val -> M Val
-- TODO - need to look up metas
forceMeta (VMeta ix sp) = case !(lookupMeta ix) of
(Unsolved k xs) => pure (VMeta ix sp)
(Solved k t) => vappSpine t sp
forceMeta x = pure x
forceMeta : Val -> Val
-- TODO - need to look up metas
forceMeta x = x
-- return renaming, the position is the new VVar
invert : Nat -> SnocList Val -> M (List Nat)
invert lvl sp = go sp []
where
go : SnocList Val -> List Nat -> M (List Nat)
go [<] acc = pure $ reverse acc
go (xs :< VVar k [<]) acc = do
if elem k acc
then throwError $ E (0,0) "non-linear pattern"
else go xs (k :: acc)
go _ _ = throwError $ E (0,0) "non-variable in pattern"
-- return renaming, the position is the new VVar
invert : Nat -> List Val -> m (List Nat)
invert lvl sp = go sp []
where
go : List Val -> List Nat -> m (List Nat)
go [] acc = pure acc
go ((VVar k []) :: xs) acc = do
if elem k acc
then throwError $ E (0,0) "non-linear pattern"
else go xs (k :: acc)
go _ _ = throwError $ E (0,0) "non-variable in pattern"
-- we have to "lift" the renaming when we go under a lambda
-- I think that essentially means our domain ix are one bigger, since we're looking at lvl
-- in the codomain, so maybe we can just keep that value
rename : Nat -> List Nat -> Nat -> Val -> m Tm
rename meta ren lvl v = go ren lvl v
where
go : List Nat -> Nat -> Val -> m Tm
goSpine : List Nat -> Nat -> Tm -> List Val -> m Tm
goSpine ren lvl tm [] = pure tm
goSpine ren lvl tm (x :: xs) = do
xtm <- go ren lvl x
goSpine ren lvl (App tm xtm) xs
-- we have to "lift" the renaming when we go under a lambda
-- I think that essentially means our domain ix are one bigger, since we're looking at lvl
-- in the codomain, so maybe we can just keep that value
rename : Nat -> List Nat -> Nat -> Val -> M Tm
rename meta ren lvl v = go ren lvl v
where
go : List Nat -> Nat -> Val -> M Tm
goSpine : List Nat -> Nat -> Tm -> SnocList Val -> M Tm
goSpine ren lvl tm [<] = pure tm
goSpine ren lvl tm (xs :< x) = do
xtm <- go ren lvl x
goSpine ren lvl (App tm xtm) xs
go ren lvl (VVar k sp) = case findIndex (== k) ren of
Nothing => throwError $ E (0,0) "scope/skolem thinger"
Just x => goSpine ren lvl (Bnd $ cast x) sp
go ren lvl (VRef nm sp) = goSpine ren lvl (Ref nm Nothing) sp
go ren lvl (VMeta ix sp) = if ix == meta
then throwError $ E (0,0) "meta occurs check"
else goSpine ren lvl (Meta ix) sp
go ren lvl (VLam n icit t) = pure (Lam n icit !(go (lvl :: ren) (S lvl) (t $$ VVar lvl [])))
go ren lvl (VPi n icit ty tm) = pure (Pi n icit !(go ren lvl ty) !(go (lvl :: ren) (S lvl) (tm $$ VVar lvl [])))
go ren lvl VU = pure U
go ren lvl (VVar k sp) = case findIndex (== k) ren of
Nothing => throwError $ E (0,0) "scope/skolem thinger"
Just x => goSpine ren lvl (Bnd $ cast x) sp
go ren lvl (VRef nm sp) = goSpine ren lvl (Ref nm Nothing) sp
go ren lvl (VMeta ix sp) = if ix == meta
then throwError $ E (0,0) "meta occurs check"
else goSpine ren lvl (Meta ix) sp
go ren lvl (VLam n icit t) = pure (Lam n icit !(go (lvl :: ren) (S lvl) !(t $$ VVar lvl [<])))
go ren lvl (VPi n icit ty tm) = pure (Pi n icit !(go ren lvl ty) !(go (lvl :: ren) (S lvl) !(tm $$ VVar lvl [<])))
go ren lvl VU = pure U
-- lams : Nat -> Tm -> Tm
-- lams 0 tm = tm
-- lams (S k) tm = Lam
lams : Nat -> Tm -> Tm
lams 0 tm = tm
lams (S k) tm = Lam "arg\{show k}" Explicit (lams k tm)
solve : Nat -> Nat -> List Val -> Val -> m ()
solve l m sp t = do
ren <- invert l sp
tm <- rename m ren l t
printLn "solution to \{show m} is \{show tm}"
solve : Nat -> Nat -> SnocList Val -> Val -> M ()
solve l m sp t = do
ren <- invert l sp
tm <- rename m ren l t
let tm = lams (length sp) tm
top <- get
soln <- eval [] CBN tm
solveMeta top m soln
pure ()
pure ()
unify : (l : Nat) -> Val -> Val -> M ()
unify : (l : Nat) -> Val -> Val -> m ()
unifySpine : Nat -> Bool -> SnocList Val -> SnocList Val -> M ()
unifySpine l False _ _ = throwError $ E (0,0) "unify failed"
unifySpine l True [<] [<] = pure ()
unifySpine l True (xs :< x) (ys :< y) = unify l x y >> unifySpine l True xs ys
unifySpine l True _ _ = throwError $ E (0,0) "meta spine length mismatch"
unifySpine : Nat -> Bool -> List Val -> List Val -> m ()
unifySpine l False _ _ = throwError $ E (0,0) "unify failed"
unifySpine l True [] [] = pure ()
unifySpine l True (x :: xs) (y :: ys) = unify l x y >> unifySpine l True xs ys
unifySpine l True _ _ = throwError $ E (0,0) "meta spine length mismatch"
unify l t u = case (forceMeta t, forceMeta u) of
(VLam _ _ t, VLam _ _ t' ) => unify (l + 1) (t $$ VVar l []) (t' $$ VVar l [])
(t, VLam _ _ t' ) => unify (l + 1) (t `vapp` VVar l []) (t' $$ VVar l [])
(VLam _ _ t, t' ) => unify (l + 1) (t $$ VVar l []) (t' `vapp` VVar l [])
(VPi _ _ a b, VPi _ _ a' b') => unify l a a' >> unify (S l) (b $$ VVar l []) (b' $$ VVar l [])
unify l t u = do
t' <- forceMeta t
u' <- forceMeta u
case (t',u') of
(VLam _ _ t, VLam _ _ t' ) => unify (l + 1) !(t $$ VVar l [<]) !(t' $$ VVar l [<])
(t, VLam _ _ t' ) => unify (l + 1) !(t `vapp` VVar l [<]) !(t' $$ VVar l [<])
(VLam _ _ t, t' ) => unify (l + 1) !(t $$ VVar l [<]) !(t' `vapp` VVar l [<])
(VPi _ _ a b, VPi _ _ a' b') => unify l a a' >> unify (S l) !(b $$ VVar l [<]) !(b' $$ VVar l [<])
(VVar k sp, VVar k' sp' ) => unifySpine l (k == k') sp sp'
(VRef n sp, VRef n' sp' ) => unifySpine l (n == n') sp sp'
(VMeta i sp, VMeta i' sp' ) => unifySpine l (i == i') sp sp'
@@ -103,75 +101,76 @@ parameters {0 m : Type -> Type} {auto _ : HasIO m} {auto _ : MonadError Error m}
_ => throwError $ E (0,0) "unify failed"
export
infer : Context -> Raw -> m (Tm, Val)
export
infer : Context -> Raw -> M (Tm, Val)
export
check : Context -> Raw -> Val -> m Tm
check ctx (RSrcPos x tm) ty = check ({pos := x} ctx) tm ty
check ctx (RLam nm icit tm) ty = case ty of
(VPi pinm icit a b) => do
-- TODO icit
let var = VVar (length ctx.env) []
let ctx' = extend ctx nm a
tm' <- check ctx' tm (b $$ var)
pure $ Lam nm icit tm'
other => error [(DS "Expected pi type, got \{show $ quote 0 ty}")]
check ctx tm ty = do
(tm', ty') <- infer ctx tm
-- This is where the conversion check / pattern unification go
unify ctx.lvl ty' ty
-- if quote 0 ty /= quote 0 ty' then
-- error [DS "type mismatch got", DD (quote 0 ty'), DS "expected", DD (quote 0 ty)]
-- else pure tm'
pure tm'
infer ctx (RVar nm) = go 0 ctx.types
where
go : Nat -> Vect n (String, Val) -> m (Tm, Val)
go i [] = case lookup nm top of
Just (MkEntry name ty (Fn t)) => pure (Ref nm (Just t), eval [] CBN ty)
Just (MkEntry name ty _) => pure (Ref nm Nothing, eval [] CBN ty)
Nothing => error [DS "\{show nm} not in scope"]
go i ((x, ty) :: xs) = if x == nm then pure $ (Bnd i, ty)
else go (i + 1) xs
-- need environment of name -> type..
infer ctx (RApp t u icit) = do
-- icit will be used for insertion, lets get this working first...
(t, tty) <- infer ctx t
case tty of
(VPi str icit' a b) => do
u <- check ctx u a
pure (App t u, b $$ eval ctx.env CBN t)
_ => error [DS "Expected Pi type"]
infer ctx RU = pure (U, VU) -- YOLO
infer ctx (RPi nm icit ty ty2) = do
ty' <- check ctx ty VU
let vty' := eval ctx.env CBN ty'
let nm := fromMaybe "_" nm
ty2' <- check (extend ctx nm vty') ty2 VU
pure (Pi nm icit ty' ty2', VU)
infer ctx (RLet str tm tm1 tm2) = ?rhs_5
infer ctx (RSrcPos x tm) = infer ({pos := x} ctx) tm
infer ctx (RAnn tm rty) = do
ty <- check ctx rty VU
let vty = eval ctx.env CBN ty
tm <- check ctx tm vty
pure (tm, vty)
export
check : Context -> Raw -> Val -> M Tm
check ctx (RSrcPos x tm) ty = check ({pos := x} ctx) tm ty
check ctx (RLam nm icit tm) ty = case ty of
(VPi pinm icit a b) => do
-- TODO icit
let var = VVar (length ctx.env) [<]
let ctx' = extend ctx nm a
tm' <- check ctx' tm !(b $$ var)
pure $ Lam nm icit tm'
other => error [(DS "Expected pi type, got \{show !(quote 0 ty)}")]
check ctx tm ty = do
(tm', ty') <- infer ctx tm
-- This is where the conversion check / pattern unification go
unify ctx.lvl ty' ty
-- if quote 0 ty /= quote 0 ty' then
-- error [DS "type mismatch got", DD (quote 0 ty'), DS "expected", DD (quote 0 ty)]
-- else pure tm'
pure tm'
infer ctx (RLam str icit tm) = error [DS "can't infer lambda"]
infer ctx RHole = do
ty <- freshMeta ctx
let vty = eval ctx.env CBN ty
tm <- freshMeta ctx
pure (tm, vty)
infer ctx tm = error [DS "Implement infer \{show tm}"]
infer ctx (RVar nm) = go 0 ctx.types
where
go : Nat -> Vect n (String, Val) -> M (Tm, Val)
go i [] = case lookup nm !(get) of
Just (MkEntry name ty (Fn t)) => pure (Ref nm (Just t), !(eval [] CBN ty))
Just (MkEntry name ty _) => pure (Ref nm Nothing, !(eval [] CBN ty))
Nothing => error [DS "\{show nm} not in scope"]
go i ((x, ty) :: xs) = if x == nm then pure $ (Bnd i, ty)
else go (i + 1) xs
-- need environment of name -> type..
infer ctx (RApp t u icit) = do
-- icit will be used for insertion, lets get this working first...
(t, tty) <- infer ctx t
case tty of
(VPi str icit' a b) => do
u <- check ctx u a
pure (App t u, !(b $$ !(eval ctx.env CBN u)))
_ => error [DS "Expected Pi type"]
infer ctx RU = pure (U, VU) -- YOLO
infer ctx (RPi nm icit ty ty2) = do
ty' <- check ctx ty VU
vty' <- eval ctx.env CBN ty'
let nm := fromMaybe "_" nm
ty2' <- check (extend ctx nm vty') ty2 VU
pure (Pi nm icit ty' ty2', VU)
infer ctx (RLet str tm tm1 tm2) = error [DS "implement RLet"]
infer ctx (RSrcPos x tm) = infer ({pos := x} ctx) tm
infer ctx (RAnn tm rty) = do
ty <- check ctx rty VU
vty <- eval ctx.env CBN ty
tm <- check ctx tm vty
pure (tm, vty)
-- I don't have types for these yet...
-- infer ctx (RLit (LString str)) = ?rhs_10
-- infer ctx (RLit (LInt i)) = ?rhs_11
-- infer ctx (RLit (LBool x)) = ?rhs_12
-- infer ctx (RCase tm xs) = ?rhs_9
-- infer ctx RHole = ?todo_meta2
-- The idea here is to insert a hole for a parse error
-- infer ctx (RParseError str) = ?todo_insert_meta
infer ctx (RLam str icit tm) = error [DS "can't infer lambda"]
infer ctx RHole = do
ty <- freshMeta ctx
vty <- eval ctx.env CBN ty
tm <- freshMeta ctx
pure (tm, vty)
infer ctx tm = error [DS "Implement infer \{show tm}"]
-- I don't have types for these yet...
-- infer ctx (RLit (LString str)) = ?rhs_10
-- infer ctx (RLit (LInt i)) = ?rhs_11
-- infer ctx (RLit (LBool x)) = ?rhs_12
-- infer ctx (RCase tm xs) = ?rhs_9
-- infer ctx RHole = ?todo_meta2
-- The idea here is to insert a hole for a parse error
-- infer ctx (RParseError str) = ?todo_insert_meta

View File

@@ -1,5 +1,5 @@
module Lib.Parser
import Lib.TT
import Lib.Types
-- The SourcePos stuff is awkward later on. We might want bounds on productions
-- But we might want to consider something more generic and closer to lean?
@@ -241,9 +241,14 @@ parseData = do
-- TODO - turn decls into something more useful
pure $ Data name ty decls
-- Not sure what I want here.
-- I can't get a Tm without a type, and then we're covered by the other stuff
parseNorm : Parser Decl
parseNorm = DCheck <$ keyword "#check" <*> typeExpr <* keyword ":" <*> typeExpr
export
parseDecl : Parser Decl
parseDecl = parseImport <|> parseSig <|> parseDef <|> parseData
parseDecl = parseImport <|> parseSig <|> parseDef <|> parseNorm <|> parseData
export
parseMod : Parser Module

View File

@@ -7,7 +7,7 @@ module Lib.TT
-- For SourcePos
import Lib.Parser.Impl
import Lib.Prettier
import Lib.Types
import Control.Monad.Error.Interface
import Data.IORef
@@ -16,302 +16,7 @@ import Data.List
import Data.Vect
import Data.SortedMap
public export
Name : Type
Name = String
public export
data Icit = Implicit | Explicit
%name Icit icit
public export
data BD = Bound | Defined
public export
data Tm : Type where
Bnd : Nat -> Tm
-- Maybe Def here instead of Maybe Tm, we'll have DCon, TCon, etc.
Ref : String -> Maybe Tm -> Tm
Meta : Nat -> Tm
-- kovacs optimization, I think we can App out meta instead
-- InsMeta : Nat -> List BD -> Tm
Lam : Name -> Icit -> Tm -> Tm
App : Tm -> Tm -> Tm
U : Tm
Pi : Name -> Icit -> Tm -> Tm -> Tm
Let : Name -> Icit -> Tm -> Tm -> Tm -> Tm
%name Tm t, u, v
public export
Show Tm where
show (Bnd k) = "(Bnd \{show k})"
show (Ref str _) = "(Ref \{show str})"
show (Lam nm Implicit t) = "(λ {\{nm}} => \{show t})"
show (Lam nm Explicit t) = "(λ \{nm} => \{show t})"
show (App t u) = "(\{show t} \{show u})"
show (Meta i) = "(Meta \{show i})"
show U = "U"
show (Pi str icit t u) = "(\{str} : \{show t} => \{show u})"
show (Let str icit t u v) = "let \{str} : \{show t} = \{show u} in \{show v}"
-- I can't really show val because it's HOAS...
-- TODO derive
export
Eq Icit where
Implicit == Implicit = True
Explicit == Explicit = True
_ == _ = False
||| Eq on Tm. We've got deBruijn indices, so I'm not comparing names
export
Eq (Tm) where
-- (Local x) == (Local y) = x == y
(Bnd x) == (Bnd y) = x == y
(Ref x _) == (Ref y _) = x == y
(Lam n icit t) == (Lam n' icit' t') = icit == icit' && t == t'
(App t u) == App t' u' = t == t' && u == u'
U == U = True
(Pi n icit t u) == (Pi n' icit' t' u') = icit == icit' && t == t' && u == u'
(Let n icit t u v) == (Let n' icit' t' u' v') = t == t' && u == u' && v == v'
_ == _ = False
public export
Pretty Tm where
pretty (Bnd k) = ?rhs_0
pretty (Ref str mt) = text str
pretty (Meta k) = text "?m\{show k}"
pretty (Lam str Implicit t) = text "(\\ {\{str}} => " <+> pretty t <+> ")"
pretty (Lam str Explicit t) = text "(\\ \{str} => " <+> pretty t <+> ")"
pretty (App t u) = text "(" <+> pretty t <+> pretty u <+> ")"
pretty U = "U"
pretty (Pi str icit t u) = text "(" <+> text str <+> ":" <+> pretty t <+> "=>" <+> pretty u <+> ")"
pretty (Let str icit t u v) = text "let" <+> text str <+> ":" <+> pretty t <+> "=" <+> pretty u
-- public export
-- data Closure : Nat -> Type
data Val : Type
-- IS/TypeTheory.idr is calling this a Kripke function space
-- Yaffle, IS/TypeTheory use a function here.
-- Kovacs, Idris use Env and Tm
-- in cctt kovacs refers to this choice as defunctionalization vs HOAS
-- https://github.com/AndrasKovacs/cctt/blob/main/README.md#defunctionalization
-- the tradeoff is access to internals of the function
-- Yaffle is vars -> vars here
public export
data Closure : Type
public export
data Val : Type where
-- This will be local / flex with spine.
VVar : (k : Nat) -> (sp : List Val) -> Val
-- I wanted the Maybe Tm in here, but for now we're always expanding.
-- Maybe this is where I glue
VRef : (nm : String) -> (sp : List Val) -> Val
-- we'll need to look this up in ctx with IO
VMeta : (ix : Nat) -> (sp : List Val) -> Val
VLam : Name -> Icit -> Closure -> Val
VPi : Name -> Icit -> Lazy Val -> Closure -> Val
VU : Val
public export
Env : Type
Env = List Val
public export
data Mode = CBN | CBV
export
eval : Env -> Mode -> Tm -> Val
data Closure = MkClosure Env Tm
public export
($$) : {auto mode : Mode} -> Closure -> Val -> Val
($$) (MkClosure env tm) u = eval (u :: env) mode tm
public export
infixl 8 $$
export
vapp : Val -> Val -> Val
vapp (VLam _ icit t) u = t $$ u
vapp (VVar k sp) u = VVar k (u :: sp)
vapp (VRef nm sp) u = VRef nm (u :: sp)
vapp (VMeta k sp) u = VMeta k (u :: sp)
vapp _ _ = ?throw_impossible
bind : Val -> Env -> Env
bind v env = v :: env
-- Do we want a def in here instead? We'll need DCon/TCon eventually
-- I need to be aggressive about reduction, I guess. I'll figure it out
-- later, maybe need lazy glued values.
eval env mode (Ref x (Just tm)) = eval env mode tm
eval env mode (Ref x Nothing) = VRef x []
eval env mode (App (Ref x (Just tm)) u) = vapp (eval env mode tm) (eval env mode u)
eval env mode (App t u) = vapp (eval env mode t) (eval env mode u)
eval env mode U = VU
eval env mode (Meta i) = VMeta i []
eval env mode (Lam x icit t) = VLam x icit (MkClosure env t)
eval env mode (Pi x icit a b) = VPi x icit (eval env mode a) (MkClosure env b)
eval env mode (Let x icit ty t u) = eval (eval env mode t :: env) mode u
eval env mode (Bnd i) = let Just rval = getAt i env | _ => ?out_of_index
in rval
export
quote : (lvl : Nat) -> Val -> Tm
quoteSp : (lvl : Nat) -> Tm -> List Val -> Tm
quoteSp lvl t [] = t
quoteSp lvl t (x :: xs) = quoteSp lvl (App t (quote lvl x)) xs
quote l (VVar k sp) = quoteSp l (Bnd ((l `minus` k) `minus` 1)) sp -- level to index
quote l (VMeta i sp) = quoteSp l (Meta i) sp
quote l (VLam x icit t) = Lam x icit (quote (S l) (t $$ VVar l []))
quote l (VPi x icit a b) = Pi x icit (quote l a) (quote (S l) (b $$ VVar l []))
quote l VU = U
quote l (VRef n sp) = quoteSp l (Ref n Nothing) sp
-- Can we assume closed terms?
-- ezoo only seems to use it at [], but essentially does this:
export
nf : Env -> Tm -> Tm
nf env t = quote (length env) (eval env CBN t)
{-
smalltt
smalltt gets into weird haskell weeds in eval - shifting top level to the left
and tagging meta vs top with a bit.
I think something subtle is going on with laziness on Elaboration.hs:300
yeah, and define is even inlined.
So it has a top context, and clears out almost everything for processing a def in
a different kind of context.
we very much need an idea of local context for metas. I don't want to abstract over
the entire program.
So I guess we have top and local then?
With haskell syntax. I think we can have Axiom for claims and rewrite to def later.
Hmm, so given ezoo, if I'm going simple, I could keep BDs short, and use the normal
context. (Zoo4.lean:222) I'd probably still need an undefined/axiom marker as a value?
ok, so with just one context, Env is List Val and we're getting Tm back from type checking.
Can I get val back? Do we need to quote? What happens if we don't?
-}
-- FIXME remove List BD
public export
data MetaEntry = Unsolved Nat (List BD) | Solved Nat Val
public export
record MetaContext where
constructor MC
metas : List MetaEntry
next : Nat
public export
data Def = Axiom | TCon (List String) | DCon Nat | Fn Tm
Show Def where
show Axiom = "axiom"
show (TCon strs) = "TCon \{show strs}"
show (DCon k) = "DCon \{show k}"
show (Fn t) = "Fn \{show t}"
||| entry in the top level context
public export
record TopEntry where
constructor MkEntry
name : String
type : Tm
def : Def
-- FIXME snoc
export
Show TopEntry where
show (MkEntry name type def) = "\{name} : \{show type} := \{show def}"
||| Top level context.
||| Most of the reason this is separate is to have a different type
||| `Def` for the entries.
|||
||| The price is that we have names in addition to levels. Do we want to
||| expand these during conversion?
public export
record TopContext where
constructor MkTop
-- We'll add a map later?
defs : List TopEntry
metas : IORef MetaContext
-- metas : TODO
-- we'll use this for typechecking, but need to keep a TopContext around too.
public export
record Context where
constructor MkCtx
lvl : Nat
-- shall we use lvl as an index?
env : Env -- Values in scope
types : Vect lvl (String, Val) -- types and names in scope
-- so we'll try "bds" determines length of local context
bds : List BD -- bound or defined
pos : SourcePos -- the last SourcePos that we saw
-- We only need this here if we don't pass TopContext
-- top : TopContext
metas : IORef MetaContext
export
freshMeta : HasIO m => Context -> m Tm
freshMeta ctx = do
mc <- readIORef ctx.metas
writeIORef ctx.metas $ { next $= S, metas $= (Unsolved mc.next ctx.bds ::) } mc
pure $ applyBDs 0 (Meta mc.next) ctx.bds
where
-- hope I got the right order here :)
applyBDs : Nat -> Tm -> List BD -> Tm
applyBDs k t [] = t
applyBDs k t (Bound :: xs) = applyBDs (S k) (App t (Bnd k)) xs
applyBDs k t (Defined :: xs) = applyBDs (S k) t xs
-- solveMeta : HasIO m => Context -> m Tm
-- solveMeta ctx = do
-- mc <- readIORef ctx.metas
-- we need more of topcontext later - Maybe switch it up so we're not passing
-- around top
export
mkCtx : IORef MetaContext -> Context
mkCtx metas = MkCtx 0 [] [] [] (0,0) metas
export partial
Show Context where
show ctx = "Context \{show $ map fst $ ctx.types}"
-- TODO Pretty Context
-- idea cribbed from pi-forall
-- Errors cribbed from pi-forall
public export
data ErrorSeg : Type where
DD : Pretty a => a -> ErrorSeg
@@ -322,15 +27,67 @@ toDoc (DD x) = pretty x
toDoc (DS str) = text str
export
error : {0 m : Type -> Type} -> {auto _ : MonadError Error m} ->
{auto ctx : Context} -> List ErrorSeg -> m a
error : {auto ctx : Context} -> List ErrorSeg -> M a
error xs = throwError $ E ctx.pos (render 80 $ spread $ map toDoc xs)
export
error' : String -> M a
error' msg = throwError $ E (0,0) msg
export
freshMeta : Context -> M Tm
freshMeta ctx = do
mc <- readIORef ctx.metas
writeIORef ctx.metas $ { next $= S, metas $= (Unsolved mc.next ctx.bds ::) } mc
pure $ applyBDs 0 (Meta mc.next) ctx.bds
where
-- hope I got the right order here :)
applyBDs : Nat -> Tm -> List BD -> Tm
applyBDs k t [] = t
applyBDs k t (Bound :: xs) = App (applyBDs (S k) t xs) (Bnd k)
applyBDs k t (Defined :: xs) = applyBDs (S k) t xs
export
lookupMeta : Nat -> M MetaEntry
lookupMeta ix = do
ctx <- get
mc <- readIORef ctx.metas
go mc.metas
where
go : List MetaEntry -> M MetaEntry
go [] = error' "Meta \{show ix} not found"
go (meta@(Unsolved k ys) :: xs) = if k == ix then pure meta else go xs
go (meta@(Solved k x) :: xs) = if k == ix then pure meta else go xs
export
solveMeta : TopContext -> Nat -> Val -> M ()
solveMeta ctx ix tm = do
mc <- readIORef ctx.metas
metas <- go mc.metas [<]
writeIORef ctx.metas $ {metas := metas} mc
where
go : List MetaEntry -> SnocList MetaEntry -> M (List MetaEntry)
go [] _ = error' "Meta \{show ix} not found"
go (meta@(Unsolved k _) :: xs) lhs = if k == ix
then pure $ lhs <>> (Solved k tm :: xs)
else go xs (lhs :< meta)
go (meta@(Solved k _) :: xs) lhs = if k == ix
then error' "Meta \{show ix} already solved!"
else go xs (lhs :< meta)
export partial
Show Context where
show ctx = "Context \{show $ map fst $ ctx.types}"
-- TODO Pretty Context
||| add a binding to environment
export
extend : Context -> String -> Val -> Context
extend ctx name ty =
{ lvl $= S, env $= (VVar ctx.lvl [] ::), types $= ((name, ty) ::), bds $= (Bound ::) } ctx
{ lvl $= S, env $= (VVar ctx.lvl [<] ::), types $= ((name, ty) ::), bds $= (Bound ::) } ctx
-- I guess we define things as values?
export
@@ -340,12 +97,86 @@ define ctx name val ty =
-- not used
lookup : {0 m : Type -> Type} -> {auto _ : MonadError String m} ->
Context -> String -> m Val
lookup : Context -> String -> M Val
lookup ctx nm = go ctx.types
where
go : Vect n (String,Val) -> m Val
go [] = throwError "Name \{nm} not in scope"
go : Vect n (String,Val) -> M Val
go [] = error' "Name \{nm} not in scope"
go ((n, ty) :: xs) = if n == nm then pure ty else go xs
-- Need to wire in the metas...
-- if it's top / ctx / IORef, I also need IO...
-- if I want errors, I need m anyway. I've already got an error down there.
export
eval : Env -> Mode -> Tm -> M Val
public export
($$) : {auto mode : Mode} -> Closure -> Val -> M Val
($$) {mode} (MkClosure env tm) u = eval (u :: env) mode tm
public export
infixl 8 $$
export
vapp : Val -> Val -> M Val
vapp (VLam _ icit t) u = t $$ u
vapp (VVar k sp) u = pure $ VVar k (sp :< u)
vapp (VRef nm sp) u = pure $ VRef nm (sp :< u)
vapp (VMeta k sp) u = pure $ VMeta k (sp :< u)
vapp t u = error' "impossible in vapp \{show t} to \{show u}"
export
vappSpine : Val -> SnocList Val -> M Val
vappSpine t [<] = pure t
vappSpine t (xs :< x) = vapp !(vappSpine t xs) x
bind : Val -> Env -> Env
bind v env = v :: env
-- Do we want a def in here instead? We'll need DCon/TCon eventually
-- I need to be aggressive about reduction, I guess. I'll figure it out
-- later, maybe need lazy glued values.
eval env mode (Ref x (Just tm)) = eval env mode tm
eval env mode (Ref x Nothing) = pure $ VRef x [<]
eval env mode (App t u) = vapp !(eval env mode t) !(eval env mode u)
eval env mode U = pure VU
eval env mode (Meta i) =
case !(lookupMeta i) of
(Unsolved k xs) => pure $ VMeta i [<]
(Solved k t) => pure $ t
eval env mode (Lam x icit t) = pure $ VLam x icit (MkClosure env t)
eval env mode (Pi x icit a b) = pure $ VPi x icit !(eval env mode a) (MkClosure env b)
eval env mode (Let x icit ty t u) = eval (!(eval env mode t) :: env) mode u
eval env mode (Bnd i) = case getAt i env of
Just rval => pure rval
Nothing => error' "Bad deBruin index \{show i}"
export
quote : (lvl : Nat) -> Val -> M Tm
quoteSp : (lvl : Nat) -> Tm -> SnocList Val -> M Tm
quoteSp lvl t [<] = pure t
quoteSp lvl t (xs :< x) =
pure $ App !(quoteSp lvl t xs) !(quote lvl x)
-- quoteSp lvl (App t !(quote lvl x)) xs -- snoc says previous is right
quote l (VVar k sp) = if k < l
then quoteSp l (Bnd ((l `minus` k) `minus` 1)) sp -- level to index
else ?borken
quote l (VMeta i sp) = quoteSp l (Meta i) sp
quote l (VLam x icit t) = pure $ Lam x icit !(quote (S l) !(t $$ VVar l [<]))
quote l (VPi x icit a b) = pure $ Pi x icit !(quote l a) !(quote (S l) !(b $$ VVar l [<]))
quote l VU = pure U
quote l (VRef n sp) = quoteSp l (Ref n Nothing) sp
-- Can we assume closed terms?
-- ezoo only seems to use it at [], but essentially does this:
export
nf : Env -> Tm -> M Tm
nf env t = quote (length env) !(eval env CBN t)

View File

@@ -1,7 +1,7 @@
module Lib.TopContext
import Data.String
import Lib.TT
import Lib.Types
import Data.IORef
-- I want unique ids, to be able to lookup, update, and a Ref so

258
src/Lib/Types.idr Normal file
View File

@@ -0,0 +1,258 @@
module Lib.Types
-- I'm not sure this is related, or just a note to self (Presheaves on Porpoise)
-- maybe watch https://www.youtube.com/watch?v=3gef0_NFz8Q
-- or drop the indices for now.
-- For SourcePos
import Lib.Parser.Impl
import Lib.Prettier
import public Control.Monad.Error.Either
import public Control.Monad.Error.Interface
import public Control.Monad.State
import Data.IORef
import Data.Fin
import Data.List
import Data.SnocList
import Data.Vect
import Data.SortedMap
public export
Name : Type
Name = String
public export
data Icit = Implicit | Explicit
%name Icit icit
public export
data BD = Bound | Defined
public export
data Tm : Type where
Bnd : Nat -> Tm
-- Maybe Def here instead of Maybe Tm, we'll have DCon, TCon, etc.
Ref : String -> Maybe Tm -> Tm
Meta : Nat -> Tm
-- kovacs optimization, I think we can App out meta instead
-- InsMeta : Nat -> List BD -> Tm
Lam : Name -> Icit -> Tm -> Tm
App : Tm -> Tm -> Tm
U : Tm
Pi : Name -> Icit -> Tm -> Tm -> Tm
Let : Name -> Icit -> Tm -> Tm -> Tm -> Tm
%name Tm t, u, v
public export
Show Tm where
show (Bnd k) = "(Bnd \{show k})"
show (Ref str _) = "(Ref \{show str})"
show (Lam nm Implicit t) = "(\\ {\{nm}} => \{show t})"
show (Lam nm Explicit t) = "(\\ \{nm} => \{show t})"
show (App t u) = "(\{show t} \{show u})"
show (Meta i) = "(Meta \{show i})"
show U = "U"
show (Pi str icit t u) = "(\{str} : \{show t} => \{show u})"
show (Let str icit t u v) = "let \{str} : \{show t} = \{show u} in \{show v}"
-- I can't really show val because it's HOAS...
-- TODO derive
export
Eq Icit where
Implicit == Implicit = True
Explicit == Explicit = True
_ == _ = False
||| Eq on Tm. We've got deBruijn indices, so I'm not comparing names
export
Eq (Tm) where
-- (Local x) == (Local y) = x == y
(Bnd x) == (Bnd y) = x == y
(Ref x _) == (Ref y _) = x == y
(Lam n icit t) == (Lam n' icit' t') = icit == icit' && t == t'
(App t u) == App t' u' = t == t' && u == u'
U == U = True
(Pi n icit t u) == (Pi n' icit' t' u') = icit == icit' && t == t' && u == u'
(Let n icit t u v) == (Let n' icit' t' u' v') = t == t' && u == u' && v == v'
_ == _ = False
public export
Pretty Tm where
pretty (Bnd k) = ?rhs_0
pretty (Ref str mt) = text str
pretty (Meta k) = text "?m\{show k}"
pretty (Lam str Implicit t) = text "(\\ {\{str}} => " <+> pretty t <+> ")"
pretty (Lam str Explicit t) = text "(\\ \{str} => " <+> pretty t <+> ")"
pretty (App t u) = text "(" <+> pretty t <+> pretty u <+> ")"
pretty U = "U"
pretty (Pi str icit t u) = text "(" <+> text str <+> ":" <+> pretty t <+> "=>" <+> pretty u <+> ")"
pretty (Let str icit t u v) = text "let" <+> text str <+> ":" <+> pretty t <+> "=" <+> pretty u
-- public export
-- data Closure : Nat -> Type
data Val : Type
-- IS/TypeTheory.idr is calling this a Kripke function space
-- Yaffle, IS/TypeTheory use a function here.
-- Kovacs, Idris use Env and Tm
-- in cctt kovacs refers to this choice as defunctionalization vs HOAS
-- https://github.com/AndrasKovacs/cctt/blob/main/README.md#defunctionalization
-- the tradeoff is access to internals of the function
-- Yaffle is vars -> vars here
public export
data Closure : Type
public export
data Val : Type where
-- This will be local / flex with spine.
VVar : (k : Nat) -> (sp : SnocList Val) -> Val
-- I wanted the Maybe Tm in here, but for now we're always expanding.
-- Maybe this is where I glue
VRef : (nm : String) -> (sp : SnocList Val) -> Val
-- we'll need to look this up in ctx with IO
VMeta : (ix : Nat) -> (sp : SnocList Val) -> Val
VLam : Name -> Icit -> Closure -> Val
VPi : Name -> Icit -> Lazy Val -> Closure -> Val
VU : Val
Show Icit where
show Implicit = "I"
show Explicit = "E"
Show Closure
covering export
Show Val where
show (VVar k sp) = "(%var \{show k} \{show sp})"
show (VRef nm sp) = "(%ref \{nm} \{show sp})"
show (VMeta ix sp) = "(%meta \{show ix} \{show sp})"
show (VLam str icit x) = "(%lam \{str} \{show icit} \{show x})"
show (VPi str icit x y) = "(%pi \{str} \{show icit} \{show x} \{show y})"
show VU = "U"
public export
Env : Type
Env = List Val
public export
data Mode = CBN | CBV
public export
data Closure = MkClosure Env Tm
covering
Show Closure where
show (MkClosure xs t) = "(%cl \{show xs} \{show t})"
{-
smalltt
smalltt gets into weird haskell weeds in eval - shifting top level to the left
and tagging meta vs top with a bit.
I think something subtle is going on with laziness on Elaboration.hs:300
yeah, and define is even inlined.
So it has a top context, and clears out almost everything for processing a def in
a different kind of context.
we very much need an idea of local context for metas. I don't want to abstract over
the entire program.
So I guess we have top and local then?
With haskell syntax. I think we can have Axiom for claims and rewrite to def later.
Hmm, so given ezoo, if I'm going simple, I could keep BDs short, and use the normal
context. (Zoo4.lean:222) I'd probably still need an undefined/axiom marker as a value?
ok, so with just one context, Env is List Val and we're getting Tm back from type checking.
Can I get val back? Do we need to quote? What happens if we don't?
-}
-- FIXME remove List BD
public export
data MetaEntry = Unsolved Nat (List BD) | Solved Nat Val
public export
record MetaContext where
constructor MC
metas : List MetaEntry
next : Nat
public export
data Def = Axiom | TCon (List String) | DCon Nat | Fn Tm
Show Def where
show Axiom = "axiom"
show (TCon strs) = "TCon \{show strs}"
show (DCon k) = "DCon \{show k}"
show (Fn t) = "Fn \{show t}"
||| entry in the top level context
public export
record TopEntry where
constructor MkEntry
name : String
type : Tm
def : Def
-- FIXME snoc
export
Show TopEntry where
show (MkEntry name type def) = "\{name} : \{show type} := \{show def}"
||| Top level context.
||| Most of the reason this is separate is to have a different type
||| `Def` for the entries.
|||
||| The price is that we have names in addition to levels. Do we want to
||| expand these during normalization?
public export
record TopContext where
constructor MkTop
-- We'll add a map later?
defs : List TopEntry
metas : IORef MetaContext
-- metas : TODO
-- we'll use this for typechecking, but need to keep a TopContext around too.
public export
record Context where
constructor MkCtx
lvl : Nat
-- shall we use lvl as an index?
env : Env -- Values in scope
types : Vect lvl (String, Val) -- types and names in scope
-- so we'll try "bds" determines length of local context
bds : List BD -- bound or defined
pos : SourcePos -- the last SourcePos that we saw
-- We only need this here if we don't pass TopContext
-- top : TopContext
metas : IORef MetaContext
public export
M : Type -> Type
M = (StateT TopContext (EitherT Impl.Error IO))
-- we need more of topcontext later - Maybe switch it up so we're not passing
-- around top
export
mkCtx : IORef MetaContext -> Context
mkCtx metas = MkCtx 0 [] [] [] (0,0) metas