First attempt at expression search

Not quite there yet because we're not saving all the information about
holes in TTCs but the basics work
This commit is contained in:
Edwin Brady 2019-06-01 22:34:59 +01:00
parent ae28f1b1f2
commit bf70aa07d2
7 changed files with 513 additions and 1 deletions

View File

@ -20,6 +20,7 @@ searchType : {auto c : Ref Ctxt Defs} ->
(defining : Name) -> (topTy : ClosedTerm) ->
Env Term vars -> (target : Term vars) -> Core (Term vars)
public export
record ArgInfo (vars : List Name) where
constructor MkArgInfo
holeID : Int
@ -28,6 +29,7 @@ record ArgInfo (vars : List Name) where
metaApp : Term vars
argType : Term vars
export
mkArgs : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} ->
FC -> RigCount ->

View File

@ -668,7 +668,8 @@ mutual
-- If they're both holes, solve the one with the bigger context
doUnifyBothApps mode loc env xfc (NMeta xn xi xargs) xargs' yfc (NMeta yn yi yargs) yargs'
= do invx <- isDefInvertible xi
if xi == yi && invx -- Invertible, (from auto implicit search)
if xi == yi && (invx || mode == InSearch)
-- Invertible, (from auto implicit search)
-- so we can also unify the arguments.
then unifyArgs mode loc env (map snd (xargs ++ xargs'))
(map snd (yargs ++ yargs'))
@ -681,6 +682,10 @@ mutual
pv : Name -> Bool
pv (PV _ _) = True
pv _ = False
doUnifyBothApps InSearch loc env xfc fx@(NRef xt hdx) xargs yfc fy@(NRef yt hdy) yargs
= if hdx == hdy
then unifyArgs InSearch loc env (map snd xargs) (map snd yargs)
else unifyApp False InSearch loc env xfc fx xargs (NApp yfc fy yargs)
doUnifyBothApps mode loc env xfc fx ax yfc fy ay
= unifyApp False mode loc env xfc fx ax (NApp yfc fy ay)

View File

@ -0,0 +1,489 @@
module TTImp.Interactive.ExprSearch
-- Expression search for interactive editing. There's a lot of similarities
-- with Core.AutoSearch (and we reuse some of it) but I think it's better for
-- them to be entirely separate: AutoSearch is a crucial part of the core
-- therefore it's good for it to be well defined and cleanly separated from
-- everything else, whereas this search mechanism is about finding plausible
-- candidates for programs.
-- We try to find as many results as possible, within the given search
-- depth.
import Core.AutoSearch
import Core.CaseTree
import Core.Context
import Core.Env
import Core.LinearCheck
import Core.Metadata
import Core.Normalise
import Core.Unify
import Core.TT
import Core.Value
import TTImp.Interactive.CaseSplit
import TTImp.Utils
%default covering
-- Data for building recursive calls - the function name, and the structure
-- of the LHS. Only recursive calls with a different structure are okay.
record RecData where
constructor MkRecData
recname : Name -- resolved name
lhsapp : Term vars
record SearchOpts where
constructor MkSearchOpts
holesOK : Bool
recOK : Bool
depth : Nat
search : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount ->
SearchOpts -> Maybe RecData ->
ClosedTerm ->
Name -> Core (List ClosedTerm)
getAllEnv : FC -> (done : List Name) ->
Env Term vars ->
List (Term (done ++ vars), Term (done ++ vars))
getAllEnv fc done [] = []
getAllEnv {vars = v :: vs} fc done (b :: env)
= let rest = getAllEnv fc (done ++ [v]) env
MkVar p = weakenVar done {inner = v :: vs} First in
(Local fc Nothing _ p,
rewrite appendAssociative done [v] vs in
weakenNs (done ++ [v]) (binderType b)) ::
rewrite appendAssociative done [v] vs in rest
-- nameIsHole : {auto c : Ref Ctxt Defs} ->
-- {auto m : Ref MD Metadata} ->
-- FC -> Name -> Core Bool
-- nameIsHole fc n
-- = do defs <- get Ctxt
-- case !(lookupDefExact n (gamma defs)) of
-- Nothing => throw (InternalError "Can't happen, name has mysteriously vanished")
-- Just def =>
-- case def of
-- Hole locs False _ => pure True
-- _ => pure False
-- Search recursively, but only if the given name wasn't solved by unification
searchIfHole : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> SearchOpts -> Maybe RecData -> ClosedTerm ->
Env Term vars -> ArgInfo vars ->
Core (List (AppInfo, Term vars))
searchIfHole fc opts defining topty env arg
= case depth opts of
Z => pure []
S k =>
do let hole = holeID arg
let rig = argRig arg
defs <- get Ctxt
Just gdef <- lookupCtxtExact (Resolved hole) (gamma defs)
| Nothing => pure []
let Hole inv = definition gdef
| _ => pure [(appInf arg,
!(normaliseHoles defs env (metaApp arg)))]
-- already solved
tms <- search fc rig (record { depth = k} opts)
defining topty (Resolved hole)
traverse (\tm => pure (appInf arg,
!(normaliseHoles defs env
(applyTo fc (embed tm) env)))) tms
mkCandidates : FC -> Term vars -> List (List (AppInfo, Term vars)) ->
List (Term vars)
mkCandidates fc f [] = pure f
mkCandidates fc f (args :: argss)
= do (p, arg) <- args
mkCandidates fc (App fc f p arg) argss
explicit : ArgInfo vars -> Bool
explicit ai = plicit (appInf ai) == Explicit
-- Apply the name to arguments and see if the result unifies with the target
-- type, then try to automatically solve any holes which were generated.
-- If there are any remaining holes, search fails.
searchName : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts ->
Env Term vars -> NF vars -> ClosedTerm ->
Maybe RecData -> (Name, GlobalDef) -> Core (List (Term vars))
searchName fc rigc opts env target topty defining (n, ndef)
= do defs <- get Ctxt
let ty = type ndef
let namety : NameType
= case definition ndef of
DCon tag arity => DataCon tag arity
TCon tag arity _ _ _ _ => TyCon tag arity
_ => Func
log 5 $ "Trying " ++ show (fullname ndef)
nty <- nf defs env (embed ty)
(args, appTy) <- mkArgs fc rigc env nty
logNF 5 "Target" env target
logNF 10 "App type" env appTy
ures <- unify InSearch fc env target appTy
let [] = constraints ures
| _ => pure []
-- Search the explicit arguments first
args' <- traverse (searchIfHole fc opts defining topty env)
(filter explicit args)
args' <- traverse (searchIfHole fc opts defining topty env)
args
let cs = mkCandidates fc (Ref fc namety n) args'
logC 10 (do strs <- traverse (\t => pure (show !(toFullNames t) ++ "\n")) cs
pure ("Candidates: " ++ concat strs))
pure cs
successful : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
List (Core a) ->
Core (List (Either Error a))
successful [] = pure []
successful (elab :: elabs)
-- Don't save any state, we'll happily keep adding variables and unification
-- problems because we might be creating multiple candidate solutions, and
-- they won't interfere.
-- We will go back to the original state once we're done!
= catch (do res <- elab
rest <- successful elabs
pure (Right res :: rest))
(\err =>
do rest <- successful elabs
pure (Left err :: rest))
getSuccessful : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts -> Bool ->
Env Term vars -> Term vars -> ClosedTerm ->
Maybe RecData ->
List (Core (List (Term vars))) ->
Core (List (Term vars))
getSuccessful {vars} fc rig opts mkHole env ty topty defining all
= do elabs <- successful all
case concat (rights elabs) of
[] => -- If no successful search, make a hole
if mkHole && holesOK opts
then do defs <- get Ctxt
let base = maybe "arg"
(\r => nameRoot (recname r) ++ "_rhs")
defining
hn <- uniqueName defs (map nameRoot vars) base
(idx, tm) <- newMeta fc rig env (UN hn) ty False
pure [tm]
else if holesOK opts
then pure []
else throw (CantSolveGoal fc [] topty)
rs => pure rs
searchNames : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts -> Env Term vars ->
Term vars -> ClosedTerm ->
Maybe RecData -> List Name -> Core (List (Term vars))
searchNames fc rig opts env ty topty defining []
= pure []
searchNames fc rig opts env ty topty defining (n :: ns)
= do defs <- get Ctxt
vis <- traverse (visible (gamma defs) (currentNS defs)) (n :: ns)
let visns = mapMaybe id vis
nfty <- nf defs env ty
logTerm 10 ("Searching " ++ show (map fst visns) ++ " for ") ty
getSuccessful fc rig opts False env ty topty defining
(map (searchName fc rig opts env nfty topty defining) visns)
where
visible : Context GlobalDef -> List String -> Name ->
Core (Maybe (Name, GlobalDef))
visible gam nspace n
= do Just def <- lookupCtxtExact n gam
| Nothing => pure Nothing
if visibleIn nspace n (visibility def)
then pure (Just (n, def))
else pure Nothing
tryRecursive : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts ->
Env Term vars -> Term vars -> ClosedTerm ->
Maybe RecData -> Core (List (Term vars))
tryRecursive fc rig opts env ty topty Nothing = pure []
tryRecursive fc rig opts env ty topty (Just rdata)
= do defs <- get Ctxt
case !(lookupCtxtExact (recname rdata) (gamma defs)) of
Nothing => pure []
Just def =>
do tms <- searchName fc rig opts env !(nf defs env ty)
topty Nothing
(recname rdata, def)
defs <- get Ctxt
pure (filter (structDiff (lhsapp rdata)) tms)
where
mutual
-- A fairly simple superficialy syntactic check to make sure that
-- the recursive call is structurally different from the lhs
-- (Essentially, meaning that there's a constructor application in
-- one where there's a local in another, or that constructor applications
-- differ somewhere)
argDiff : Term vs -> Term vs' -> Bool
argDiff (Local _ _ _ _) (Local _ _ _ _) = False
argDiff (Ref _ _ fn) (Ref _ _ fn') = fn /= fn'
argDiff (Bind _ _ _ _) (Bind _ _ _ _) = False
argDiff (App _ f _ a) (App _ f' _ a')
= structDiff f f' || structDiff a a'
argDiff (PrimVal _ c) (PrimVal _ c') = c /= c'
argDiff (Erased _) _ = False
argDiff _ (Erased _) = False
argDiff (TType _) (TType _) = False
argDiff _ _ = True
appsDiff : Term vs -> Term vs' -> List (Term vs) -> List (Term vs') ->
Bool
appsDiff (Ref _ (DataCon _ _) f) (Ref _ (DataCon _ _) f') args args'
= if f == f' then or (map Delay (zipWith argDiff args args'))
else True
appsDiff (Ref _ (TyCon _ _) f) (Ref _ (TyCon _ _) f') args args'
= if f == f' then or (map Delay (zipWith argDiff args args'))
else True
appsDiff (Ref _ _ f) (Ref _ _ f') args args'
= if f == f' && length args == length args'
then or (map Delay (zipWith argDiff args args'))
else False -- can't be sure
appsDiff (Ref _ (DataCon _ _) f) (Local _ _ _ _) _ _ = True
appsDiff (Local _ _ _ _) (Ref _ (DataCon _ _) f) _ _ = True
appsDiff f f' [] [] = argDiff f f'
appsDiff _ _ _ _ = False -- can't be sure...
-- Reject if the recursive call is the same in structure as the lhs
structDiff : Term vs -> Term vs' -> Bool
structDiff tm tm'
= let (f, args) = getFnArgs tm
(f', args') = getFnArgs tm' in
appsDiff f f' (map snd args) (map snd args')
-- A local is usable as long as its type isn't a hole
usableLocal : FC -> Env Term vars -> NF vars -> Bool
usableLocal loc env (NApp _ (NMeta _ _ _) args) = False
usableLocal loc _ _ = True
searchLocalWith : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts -> Env Term vars ->
List (Term vars, Term vars) ->
Term vars -> ClosedTerm -> Maybe RecData ->
Core (List (Term vars))
searchLocalWith fc rig opts env [] ty topty defining
= pure []
searchLocalWith {vars} fc rig opts env ((p, pty) :: rest) ty topty defining
= do defs <- get Ctxt
nty <- nf defs env ty
getSuccessful fc rig opts False env ty topty defining
[findPos defs p id !(nf defs env pty) nty,
searchLocalWith fc rig opts env rest ty
topty defining]
where
findDirect : Defs -> Term vars ->
(Term vars -> Term vars) ->
NF vars -> -- local variable's type
NF vars -> -- type we're looking for
Core (List (Term vars))
findDirect defs prf f nty ty
= do (args, appTy) <- mkArgs fc rig env nty
-- We can only use a local variable if it's not an unsolved
-- hole
if usableLocal fc env nty
then
tryUnify -- try with no arguments first
(do ures <- unify InTerm fc env ty nty
let [] = constraints ures
| _ => throw (InternalError "Can't use directly")
pure (mkCandidates fc (f prf) []))
(do ures <- unify InTerm fc env ty appTy
let [] = constraints ures
| _ => pure []
args' <- traverse (searchIfHole fc opts defining topty env)
args
pure (mkCandidates fc (f prf) args'))
else pure []
findPos : Defs -> Term vars ->
(Term vars -> Term vars) ->
NF vars -> NF vars -> Core (List (Term vars))
findPos defs prf f x@(NTCon pfc pn _ _ [(xp, xty), (yp, yty)]) target
= getSuccessful fc rig opts False env ty topty defining
[findDirect defs prf f x target,
(do fname <- maybe (throw (InternalError "No fst"))
pure
(fstName defs)
sname <- maybe (throw (InternalError "No snd"))
pure
(sndName defs)
if isPairType pn defs
then do empty <- clearDefs defs
xtytm <- quote empty env xty
ytytm <- quote empty env yty
getSuccessful fc rig opts False env ty topty defining
[(do xtynf <- evalClosure defs xty
findPos defs prf
(\arg => applyInfo fc (Ref fc Func fname)
[(xp, xtytm),
(yp, ytytm),
(explApp Nothing, f arg)])
xtynf target),
(do ytynf <- evalClosure defs yty
findPos defs prf
(\arg => applyInfo fc (Ref fc Func sname)
[(xp, xtytm),
(yp, ytytm),
(explApp Nothing, f arg)])
ytynf target)]
else pure [])]
findPos defs prf f nty target = findDirect defs prf f nty target
searchLocal : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts ->
Env Term vars -> Term vars -> ClosedTerm ->
Maybe RecData -> Core (List (Term vars))
searchLocal fc rig opts env ty topty defining
= do defs <- get Ctxt
searchLocalWith fc rig opts env (getAllEnv fc [] env)
ty topty defining
searchType : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts -> Env Term vars -> Maybe RecData ->
ClosedTerm ->
Nat -> Term vars -> Core (List (Term vars))
searchType fc rig opts env defining topty k (Bind bfc n (Pi c info ty) sc)
= do let env' : Env Term (n :: _) = Pi c info ty :: env
log 10 $ "Introduced lambda, search for " ++ show sc
scVal <- searchType fc rig opts env' defining topty k sc
pure (map (Bind bfc n (Lam c info ty)) scVal)
searchType {vars} fc rig opts env defining topty Z (Bind bfc n (Pi c info ty) sc)
= -- try a local before creating a lambda...
tryUnify
(searchLocal fc rig opts env (Bind bfc n (Pi c info ty) sc) topty defining)
(do defs <- get Ctxt
let n' = UN (getArgName defs n vars !(nf defs env ty))
let env' : Env Term (n' :: _) = Pi c info ty :: env
let sc' = renameTop n' sc
log 10 $ "Introduced lambda, search for " ++ show sc'
scVal <- searchType fc rig opts env' defining topty Z sc'
pure (map (Bind bfc n' (Lam c info ty)) scVal))
searchType fc rig opts env defining topty _ ty
= case getFnArgs ty of
(Ref rfc (TyCon t ar) n, args) =>
do defs <- get Ctxt
if length args == ar
then do sd <- getSearchData fc False n
let allHints = concat (hintGroups sd)
-- Solutions is either:
-- First try the locals,
-- Then try the hints in order
-- Then try a recursive call
log 10 $ "Hints found for " ++ show n ++ " "
++ show allHints
getSuccessful fc rig opts True env ty topty defining
([searchLocal fc rig opts env ty topty defining,
searchNames fc rig opts env ty topty defining
allHints]
++ if recOK opts
then [tryRecursive fc rig opts env ty topty defining]
else [])
else pure []
_ => do logTerm 10 "Searching locals only at" ty
getSuccessful fc rig opts True env ty topty defining
([searchLocal fc rig opts env ty topty defining]
++ if recOK opts
then [tryRecursive fc rig opts env ty topty defining]
else [])
searchHole : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> RigCount -> SearchOpts -> Maybe RecData -> Name ->
Nat -> ClosedTerm ->
Defs -> GlobalDef -> Core (List ClosedTerm)
searchHole fc rig opts defining n locs topty defs glob
= do searchty <- normaliseScope defs [] (type glob)
logTerm 10 "Normalised type" searchty
searchType fc rig opts [] defining topty locs searchty
-- Declared at the top
search fc rig opts defining topty n_in
= do defs <- get Ctxt
case !(lookupHoleName n_in (gamma defs)) of
Just (n, i, gdef) =>
-- The definition should be 'BySearch' at this stage,
-- if it's arising from an auto implicit
case definition gdef of
Hole _ => searchHole fc rig opts defining n Z topty defs gdef
BySearch _ _ _ => searchHole fc rig opts defining n
!(getArity defs [] (type gdef))
topty defs gdef
_ => do log 10 $ show n_in ++ " not a hole"
throw (InternalError $ "Not a hole: " ++ show n ++ " in " ++ show (map recname defining))
_ => do log 10 $ show n_in ++ " not found"
throw (UndefinedName fc n_in)
where
lookupHoleName : Name -> Context GlobalDef ->
Core (Maybe (Name, Int, GlobalDef))
lookupHoleName n defs
= case !(lookupCtxtExactI n defs) of
Just (i, res) => pure $ Just (n, i, res)
Nothing => case !(lookupCtxtName n defs) of
[res] => pure $ Just res
_ => pure Nothing
getLHSData : Defs -> Maybe ClosedTerm -> Core (Maybe RecData)
getLHSData defs Nothing = pure Nothing
getLHSData defs (Just tm) = pure $ getLHS !(normaliseHoles defs [] tm)
where
getLHS : Term vars -> Maybe RecData
getLHS (Bind _ _ (PVar _ _) sc) = getLHS sc
getLHS (Bind _ _ (PLet _ _ _) sc) = getLHS sc
getLHS sc
= case getFn sc of
Ref _ _ n => Just (MkRecData n sc)
_ => Nothing
dropLinearErrors : {auto c : Ref Ctxt Defs} ->
{auto u : Ref UST UState} ->
FC -> List ClosedTerm ->
Core (List ClosedTerm)
dropLinearErrors fc [] = pure []
dropLinearErrors fc (t :: ts)
= tryUnify
(do linearCheck fc Rig1 False [] t
ts' <- dropLinearErrors fc ts
pure (t :: ts'))
(dropLinearErrors fc ts)
export
exprSearch : {auto c : Ref Ctxt Defs} ->
{auto m : Ref MD Metadata} ->
{auto u : Ref UST UState} ->
FC -> Name -> List Name -> Core (List ClosedTerm)
exprSearch fc n hints
= do lhs <- findHoleLHS n
defs <- get Ctxt
Just gdef <- lookupCtxtExact n (gamma defs)
| Nothing => throw (UndefinedName fc n)
rs <- search fc (multiplicity gdef) (MkSearchOpts False True 5)
!(getLHSData defs lhs) (type gdef) n
dropLinearErrors fc rs

View File

@ -691,6 +691,9 @@ command
<|> do symbol ":"; exactIdent "s"
n <- name
pure (ProofSearch n)
<|> do symbol ":"; exactIdent "es"
n <- name
pure (ExprSearch n)
<|> do symbol ":"; exactIdent "di"
n <- name
pure (DebugInfo n)

View File

@ -298,6 +298,7 @@ data ImpREPL : Type where
Eval : RawImp -> ImpREPL
Check : RawImp -> ImpREPL
ProofSearch : Name -> ImpREPL
ExprSearch : Name -> ImpREPL
DebugInfo : Name -> ImpREPL
Quit : ImpREPL

View File

@ -13,6 +13,7 @@ import Core.Value
import TTImp.Elab
import TTImp.Elab.Check
import TTImp.Interactive.ExprSearch
import TTImp.Parser
import TTImp.ProcessDecls
import TTImp.TTImp
@ -67,6 +68,16 @@ process (ProofSearch n_in)
defnf <- normaliseHoles defs [] def
coreLift (printLn !(toFullNames defnf))
pure True
process (ExprSearch n_in)
= do defs <- get Ctxt
[(n, i, ty)] <- lookupTyName n_in (gamma defs)
| [] => throw (UndefinedName toplevelFC n_in)
| ns => throw (AmbiguousName toplevelFC (map fst ns))
results <- exprSearch toplevelFC n []
defs <- get Ctxt
defnfs <- traverse (normaliseHoles defs []) results
traverse_ (\d => coreLift (printLn !(toFullNames d))) defnfs
pure True
process (DebugInfo n)
= do defs <- get Ctxt
traverse showInfo !(lookupDefName n (gamma defs))

View File

@ -72,6 +72,7 @@ modules =
TTImp.Elab.Rewrite,
TTImp.Elab.Term,
TTImp.Interactive.CaseSplit,
TTImp.Interactive.ExprSearch,
TTImp.Parser,
TTImp.ProcessData,
TTImp.ProcessDecls,