mirror of
https://github.com/edwinb/Idris2-boot.git
synced 2024-12-26 14:21:52 +03:00
Initial implementation of auto implicit search
This only searches for lambdas and locals so far, and doesn't take into account that it shouldn't search for holes, but it does set up the basic mechanism.
This commit is contained in:
parent
190690e925
commit
29a41ed975
118
src/Core/AutoSearch.idr
Normal file
118
src/Core/AutoSearch.idr
Normal file
@ -0,0 +1,118 @@
|
|||||||
|
module Core.AutoSearch
|
||||||
|
|
||||||
|
import Core.Context
|
||||||
|
import Core.Core
|
||||||
|
import Core.Env
|
||||||
|
import Core.Normalise
|
||||||
|
import Core.TT
|
||||||
|
import Core.Unify
|
||||||
|
import Core.Value
|
||||||
|
|
||||||
|
%default covering
|
||||||
|
|
||||||
|
SearchEnv : List Name -> Type
|
||||||
|
SearchEnv vars = List (NF vars, Closure vars)
|
||||||
|
|
||||||
|
searchType : {auto c : Ref Ctxt Defs} ->
|
||||||
|
{auto u : Ref UST UState} ->
|
||||||
|
FC -> RigCount ->
|
||||||
|
(defining : Name) -> (topTy : Term vars) ->
|
||||||
|
Env Term vars -> SearchEnv vars ->
|
||||||
|
(target : NF vars) -> Core (NF vars)
|
||||||
|
|
||||||
|
export
|
||||||
|
search : {auto c : Ref Ctxt Defs} ->
|
||||||
|
{auto u : Ref UST UState} ->
|
||||||
|
FC -> RigCount ->
|
||||||
|
(defining : Name) -> (topTy : Term vars) -> Env Term vars ->
|
||||||
|
Core (Term vars)
|
||||||
|
|
||||||
|
record ArgInfo (vars : List Name) where
|
||||||
|
constructor MkArgInfo
|
||||||
|
holeID : Int
|
||||||
|
argRig : RigCount
|
||||||
|
appInf : AppInfo
|
||||||
|
metaApp : Term vars
|
||||||
|
argType : NF vars
|
||||||
|
|
||||||
|
mkArgs : {auto c : Ref Ctxt Defs} ->
|
||||||
|
{auto u : Ref UST UState} ->
|
||||||
|
FC -> RigCount ->
|
||||||
|
Env Term vars -> NF vars ->
|
||||||
|
Core (List (ArgInfo vars), NF vars)
|
||||||
|
mkArgs fc rigc env (NBind nfc x (Pi c p ty) sc)
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
empty <- clearDefs defs
|
||||||
|
nm <- genName "sa"
|
||||||
|
argTy <- quote empty env ty
|
||||||
|
let argRig = rigMult rigc c
|
||||||
|
(idx, arg) <- newMeta fc argRig env nm argTy
|
||||||
|
-- setInvertible fc argName
|
||||||
|
(rest, restTy) <- mkArgs fc rigc env
|
||||||
|
!(sc (toClosure defaultOpts env arg))
|
||||||
|
pure (MkArgInfo idx argRig (appInf Nothing p) arg ty :: rest, restTy)
|
||||||
|
mkArgs fc rigc env ty = pure ([], ty)
|
||||||
|
|
||||||
|
closureApply : FC -> Closure vars -> List (ArgInfo vars) ->
|
||||||
|
Core (Term vars)
|
||||||
|
closureApply fc (MkClosure _ [] _ tm) args
|
||||||
|
= pure (applyInfo fc tm (map (\i => (appInf i, metaApp i)) args))
|
||||||
|
closureApply _ _ _ = throw (InternalError "Wrong argument form in AutoSearch")
|
||||||
|
|
||||||
|
searchIfHole : {auto c : Ref Ctxt Defs} ->
|
||||||
|
{auto u : Ref UST UState} ->
|
||||||
|
FC ->
|
||||||
|
(defining : Name) -> (topTy : Term vars) -> Env Term vars ->
|
||||||
|
SearchEnv vars ->
|
||||||
|
(arg : ArgInfo vars) ->
|
||||||
|
Core ()
|
||||||
|
searchIfHole fc def top env locs arg
|
||||||
|
= do let hole = holeID arg
|
||||||
|
let rig = argRig arg
|
||||||
|
|
||||||
|
defs <- get Ctxt
|
||||||
|
Just gdef <- lookupCtxtExact (Resolved hole) (gamma defs)
|
||||||
|
| Nothing => throw (CantSolveGoal fc env top)
|
||||||
|
let Hole inv = definition gdef
|
||||||
|
| _ => pure () -- already solved
|
||||||
|
argdef <- searchType fc rig def top env locs (argType arg)
|
||||||
|
vs <- unify InTerm fc env !(nf defs env (metaApp arg)) argdef
|
||||||
|
let [] = constraints vs
|
||||||
|
| _ => throw (CantSolveGoal fc env top)
|
||||||
|
pure ()
|
||||||
|
|
||||||
|
searchLocal : {auto c : Ref Ctxt Defs} ->
|
||||||
|
{auto u : Ref UST UState} ->
|
||||||
|
FC -> RigCount ->
|
||||||
|
(defining : Name) -> (topTy : Term vars) ->
|
||||||
|
Env Term vars -> SearchEnv vars -> SearchEnv vars ->
|
||||||
|
(target : NF vars) -> Core (NF vars)
|
||||||
|
searchLocal fc rigc def top env locs [] target
|
||||||
|
= throw (CantSolveGoal fc env top)
|
||||||
|
searchLocal fc rigc def top env locs ((ty, c) :: rest) target
|
||||||
|
= tryUnify
|
||||||
|
(do (args, appTy) <- mkArgs fc rigc env ty
|
||||||
|
-- TODO: Can only use the local if it's not an unsolved hole
|
||||||
|
ures <- unify InTerm fc env target appTy
|
||||||
|
let [] = constraints ures
|
||||||
|
| _ => throw (CantSolveGoal fc env top)
|
||||||
|
candidate <- closureApply fc c args
|
||||||
|
traverse (searchIfHole fc def top env locs) args
|
||||||
|
defs <- get Ctxt
|
||||||
|
nf defs env candidate)
|
||||||
|
(searchLocal fc rigc def top env locs rest target)
|
||||||
|
|
||||||
|
searchType fc rigc def top env locs (NBind nfc x (Pi c p ty) sc)
|
||||||
|
= pure (NBind nfc x (Lam c p ty)
|
||||||
|
(\arg => searchType fc rigc def top env ((ty, arg) :: locs) !(sc arg)))
|
||||||
|
searchType fc rigc def top env locs target
|
||||||
|
= searchLocal fc rigc def top env locs locs target
|
||||||
|
|
||||||
|
-- Declared at the top
|
||||||
|
search fc rigc def top env
|
||||||
|
= do defs <- get Ctxt
|
||||||
|
tm <- searchType fc rigc def top env []
|
||||||
|
!(nf defs env top)
|
||||||
|
empty <- clearDefs defs
|
||||||
|
quote empty env tm
|
||||||
|
|
@ -498,6 +498,10 @@ export
|
|||||||
lookupTyExact : Name -> Context GlobalDef -> Core (Maybe ClosedTerm)
|
lookupTyExact : Name -> Context GlobalDef -> Core (Maybe ClosedTerm)
|
||||||
lookupTyExact = lookupExactBy type
|
lookupTyExact = lookupExactBy type
|
||||||
|
|
||||||
|
export
|
||||||
|
lookupTyName : Name -> Context GlobalDef -> Core (List (Name, Int, ClosedTerm))
|
||||||
|
lookupTyName = lookupNameBy type
|
||||||
|
|
||||||
-- Set the default namespace for new definitions
|
-- Set the default namespace for new definitions
|
||||||
export
|
export
|
||||||
setNS : {auto c : Ref Ctxt Defs} ->
|
setNS : {auto c : Ref Ctxt Defs} ->
|
||||||
|
@ -52,7 +52,7 @@ data Error
|
|||||||
| IncompatibleFieldUpdate FC (List String)
|
| IncompatibleFieldUpdate FC (List String)
|
||||||
| InvalidImplicits FC (Env Term vars) (List (Maybe Name)) (Term vars)
|
| InvalidImplicits FC (Env Term vars) (List (Maybe Name)) (Term vars)
|
||||||
| TryWithImplicits FC (Env Term vars) (List (Name, Term vars))
|
| TryWithImplicits FC (Env Term vars) (List (Name, Term vars))
|
||||||
| CantSolveGoal FC (Term [])
|
| CantSolveGoal FC (Env Term vars) (Term vars)
|
||||||
| DeterminingArg FC Name (Env Term vars) (Term vars)
|
| DeterminingArg FC Name (Env Term vars) (Term vars)
|
||||||
| UnsolvedHoles (List (FC, Name))
|
| UnsolvedHoles (List (FC, Name))
|
||||||
| CantInferArgType FC (Env Term vars) Name Name (Term vars)
|
| CantInferArgType FC (Env Term vars) Name Name (Term vars)
|
||||||
@ -171,7 +171,8 @@ Show Error where
|
|||||||
= show fc ++ ":Need to bind implicits "
|
= show fc ++ ":Need to bind implicits "
|
||||||
++ showSep "," (map (\x => show (fst x) ++ " : " ++ show (snd x)) imps)
|
++ showSep "," (map (\x => show (fst x) ++ " : " ++ show (snd x)) imps)
|
||||||
++ "\n(The front end should probably have done this for you. Please report!)"
|
++ "\n(The front end should probably have done this for you. Please report!)"
|
||||||
show (CantSolveGoal fc g) = show fc ++ ":Can't solve goal " ++ assert_total (show g)
|
show (CantSolveGoal fc env g)
|
||||||
|
= show fc ++ ":Can't solve goal " ++ assert_total (show g)
|
||||||
show (DeterminingArg fc n env g)
|
show (DeterminingArg fc n env g)
|
||||||
= show fc ++ ":Can't solve goal " ++ assert_total (show g) ++
|
= show fc ++ ":Can't solve goal " ++ assert_total (show g) ++
|
||||||
" since argument " ++ show n ++ " can't be inferred"
|
" since argument " ++ show n ++ " can't be inferred"
|
||||||
@ -263,7 +264,7 @@ getErrorLoc (NotRecordType loc _) = Just loc
|
|||||||
getErrorLoc (IncompatibleFieldUpdate loc _) = Just loc
|
getErrorLoc (IncompatibleFieldUpdate loc _) = Just loc
|
||||||
getErrorLoc (InvalidImplicits loc _ y tm) = Just loc
|
getErrorLoc (InvalidImplicits loc _ y tm) = Just loc
|
||||||
getErrorLoc (TryWithImplicits loc _ _) = Just loc
|
getErrorLoc (TryWithImplicits loc _ _) = Just loc
|
||||||
getErrorLoc (CantSolveGoal loc tm) = Just loc
|
getErrorLoc (CantSolveGoal loc _ tm) = Just loc
|
||||||
getErrorLoc (DeterminingArg loc y env tm) = Just loc
|
getErrorLoc (DeterminingArg loc y env tm) = Just loc
|
||||||
getErrorLoc (UnsolvedHoles ((loc, _) :: xs)) = Just loc
|
getErrorLoc (UnsolvedHoles ((loc, _) :: xs)) = Just loc
|
||||||
getErrorLoc (UnsolvedHoles []) = Nothing
|
getErrorLoc (UnsolvedHoles []) = Nothing
|
||||||
|
@ -276,14 +276,14 @@ export
|
|||||||
newMeta : {auto c : Ref Ctxt Defs} ->
|
newMeta : {auto c : Ref Ctxt Defs} ->
|
||||||
{auto u : Ref UST UState} ->
|
{auto u : Ref UST UState} ->
|
||||||
FC -> RigCount ->
|
FC -> RigCount ->
|
||||||
Env Term vars -> Name -> Term vars -> Core (Term vars)
|
Env Term vars -> Name -> Term vars -> Core (Int, Term vars)
|
||||||
newMeta {vars} fc rig env n ty
|
newMeta {vars} fc rig env n ty
|
||||||
= do let hty = abstractEnvType fc env ty
|
= do let hty = abstractEnvType fc env ty
|
||||||
let hole = newDef fc n rig hty Public (Hole False)
|
let hole = newDef fc n rig hty Public (Hole False)
|
||||||
log 10 $ "Adding new meta " ++ show n
|
log 10 $ "Adding new meta " ++ show n
|
||||||
idx <- addDef n hole
|
idx <- addDef n hole
|
||||||
addHoleName fc n idx
|
addHoleName fc n idx
|
||||||
pure (Meta fc n idx envArgs)
|
pure (idx, Meta fc n idx envArgs)
|
||||||
where
|
where
|
||||||
envArgs : List (Term vars)
|
envArgs : List (Term vars)
|
||||||
envArgs = let args = reverse (mkConstantAppArgs {done = []} False fc env []) in
|
envArgs = let args = reverse (mkConstantAppArgs {done = []} False fc env []) in
|
||||||
|
@ -263,7 +263,7 @@ mutual
|
|||||||
let argTyG = gnf defs env argTy
|
let argTyG = gnf defs env argTy
|
||||||
-- Can't use 'metaVar' to record it for binding because it's
|
-- Can't use 'metaVar' to record it for binding because it's
|
||||||
-- in a different scope... so it'll stay global
|
-- in a different scope... so it'll stay global
|
||||||
retTy <- newMeta {vars = argn :: vars}
|
(_, retTy) <- newMeta {vars = argn :: vars}
|
||||||
fc Rig0 (Pi RigW Explicit argTy :: env)
|
fc Rig0 (Pi RigW Explicit argTy :: env)
|
||||||
retn (TType fc)
|
retn (TType fc)
|
||||||
(argv, argt) <- check rig (nextLevel elabinfo)
|
(argv, argt) <- check rig (nextLevel elabinfo)
|
||||||
|
@ -186,7 +186,8 @@ metaVar : {auto c : Ref Ctxt Defs} ->
|
|||||||
FC -> RigCount ->
|
FC -> RigCount ->
|
||||||
Env Term vars -> Name -> Term vars -> Core (Term vars)
|
Env Term vars -> Name -> Term vars -> Core (Term vars)
|
||||||
metaVar fc rig env n ty
|
metaVar fc rig env n ty
|
||||||
= newMeta fc rig env n ty
|
= do (_, tm) <- newMeta fc rig env n ty
|
||||||
|
pure tm
|
||||||
|
|
||||||
-- Elaboration info (passed to recursive calls)
|
-- Elaboration info (passed to recursive calls)
|
||||||
public export
|
public export
|
||||||
|
@ -1,5 +1,6 @@
|
|||||||
module Yaffle.REPL
|
module Yaffle.REPL
|
||||||
|
|
||||||
|
import Core.AutoSearch
|
||||||
import Core.Context
|
import Core.Context
|
||||||
import Core.Core
|
import Core.Core
|
||||||
import Core.Env
|
import Core.Env
|
||||||
@ -7,6 +8,7 @@ import Core.FC
|
|||||||
import Core.Normalise
|
import Core.Normalise
|
||||||
import Core.TT
|
import Core.TT
|
||||||
import Core.Unify
|
import Core.Unify
|
||||||
|
import Core.Value
|
||||||
|
|
||||||
import TTImp.Elab
|
import TTImp.Elab
|
||||||
import TTImp.Elab.Check
|
import TTImp.Elab.Check
|
||||||
@ -40,13 +42,16 @@ process (Check ttimp)
|
|||||||
ty <- normaliseHoles defs [] tyh
|
ty <- normaliseHoles defs [] tyh
|
||||||
coreLift (printLn !(toFullNames ty))
|
coreLift (printLn !(toFullNames ty))
|
||||||
pure True
|
pure True
|
||||||
process (ProofSearch n)
|
process (ProofSearch n_in)
|
||||||
= do throw (InternalError "Not implemented")
|
= do defs <- get Ctxt
|
||||||
-- tm <- search () False 1000 [] (UN "(interactive)") Nothing n
|
[(n, i, ty)] <- lookupTyName n_in (gamma defs)
|
||||||
-- gam <- get Ctxt
|
| [] => throw (UndefinedName toplevelFC n_in)
|
||||||
-- coreLift (putStrLn (show (normalise gam [] tm)))
|
| ns => throw (AmbiguousName toplevelFC (map fst ns))
|
||||||
-- dumpConstraints 0 True
|
def <- search toplevelFC RigW n ty []
|
||||||
-- pure True
|
defs <- get Ctxt
|
||||||
|
defnf <- normaliseHoles defs [] def
|
||||||
|
coreLift (printLn !(toFullNames defnf))
|
||||||
|
pure True
|
||||||
process (DebugInfo n)
|
process (DebugInfo n)
|
||||||
= do defs <- get Ctxt
|
= do defs <- get Ctxt
|
||||||
traverse showInfo !(lookupDefName n (gamma defs))
|
traverse showInfo !(lookupDefName n (gamma defs))
|
||||||
|
@ -3,6 +3,7 @@ package yaffle
|
|||||||
modules =
|
modules =
|
||||||
Control.Delayed,
|
Control.Delayed,
|
||||||
|
|
||||||
|
Core.AutoSearch,
|
||||||
Core.Binary,
|
Core.Binary,
|
||||||
Core.CaseBuilder,
|
Core.CaseBuilder,
|
||||||
Core.CaseTree,
|
Core.CaseTree,
|
||||||
|
Loading…
Reference in New Issue
Block a user