Typeclass works for Monad
This commit is contained in:
@@ -93,7 +93,7 @@ compileTerm tm@(App _ _ _) with (funArgs tm)
|
||||
-- apply (CRef "Meta\{show k}") args' [<] 0
|
||||
arity <- case meta of
|
||||
-- maybe throw
|
||||
(Unsolved x j ctx _ _) => pure 0 -- FIXME # of Bound in ctx.bds
|
||||
(Unsolved x j ctx _ _ _) => pure 0 -- FIXME # of Bound in ctx.bds
|
||||
(Solved j tm) => pure $ getArity !(quote 0 tm)
|
||||
apply (CRef "Meta\{show k}") args' [<] arity
|
||||
_ | (t@(Ref fc nm _), args) = do
|
||||
|
||||
@@ -9,6 +9,7 @@ import Lib.Prettier
|
||||
import Data.List
|
||||
import Data.Vect
|
||||
import Data.String
|
||||
import Data.IORef
|
||||
import Lib.Types
|
||||
import Lib.Eval
|
||||
import Lib.TopContext
|
||||
@@ -53,7 +54,7 @@ lookupDef ctx name = go 0 ctx.types ctx.env
|
||||
export
|
||||
forceMeta : Val -> M Val
|
||||
forceMeta (VMeta fc ix sp) = case !(lookupMeta ix) of
|
||||
(Unsolved pos k xs _ _) => pure (VMeta fc ix sp)
|
||||
(Unsolved pos k xs _ _ _) => pure (VMeta fc ix sp)
|
||||
(Solved k t) => vappSpine t sp >>= forceMeta
|
||||
forceMeta x = pure x
|
||||
|
||||
@@ -70,7 +71,7 @@ tryEval k sp =
|
||||
-- Force far enough to compare types
|
||||
forceType : Val -> M Val
|
||||
forceType (VMeta fc ix sp) = case !(lookupMeta ix) of
|
||||
(Unsolved x k xs _ _) => pure (VMeta fc ix sp)
|
||||
(Unsolved x k xs _ _ _) => pure (VMeta fc ix sp)
|
||||
(Solved k t) => vappSpine t sp >>= forceType
|
||||
forceType x@(VRef fc nm _ sp) = fromMaybe x <$> tryEval nm sp
|
||||
forceType x = pure x
|
||||
@@ -87,6 +88,30 @@ Semigroup UnifyResult where
|
||||
Monoid UnifyResult where
|
||||
neutral = MkResult []
|
||||
|
||||
-- We need to switch to SortedMap here
|
||||
export
|
||||
updateMeta : Context -> Nat -> (MetaEntry -> M MetaEntry) -> M ()
|
||||
updateMeta ctx ix f = do
|
||||
mc <- readIORef ctx.metas
|
||||
metas <- go mc.metas
|
||||
writeIORef ctx.metas $ {metas := metas} mc
|
||||
where
|
||||
go : List MetaEntry -> M (List MetaEntry)
|
||||
go [] = error' "Meta \{show ix} not found"
|
||||
go (x@((Unsolved y k z w v ys)) :: xs) = if k == ix then (::xs) <$> f x else (x ::) <$> go xs
|
||||
go (x@((Solved k y)) :: xs) = if k == ix then (::xs) <$> f x else (x ::) <$> go xs
|
||||
|
||||
export
|
||||
addConstraint : Context -> Nat -> SnocList Val -> Val -> M ()
|
||||
addConstraint ctx ix sp tm = do
|
||||
mc <- readIORef ctx.metas
|
||||
updateMeta ctx ix $ \case
|
||||
(Unsolved pos k a b c cons) => do
|
||||
info (getFC tm) "Add constraint m\{show ix} \{show sp} =?= \{show tm}"
|
||||
pure (Unsolved pos k a b c (MkMc (getFC tm) ctx sp tm :: cons))
|
||||
(Solved k tm) => error' "Meta \{show k} already solved"
|
||||
pure ()
|
||||
|
||||
parameters (ctx: Context)
|
||||
-- return renaming, the position is the new VVar
|
||||
invert : Nat -> SnocList Val -> M (List Nat)
|
||||
@@ -140,28 +165,51 @@ parameters (ctx: Context)
|
||||
-- REVIEW can I get better names in here?
|
||||
lams (S k) tm = Lam emptyFC "arg_\{show k}" (lams k tm)
|
||||
|
||||
|
||||
export
|
||||
unify : (l : Nat) -> Val -> Val -> M UnifyResult
|
||||
|
||||
export
|
||||
solve : (lvl : Nat) -> (k : Nat) -> SnocList Val -> Val -> M ()
|
||||
solve l m sp t = do
|
||||
meta@(Unsolved metaFC ix ctx ty kind) <- lookupMeta m
|
||||
meta@(Unsolved metaFC ix ctx ty kind cons) <- lookupMeta m
|
||||
| _ => error (getFC t) "Meta \{show m} already solved!"
|
||||
debug "SOLVE \{show m} \{show kind} lvl \{show l} sp \{show sp} is \{show t}"
|
||||
let size = length $ filter (\x => x == Bound) $ toList ctx.bds
|
||||
debug "\{show m} size is \{show size}"
|
||||
if (length sp /= size) then do
|
||||
debug "\{show m} size is \{show size} sps \{show $ length sp}"
|
||||
let True = length sp == size
|
||||
| _ => do
|
||||
-- need INFO that works like debug.
|
||||
-- FIXME we probably need to hold onto the constraint and recheck when we solve the meta?
|
||||
info (getFC t) "meta \{show m} (\{show ix}) applied to \{show $ length sp} args instead of \{show size}"
|
||||
debug "CONSTRAINT m\{show ix} \{show sp} =?= \{show t}"
|
||||
-- error (getFC t) "meta \{show m} applied to \{show $ length sp} args instead of \{show size}"
|
||||
else do
|
||||
debug "meta \{show meta}"
|
||||
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 ()
|
||||
|
||||
-- add constraint to meta m
|
||||
-- we can keep a list and run them when it is solved.
|
||||
addConstraint ctx m sp t
|
||||
|
||||
debug "meta \{show meta}"
|
||||
ren <- invert l sp
|
||||
tm <- rename m ren l t
|
||||
let tm = lams (length sp) tm
|
||||
top <- get
|
||||
soln <- eval [] CBN tm
|
||||
|
||||
updateMeta ctx m $ \case
|
||||
(Unsolved pos k _ _ _ cons) => do
|
||||
putStrLn "INFO at \{show pos}: solve \{show k} as \{pprint [] !(quote 0 soln)}"
|
||||
pure $ Solved k soln
|
||||
(Solved k x) => error' "Meta \{show ix} already solved!"
|
||||
-- We're not breaking anything, but not quite getting an answer?
|
||||
for_ cons $ \case
|
||||
MkMc fc ctx sp rhs => do
|
||||
val <- vappSpine soln sp
|
||||
debug "discharge l=\{show ctx.lvl} \{(show val)} =?= \{(show rhs)}"
|
||||
-- is this the right depth?
|
||||
unify ctx.lvl val rhs
|
||||
|
||||
pure ()
|
||||
|
||||
trySolve : Nat -> Nat -> SnocList Val -> Val -> M ()
|
||||
trySolve l m sp t = do
|
||||
@@ -170,8 +218,6 @@ parameters (ctx: Context)
|
||||
pure ())
|
||||
|
||||
|
||||
export
|
||||
unify : (l : Nat) -> Val -> Val -> M UnifyResult
|
||||
|
||||
unifySpine : Nat -> Bool -> SnocList Val -> SnocList Val -> M UnifyResult
|
||||
unifySpine l False _ _ = error emptyFC "unify failed at head" -- unreachable now
|
||||
|
||||
@@ -79,7 +79,7 @@ eval env mode (App _ t u) = vapp !(eval env mode t) !(eval env mode u)
|
||||
eval env mode (U fc) = pure (VU fc)
|
||||
eval env mode (Meta fc i) =
|
||||
case !(lookupMeta i) of
|
||||
(Unsolved _ k xs _ _) => pure $ VMeta fc i [<]
|
||||
(Unsolved _ k xs _ _ _) => pure $ VMeta fc i [<]
|
||||
(Solved k t) => pure $ t
|
||||
eval env mode (Lam fc x t) = pure $ VLam fc x (MkClosure env t)
|
||||
eval env mode (Pi fc x icit a b) = pure $ VPi fc x icit !(eval env mode a) (MkClosure env b)
|
||||
@@ -112,7 +112,7 @@ quote l (VVar fc k sp) = if k < l
|
||||
else ?borken
|
||||
quote l (VMeta fc i sp) =
|
||||
case !(lookupMeta i) of
|
||||
(Unsolved _ k xs _ _) => quoteSp l (Meta fc i) sp
|
||||
(Unsolved _ k xs _ _ _) => quoteSp l (Meta fc i) sp
|
||||
(Solved k t) => quote l !(vappSpine t sp)
|
||||
quote l (VLam fc x t) = pure $ Lam fc x !(quote (S l) !(t $$ VVar emptyFC l [<]))
|
||||
quote l (VPi fc x icit a b) = pure $ Pi fc x icit !(quote l a) !(quote (S l) !(b $$ VVar emptyFC l [<]))
|
||||
@@ -128,35 +128,10 @@ export
|
||||
nf : Env -> Tm -> M Tm
|
||||
nf env t = quote (length env) !(eval env CBN t)
|
||||
|
||||
export
|
||||
prval : Val -> M String
|
||||
prval v = pure $ pprint [] !(quote 0 v)
|
||||
|
||||
export
|
||||
prvalCtx : {auto ctx : Context} -> Val -> M String
|
||||
prvalCtx v = pure $ pprint (toList $ map fst ctx.types) !(quote ctx.lvl v)
|
||||
|
||||
|
||||
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 pos k _ _ _) :: xs) lhs = if k == ix
|
||||
then do
|
||||
-- empty context should be ok, because this needs to be closed
|
||||
putStrLn "INFO at \{show pos}: solve \{show k} as \{!(prval tm)}"
|
||||
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)
|
||||
|
||||
|
||||
-- REVIEW - might be easier if we inserted the meta without a bunch of explicit App
|
||||
-- I believe Kovacs is doing that.
|
||||
|
||||
@@ -183,6 +158,7 @@ appSpine : Tm -> List Tm -> Tm
|
||||
appSpine t [] = t
|
||||
appSpine t (x :: xs) = appSpine (App (getFC t) t x) xs
|
||||
|
||||
-- TODO replace this with a variant on nf
|
||||
zonkApp : TopContext -> Nat -> Env -> Tm -> List Tm -> M Tm
|
||||
zonkApp top l env (App fc t u) sp = zonkApp top l env t (!(zonk top l env u) :: sp)
|
||||
zonkApp top l env t@(Meta fc k) sp = case !(lookupMeta k) of
|
||||
@@ -192,7 +168,7 @@ zonkApp top l env t@(Meta fc k) sp = case !(lookupMeta k) of
|
||||
foo <- vappSpine v ([<] <>< sp')
|
||||
debug "-> result is \{show foo}"
|
||||
quote l foo
|
||||
(Unsolved x j xs _ _) => pure $ appSpine t sp
|
||||
(Unsolved x j xs _ _ _) => pure $ appSpine t sp
|
||||
zonkApp top l env t sp = pure $ appSpine !(zonk top l env t) sp
|
||||
|
||||
zonkAlt : TopContext -> Nat -> Env -> CaseAlt -> M CaseAlt
|
||||
|
||||
@@ -25,7 +25,7 @@ isCandidate _ _ = False
|
||||
|
||||
-- This is a crude first pass
|
||||
-- TODO consider ctx
|
||||
findMatches : Val -> List TopEntry -> M (List Tm)
|
||||
findMatches : Val -> List TopEntry -> M (List (Tm, MetaContext))
|
||||
findMatches ty [] = pure []
|
||||
findMatches ty ((MkEntry name type def@(Fn t)) :: xs) = do
|
||||
let True = isCandidate ty type | False => findMatches ty xs
|
||||
@@ -38,10 +38,12 @@ findMatches ty ((MkEntry name type def@(Fn t)) :: xs) = do
|
||||
-- TODO sort out the FC here
|
||||
let fc = getFC ty
|
||||
debug "TRY \{name} : \{pprint [] type} for \{show ty}"
|
||||
tm <- check (mkCtx top.metas fc) (RVar fc name) ty
|
||||
-- This check is solving metas, so we save mc below in case we want this solution
|
||||
tm <- check (mkCtx top.metas fc) (RVar fc name) ty
|
||||
debug "Found \{pprint [] tm} for \{show ty}"
|
||||
mc' <- readIORef top.metas
|
||||
writeIORef top.metas mc
|
||||
(tm ::) <$> findMatches ty xs)
|
||||
((tm, mc') ::) <$> findMatches ty xs)
|
||||
(\ err => do
|
||||
debug "No match \{show ty} \{pprint [] type} \{showError "" err}"
|
||||
writeIORef top.metas mc
|
||||
@@ -127,14 +129,16 @@ processDecl (Def fc nm clauses) = do
|
||||
|
||||
mc <- readIORef top.metas
|
||||
let mlen = length mc.metas `minus` mstart
|
||||
-- FIXME every time we hit solve in this loop, we need to start over with mc.metas
|
||||
for_ (take mlen mc.metas) $ \case
|
||||
(Unsolved fc k ctx ty AutoSolve) => do
|
||||
debug "auto solving \{show k} : \{show ty}"
|
||||
(Unsolved fc k ctx ty AutoSolve _) => do
|
||||
debug "AUTO solving \{show k} : \{show ty}"
|
||||
-- we want the context here too.
|
||||
[tm] <- findMatches ty top.defs
|
||||
| res => error fc "Failed to solve \{show ty}, matches: \{show $ map (pprint []) res}"
|
||||
[(tm,mc)] <- findMatches ty top.defs
|
||||
| res => error fc "FAILED to solve \{show ty}, matches: \{show $ map (pprint [] . fst) res}"
|
||||
writeIORef top.metas mc
|
||||
val <- eval ctx.env CBN tm
|
||||
debug "solution \{pprint [] tm} evaled to \{show val}"
|
||||
debug "SOLUTION \{pprint [] tm} evaled to \{show val}"
|
||||
let sp = makeSpine ctx.lvl ctx.bds
|
||||
solve ctx ctx.lvl k sp val
|
||||
pure ()
|
||||
@@ -146,13 +150,15 @@ processDecl (Def fc nm clauses) = do
|
||||
mc <- readIORef top.metas
|
||||
for_ (take mlen mc.metas) $ \case
|
||||
(Solved k x) => pure ()
|
||||
(Unsolved (l,c) k ctx ty User) => do
|
||||
(Unsolved (l,c) k ctx ty User cons) => do
|
||||
-- TODO print here instead of during Elab
|
||||
pure ()
|
||||
(Unsolved (l,c) k ctx ty kind) => do
|
||||
(Unsolved (l,c) k ctx ty kind cons) => do
|
||||
tm <- quote ctx.lvl !(forceMeta ty)
|
||||
-- Now that we're collecting errors, maybe we simply check at the end
|
||||
addError $ E (l,c) "Unsolved meta \{show k} \{show kind} type \{pprint (names ctx) tm}"
|
||||
-- TODO - log constraints?
|
||||
addError $ E (l,c) "Unsolved meta \{show k} \{show kind} type \{pprint (names ctx) tm} \{show $ length cons} constraints"
|
||||
|
||||
debug "Add def \{nm} \{pprint [] tm'} : \{pprint [] ty}"
|
||||
modify $ setDef nm ty (Fn tm')
|
||||
|
||||
|
||||
@@ -325,8 +325,12 @@ Show MetaKind where
|
||||
show User = "User"
|
||||
show AutoSolve = "Auto"
|
||||
|
||||
-- constrain meta applied to val to be a val
|
||||
public export
|
||||
data MetaEntry = Unsolved FC Nat Context Val MetaKind | Solved Nat Val
|
||||
data MConstraint = MkMc FC Context (SnocList Val) Val
|
||||
|
||||
public export
|
||||
data MetaEntry = Unsolved FC Nat Context Val MetaKind (List MConstraint) | Solved Nat Val
|
||||
|
||||
|
||||
public export
|
||||
@@ -399,6 +403,7 @@ record Context where
|
||||
metas : IORef MetaContext
|
||||
fc : FC
|
||||
|
||||
%name Context ctx
|
||||
|
||||
||| add a binding to environment
|
||||
export
|
||||
@@ -416,7 +421,7 @@ define ctx name val ty =
|
||||
export
|
||||
covering
|
||||
Show MetaEntry where
|
||||
show (Unsolved pos k ctx ty kind) = "Unsolved \{show pos} \{show k} \{show kind} : \{show ty} \{show ctx.bds}"
|
||||
show (Unsolved pos k ctx ty kind constraints) = "Unsolved \{show pos} \{show k} \{show kind} : \{show ty} \{show ctx.bds} cs \{show $ length constraints}"
|
||||
show (Solved k x) = "Solved \{show k} \{show x}"
|
||||
|
||||
export withPos : Context -> FC -> Context
|
||||
@@ -447,7 +452,7 @@ freshMeta : Context -> FC -> Val -> MetaKind -> M Tm
|
||||
freshMeta ctx fc ty kind = do
|
||||
mc <- readIORef ctx.metas
|
||||
putStrLn "INFO at \{show fc}: fresh meta \{show mc.next} : \{show ty}"
|
||||
writeIORef ctx.metas $ { next $= S, metas $= (Unsolved fc mc.next ctx ty kind ::) } mc
|
||||
writeIORef ctx.metas $ { next $= S, metas $= (Unsolved fc mc.next ctx ty kind [] ::) } mc
|
||||
pure $ applyBDs 0 (Meta emptyFC mc.next) ctx.bds
|
||||
where
|
||||
-- hope I got the right order here :)
|
||||
@@ -471,7 +476,7 @@ lookupMeta ix = do
|
||||
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@(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
|
||||
|
||||
-- we need more of topcontext later - Maybe switch it up so we're not passing
|
||||
|
||||
Reference in New Issue
Block a user