mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-11-24 15:07:37 +03:00
REPL commands: load-package & fuzzy-search (#1318)
REPL commands: load-package & fuzzy-search Move REPL-related code over to its own namespace
This commit is contained in:
parent
489b85aae5
commit
70158b9dd1
@ -83,8 +83,6 @@ modules =
|
||||
Idris.Pretty.Render,
|
||||
Idris.ProcessIdr,
|
||||
Idris.REPL,
|
||||
Idris.REPLCommon,
|
||||
Idris.REPLOpts,
|
||||
Idris.Resugar,
|
||||
Idris.SetOptions,
|
||||
Idris.Syntax,
|
||||
@ -102,6 +100,10 @@ modules =
|
||||
Idris.IDEMode.SyntaxHighlight,
|
||||
Idris.IDEMode.TokenLine,
|
||||
|
||||
Idris.REPL.Common,
|
||||
Idris.REPL.Opts,
|
||||
Idris.REPL.FuzzySearch,
|
||||
|
||||
Libraries.Control.ANSI,
|
||||
Libraries.Control.ANSI.CSI,
|
||||
Libraries.Control.ANSI.SGR,
|
||||
|
@ -60,3 +60,8 @@ lowerMaybe (Just x) = x
|
||||
export
|
||||
raiseToMaybe : (Monoid a, Eq a) => a -> Maybe a
|
||||
raiseToMaybe x = if x == neutral then Nothing else Just x
|
||||
|
||||
public export
|
||||
filter : (a -> Bool) -> Maybe a -> Maybe a
|
||||
filter _ Nothing = Nothing
|
||||
filter f (Just x) = toMaybe (f x) x
|
||||
|
@ -12,7 +12,7 @@ import System.Path
|
||||
------------------------------------------------------------------------------
|
||||
-- Filenames
|
||||
|
||||
||| A `Filename root` is anchored in `root`.
|
||||
||| A `Filename root` is anchored in `root`.
|
||||
||| We use a `data` type so that Idris can easily infer `root` when passing
|
||||
||| a `FileName` around. We do not use a `record` because we do not want to
|
||||
||| allow users to manufacture their own `FileName`.
|
||||
@ -26,6 +26,11 @@ export
|
||||
fileName : FileName root -> String
|
||||
fileName (MkFileName str) = str
|
||||
|
||||
namespace FileName
|
||||
export
|
||||
toRelative : FileName root -> FileName (parse "")
|
||||
toRelative (MkFileName x) = MkFileName x
|
||||
|
||||
||| Convert a filename anchored in `root` to a filepath by appending the name
|
||||
||| to the root path.
|
||||
export
|
||||
@ -63,6 +68,13 @@ export
|
||||
emptyTree : Tree root
|
||||
emptyTree = MkTree [] []
|
||||
|
||||
namespace Tree
|
||||
||| No run time information is changed,
|
||||
||| so we assert the identity.
|
||||
export
|
||||
toRelative : Tree root -> Tree (parse "")
|
||||
toRelative x = believe_me x
|
||||
|
||||
||| Filter out files and directories that do not satisfy a given predicate.
|
||||
export
|
||||
filter : (filePred, dirPred : {root : _} -> FileName root -> Bool) ->
|
||||
|
@ -523,7 +523,8 @@ fileStem path = pure $ fst $ splitFileName !(fileName path)
|
||||
||| - Otherwise, the portion of the file name after the last ".".
|
||||
export
|
||||
extension : String -> Maybe String
|
||||
extension path = pure $ snd $ splitFileName !(fileName path)
|
||||
extension path = fileName path >>=
|
||||
filter (/= "") . Just . snd . splitFileName
|
||||
|
||||
||| Updates the file name in the path.
|
||||
|||
|
||||
|
@ -6,7 +6,7 @@ import Core.Env
|
||||
import Core.TT
|
||||
|
||||
import Idris.Pretty.Render
|
||||
import Idris.REPLOpts
|
||||
import Idris.REPL.Opts
|
||||
import Idris.Resugar
|
||||
import Idris.Syntax
|
||||
|
||||
|
@ -164,7 +164,7 @@ stMain cgs opts
|
||||
let ideSocket = ideModeSocket opts
|
||||
let outmode = if ide then IDEMode 0 stdin stdout else REPL False
|
||||
let fname = findInput opts
|
||||
o <- newRef ROpts (REPLOpts.defaultOpts fname outmode cgs)
|
||||
o <- newRef ROpts (REPL.Opts.defaultOpts fname outmode cgs)
|
||||
|
||||
finish <- showInfo opts
|
||||
when (not finish) $ do
|
||||
|
@ -7,7 +7,7 @@ import Core.Env
|
||||
import Core.Options
|
||||
import Core.Value
|
||||
|
||||
import Idris.REPLOpts
|
||||
import Idris.REPL.Opts
|
||||
import Idris.Resugar
|
||||
import Idris.Syntax
|
||||
import Idris.Pretty
|
||||
|
@ -14,7 +14,7 @@ import TTImp.TTImp
|
||||
import TTImp.Utils
|
||||
|
||||
import Idris.IDEMode.TokenLine
|
||||
import Idris.REPLOpts
|
||||
import Idris.REPL.Opts
|
||||
import Idris.Resugar
|
||||
import Idris.Syntax
|
||||
|
||||
|
@ -2,7 +2,7 @@ module Idris.IDEMode.Commands
|
||||
|
||||
import Core.Core
|
||||
import Core.Name
|
||||
import public Idris.REPLOpts
|
||||
import public Idris.REPL.Opts
|
||||
import Libraries.Utils.Hex
|
||||
|
||||
import System.File
|
||||
|
@ -15,7 +15,7 @@ import Idris.Desugar
|
||||
import Idris.Error
|
||||
import Idris.Parser
|
||||
import Idris.ProcessIdr
|
||||
import Idris.REPLCommon
|
||||
import Idris.REPL.Common
|
||||
import Idris.Syntax
|
||||
import Idris.Pretty
|
||||
|
||||
|
@ -31,8 +31,8 @@ import Idris.CommandLine
|
||||
import Idris.ModTree
|
||||
import Idris.ProcessIdr
|
||||
import Idris.REPL
|
||||
import Idris.REPLCommon
|
||||
import Idris.REPLOpts
|
||||
import Idris.REPL.Common
|
||||
import Idris.REPL.Opts
|
||||
import Idris.SetOptions
|
||||
import Idris.Syntax
|
||||
import Idris.Version
|
||||
|
@ -1890,6 +1890,8 @@ parserCommandsForHelp =
|
||||
, noArgCmd (ParseREPLCmd ["version"]) ShowVersion "Display the Idris version"
|
||||
, noArgCmd (ParseREPLCmd ["?", "h", "help"]) Help "Display this help text"
|
||||
, declsArgCmd (ParseKeywordCmd "let") NewDefn "Define a new value"
|
||||
, stringArgCmd (ParseREPLCmd ["lp", "loadpackage"]) ImportPackage "Load all modules of the package"
|
||||
, exprArgCmd (ParseREPLCmd ["fs", "fsearch"]) FuzzyTypeSearch "Search for global definitions by sketching the names distribution of the wanted type(s)."
|
||||
]
|
||||
|
||||
export
|
||||
|
@ -11,7 +11,7 @@ import public Libraries.Text.PrettyPrint.Prettyprinter
|
||||
import public Libraries.Text.PrettyPrint.Prettyprinter.Util
|
||||
|
||||
import Algebra
|
||||
import Idris.REPLOpts
|
||||
import Idris.REPL.Opts
|
||||
import Idris.Syntax
|
||||
|
||||
%default covering
|
||||
|
@ -2,7 +2,7 @@ module Idris.Pretty.Render
|
||||
|
||||
import Core.Core
|
||||
|
||||
import Idris.REPLOpts
|
||||
import Idris.REPL.Opts
|
||||
|
||||
import Libraries.Control.ANSI.SGR
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter
|
||||
|
@ -21,8 +21,8 @@ import TTImp.TTImp
|
||||
|
||||
import Idris.Desugar
|
||||
import Idris.Parser
|
||||
import Idris.REPLCommon
|
||||
import Idris.REPLOpts
|
||||
import Idris.REPL.Common
|
||||
import Idris.REPL.Opts
|
||||
import Idris.Syntax
|
||||
import Idris.Pretty
|
||||
|
||||
|
@ -38,10 +38,12 @@ import Idris.Parser
|
||||
import Idris.Pretty
|
||||
import Idris.ProcessIdr
|
||||
import Idris.Resugar
|
||||
import public Idris.REPLCommon
|
||||
import Idris.Syntax
|
||||
import Idris.Version
|
||||
|
||||
import public Idris.REPL.Common
|
||||
import Idris.REPL.FuzzySearch
|
||||
|
||||
import TTImp.Elab
|
||||
import TTImp.Elab.Check
|
||||
import TTImp.Elab.Local
|
||||
@ -61,13 +63,18 @@ import Libraries.Data.NameMap
|
||||
import Libraries.Data.PosMap
|
||||
import Data.Stream
|
||||
import Data.Strings
|
||||
import Data.DPair
|
||||
import Libraries.Data.String.Extra
|
||||
import Libraries.Data.List.Extra
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter.Util
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter.Render.Terminal
|
||||
import Libraries.Utils.Path
|
||||
import Libraries.System.Directory.Tree
|
||||
|
||||
import System
|
||||
import System.File
|
||||
import System.Directory
|
||||
|
||||
%hide Data.Strings.lines
|
||||
%hide Data.Strings.lines'
|
||||
@ -237,14 +244,6 @@ lookupDefTyName : Name -> Context ->
|
||||
Core (List (Name, Int, (Def, ClosedTerm)))
|
||||
lookupDefTyName = lookupNameBy (\g => (definition g, type g))
|
||||
|
||||
public export
|
||||
data EditResult : Type where
|
||||
DisplayEdit : Doc IdrisAnn -> EditResult
|
||||
EditError : Doc IdrisAnn -> EditResult
|
||||
MadeLemma : Maybe String -> Name -> PTerm -> String -> EditResult
|
||||
MadeWith : Maybe String -> List String -> EditResult
|
||||
MadeCase : Maybe String -> List String -> EditResult
|
||||
|
||||
updateFile : {auto r : Ref ROpts REPLOpts} ->
|
||||
(List String -> List String) -> Core EditResult
|
||||
updateFile update
|
||||
@ -541,43 +540,6 @@ processEdit (MakeWith upd line name)
|
||||
then updateFile (addMadeCase markM w (max 0 (integerToNat (cast (line - 1)))))
|
||||
else pure $ MadeWith markM w
|
||||
|
||||
public export
|
||||
data MissedResult : Type where
|
||||
CasesMissing : Name -> List String -> MissedResult
|
||||
CallsNonCovering : Name -> List Name -> MissedResult
|
||||
AllCasesCovered : Name -> MissedResult
|
||||
|
||||
public export
|
||||
data REPLResult : Type where
|
||||
Done : REPLResult
|
||||
REPLError : Doc IdrisAnn -> REPLResult
|
||||
Executed : PTerm -> REPLResult
|
||||
RequestedHelp : REPLResult
|
||||
Evaluated : PTerm -> (Maybe PTerm) -> REPLResult
|
||||
Printed : Doc IdrisAnn -> REPLResult
|
||||
TermChecked : PTerm -> PTerm -> REPLResult
|
||||
FileLoaded : String -> REPLResult
|
||||
ModuleLoaded : String -> REPLResult
|
||||
ErrorLoadingModule : String -> Error -> REPLResult
|
||||
ErrorLoadingFile : String -> FileError -> REPLResult
|
||||
ErrorsBuildingFile : String -> List Error -> REPLResult
|
||||
NoFileLoaded : REPLResult
|
||||
CurrentDirectory : String -> REPLResult
|
||||
CompilationFailed: REPLResult
|
||||
Compiled : String -> REPLResult
|
||||
ProofFound : PTerm -> REPLResult
|
||||
Missed : List MissedResult -> REPLResult
|
||||
CheckedTotal : List (Name, Totality) -> REPLResult
|
||||
FoundHoles : List HoleData -> REPLResult
|
||||
OptionsSet : List REPLOpt -> REPLResult
|
||||
LogLevelSet : Maybe LogLevel -> REPLResult
|
||||
ConsoleWidthSet : Maybe Nat -> REPLResult
|
||||
ColorSet : Bool -> REPLResult
|
||||
VersionIs : Version -> REPLResult
|
||||
DefDeclared : REPLResult
|
||||
Exited : REPLResult
|
||||
Edited : EditResult -> REPLResult
|
||||
|
||||
getItDecls :
|
||||
{auto o : Ref ROpts REPLOpts} ->
|
||||
Core (List ImpDecl)
|
||||
@ -682,47 +644,6 @@ loadMainFile f
|
||||
[] => pure (FileLoaded f)
|
||||
_ => pure (ErrorsBuildingFile f errs)
|
||||
|
||||
docsOrSignature : {auto o : Ref ROpts REPLOpts} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
FC -> Name -> Core (List String)
|
||||
docsOrSignature fc n
|
||||
= do syn <- get Syn
|
||||
defs <- get Ctxt
|
||||
all@(_ :: _) <- lookupCtxtName n (gamma defs)
|
||||
| _ => undefinedName fc n
|
||||
let ns@(_ :: _) = concatMap (\n => lookupName n (docstrings syn))
|
||||
(map fst all)
|
||||
| [] => typeSummary defs
|
||||
pure <$> getDocsForName fc n
|
||||
where
|
||||
typeSummary : Defs -> Core (List String)
|
||||
typeSummary defs = do Just def <- lookupCtxtExact n (gamma defs)
|
||||
| Nothing => pure []
|
||||
ty <- normaliseHoles defs [] (type def)
|
||||
pure [(show n) ++ " : " ++ (show !(resugar [] ty))]
|
||||
|
||||
equivTypes : {auto c : Ref Ctxt Defs} ->
|
||||
(ty1 : ClosedTerm) ->
|
||||
(ty2 : ClosedTerm) ->
|
||||
Core Bool
|
||||
equivTypes ty1 ty2 =
|
||||
do let False = isErased ty1
|
||||
| _ => pure False
|
||||
logTerm "typesearch.equiv" 10 "Candidate: " ty1
|
||||
defs <- get Ctxt
|
||||
True <- pure (!(getArity defs [] ty1) == !(getArity defs [] ty2))
|
||||
| False => pure False
|
||||
_ <- newRef UST initUState
|
||||
b <- catch
|
||||
(do res <- unify inTerm replFC [] ty1 ty2
|
||||
case res of
|
||||
(MkUnifyResult [] _ [] NoLazy) => pure True
|
||||
_ => pure False)
|
||||
(\err => pure False)
|
||||
when b $ logTerm "typesearch.equiv" 20 "Accepted: " ty1
|
||||
pure b
|
||||
|
||||
||| Process a single `REPLCmd`
|
||||
|||
|
||||
||| Returns `REPLResult` for display by the higher level shell which
|
||||
@ -856,11 +777,12 @@ process (TypeSearch searchTerm)
|
||||
let defs = flip mapMaybe defs $ \ md =>
|
||||
do d <- md
|
||||
guard (visibleIn curr (fullname d) (visibility d))
|
||||
guard (isJust $ userNameRoot (fullname d))
|
||||
pure d
|
||||
allDefs <- traverse (resolved ctxt) defs
|
||||
filterM (\def => equivTypes def.type ty') allDefs
|
||||
put Ctxt defs
|
||||
doc <- traverse (docsOrSignature replFC) $ (.fullname) <$> filteredDefs
|
||||
doc <- traverse (docsOrSignature replFC) $ fullname <$> filteredDefs
|
||||
pure $ Printed $ vsep $ pretty <$> (intersperse "\n" $ join doc)
|
||||
process (Missing n)
|
||||
= do defs <- get Ctxt
|
||||
@ -949,6 +871,34 @@ process NOP
|
||||
= pure Done
|
||||
process ShowVersion
|
||||
= pure $ VersionIs version
|
||||
process (ImportPackage package) = do
|
||||
defs <- get Ctxt
|
||||
let searchDirs = defs.options.dirs.extra_dirs
|
||||
let Just packageDir = find
|
||||
(\d => isInfixOf package (fromMaybe d (fileName d)))
|
||||
searchDirs
|
||||
| _ => pure (REPLError (pretty "Package not found in the known search directories"))
|
||||
let packageDirPath = parse packageDir
|
||||
tree <- coreLift $ explore packageDirPath
|
||||
fentries <- coreLift $ toPaths (toRelative tree)
|
||||
errs <- for fentries \entry => do
|
||||
let entry' = dropExtension entry
|
||||
let sp = forget $ split (== dirSeparator) entry'
|
||||
let ns = concat $ intersperse "." sp
|
||||
let ns' = mkNamespace ns
|
||||
catch (do addImport (MkImport emptyFC False (nsAsModuleIdent ns') ns'); pure Nothing)
|
||||
(\err => pure (Just err))
|
||||
let errs' = catMaybes errs
|
||||
res <- case errs' of
|
||||
[] => pure (pretty "Done")
|
||||
onePlus => pure $ vsep !(traverse display onePlus)
|
||||
pure (Printed res)
|
||||
where
|
||||
toPaths : {root : _} -> Tree root -> IO (List String)
|
||||
toPaths tree =
|
||||
depthFirst (\x => map (toFilePath x ::) . force) tree (pure [])
|
||||
|
||||
process (FuzzyTypeSearch expr) = fuzzySearch expr
|
||||
|
||||
processCatch : {auto c : Ref Ctxt Defs} ->
|
||||
{auto u : Ref UST UState} ->
|
||||
|
@ -1,20 +1,29 @@
|
||||
module Idris.REPLCommon
|
||||
module Idris.REPL.Common
|
||||
|
||||
import Core.Context
|
||||
import Core.Env
|
||||
import Core.Context.Log
|
||||
import Core.InitPrimitives
|
||||
import Core.Metadata
|
||||
import Core.Options
|
||||
import Core.Primitives
|
||||
import Core.TT
|
||||
import Core.Unify
|
||||
import Core.UnifyState
|
||||
|
||||
import Idris.DocString
|
||||
import Idris.Error
|
||||
import Idris.IDEMode.Commands
|
||||
import public Idris.REPLOpts
|
||||
import Idris.Syntax
|
||||
import Idris.IDEMode.Holes
|
||||
import Idris.Pretty
|
||||
import public Idris.REPL.Opts
|
||||
import Idris.Resugar
|
||||
import Idris.Syntax
|
||||
import Idris.Version
|
||||
|
||||
import Libraries.Data.ANameMap
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter
|
||||
|
||||
import Data.List
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter.Render.Terminal
|
||||
import System.File
|
||||
|
||||
%default covering
|
||||
|
||||
@ -135,3 +144,92 @@ resetContext
|
||||
put UST initUState
|
||||
put Syn initSyntax
|
||||
put MD initMetadata
|
||||
|
||||
public export
|
||||
data EditResult : Type where
|
||||
DisplayEdit : Doc IdrisAnn -> EditResult
|
||||
EditError : Doc IdrisAnn -> EditResult
|
||||
MadeLemma : Maybe String -> Name -> PTerm -> String -> EditResult
|
||||
MadeWith : Maybe String -> List String -> EditResult
|
||||
MadeCase : Maybe String -> List String -> EditResult
|
||||
|
||||
public export
|
||||
data MissedResult : Type where
|
||||
CasesMissing : Name -> List String -> MissedResult
|
||||
CallsNonCovering : Name -> List Name -> MissedResult
|
||||
AllCasesCovered : Name -> MissedResult
|
||||
|
||||
public export
|
||||
data REPLResult : Type where
|
||||
Done : REPLResult
|
||||
REPLError : Doc IdrisAnn -> REPLResult
|
||||
Executed : PTerm -> REPLResult
|
||||
RequestedHelp : REPLResult
|
||||
Evaluated : PTerm -> (Maybe PTerm) -> REPLResult
|
||||
Printed : Doc IdrisAnn -> REPLResult
|
||||
TermChecked : PTerm -> PTerm -> REPLResult
|
||||
FileLoaded : String -> REPLResult
|
||||
ModuleLoaded : String -> REPLResult
|
||||
ErrorLoadingModule : String -> Error -> REPLResult
|
||||
ErrorLoadingFile : String -> FileError -> REPLResult
|
||||
ErrorsBuildingFile : String -> List Error -> REPLResult
|
||||
NoFileLoaded : REPLResult
|
||||
CurrentDirectory : String -> REPLResult
|
||||
CompilationFailed: REPLResult
|
||||
Compiled : String -> REPLResult
|
||||
ProofFound : PTerm -> REPLResult
|
||||
Missed : List MissedResult -> REPLResult
|
||||
CheckedTotal : List (Name, Totality) -> REPLResult
|
||||
FoundHoles : List HoleData -> REPLResult
|
||||
OptionsSet : List REPLOpt -> REPLResult
|
||||
LogLevelSet : Maybe LogLevel -> REPLResult
|
||||
ConsoleWidthSet : Maybe Nat -> REPLResult
|
||||
ColorSet : Bool -> REPLResult
|
||||
VersionIs : Version -> REPLResult
|
||||
DefDeclared : REPLResult
|
||||
Exited : REPLResult
|
||||
Edited : EditResult -> REPLResult
|
||||
|
||||
export
|
||||
docsOrSignature : {auto o : Ref ROpts REPLOpts} ->
|
||||
{auto c : Ref Ctxt Defs} ->
|
||||
{auto s : Ref Syn SyntaxInfo} ->
|
||||
FC -> Name -> Core (List String)
|
||||
docsOrSignature fc n
|
||||
= do syn <- get Syn
|
||||
defs <- get Ctxt
|
||||
all@(_ :: _) <- lookupCtxtName n (gamma defs)
|
||||
| _ => undefinedName fc n
|
||||
let ns@(_ :: _) = concatMap (\n => lookupName n (docstrings syn))
|
||||
(map fst all)
|
||||
| [] => typeSummary defs
|
||||
pure <$> getDocsForName fc n
|
||||
where
|
||||
typeSummary : Defs -> Core (List String)
|
||||
typeSummary defs = do Just def <- lookupCtxtExact n (gamma defs)
|
||||
| Nothing => pure []
|
||||
ty <- normaliseHoles defs [] (type def)
|
||||
pure [(show n) ++ " : " ++ (show !(resugar [] ty))]
|
||||
|
||||
export
|
||||
equivTypes : {auto c : Ref Ctxt Defs} ->
|
||||
(ty1 : ClosedTerm) ->
|
||||
(ty2 : ClosedTerm) ->
|
||||
Core Bool
|
||||
equivTypes ty1 ty2 =
|
||||
do let False = isErased ty1
|
||||
| _ => pure False
|
||||
logTerm "typesearch.equiv" 10 "Candidate: " ty1
|
||||
defs <- get Ctxt
|
||||
True <- pure (!(getArity defs [] ty1) == !(getArity defs [] ty2))
|
||||
| False => pure False
|
||||
_ <- newRef UST initUState
|
||||
b <- catch
|
||||
(do res <- unify inTerm replFC [] ty1 ty2
|
||||
case res of
|
||||
(MkUnifyResult [] _ [] NoLazy) => pure True
|
||||
_ => pure False)
|
||||
(\err => pure False)
|
||||
when b $ logTerm "typesearch.equiv" 20 "Accepted: " ty1
|
||||
pure b
|
||||
|
215
src/Idris/REPL/FuzzySearch.idr
Normal file
215
src/Idris/REPL/FuzzySearch.idr
Normal file
@ -0,0 +1,215 @@
|
||||
module Idris.REPL.FuzzySearch
|
||||
|
||||
import Core.AutoSearch
|
||||
import Core.CaseTree
|
||||
import Core.CompileExpr
|
||||
import Core.Context
|
||||
import Core.Context.Log
|
||||
import Core.Env
|
||||
import Core.InitPrimitives
|
||||
import Core.LinearCheck
|
||||
import Core.Metadata
|
||||
import Core.Normalise
|
||||
import Core.Options
|
||||
import Core.Termination
|
||||
import Core.TT
|
||||
import Core.Unify
|
||||
|
||||
import Idris.Desugar
|
||||
import Idris.DocString
|
||||
import Idris.Error
|
||||
import Idris.IDEMode.CaseSplit
|
||||
import Idris.IDEMode.Commands
|
||||
import Idris.IDEMode.MakeClause
|
||||
import Idris.IDEMode.Holes
|
||||
import Idris.ModTree
|
||||
import Idris.Parser
|
||||
import Idris.Pretty
|
||||
import Idris.ProcessIdr
|
||||
import Idris.Resugar
|
||||
import Idris.Syntax
|
||||
import Idris.Version
|
||||
|
||||
import public Idris.REPL.Common
|
||||
|
||||
import Data.List
|
||||
import Data.List1
|
||||
import Data.Maybe
|
||||
import Libraries.Data.ANameMap
|
||||
import Libraries.Data.NameMap
|
||||
import Libraries.Data.PosMap
|
||||
import Data.Stream
|
||||
import Data.Strings
|
||||
import Data.DPair
|
||||
import Libraries.Data.String.Extra
|
||||
import Libraries.Data.List.Extra
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter.Util
|
||||
import Libraries.Text.PrettyPrint.Prettyprinter.Render.Terminal
|
||||
import Libraries.Utils.Path
|
||||
import Libraries.System.Directory.Tree
|
||||
|
||||
import System
|
||||
import System.File
|
||||
import System.Directory
|
||||
|
||||
export
|
||||
fuzzySearch : {auto c : Ref Ctxt Defs}
|
||||
-> {auto u : Ref UST UState}
|
||||
-> {auto s : Ref Syn SyntaxInfo}
|
||||
-> {auto m : Ref MD Metadata}
|
||||
-> {auto o : Ref ROpts REPLOpts}
|
||||
-> PTerm
|
||||
-> Core REPLResult
|
||||
fuzzySearch expr = do
|
||||
let Just (neg, pos) = parseExpr expr
|
||||
| _ => pure (REPLError (pretty "Bad expression, expected"
|
||||
<++> code (pretty "B")
|
||||
<++> pretty "or"
|
||||
<++> code (pretty "_ -> B")
|
||||
<++> pretty "or"
|
||||
<++> code (pretty "A -> B")
|
||||
<+> pretty ", where"
|
||||
<++> code (pretty "A")
|
||||
<++> pretty "and"
|
||||
<++> code (pretty "B")
|
||||
<++> pretty "are spines of global names"))
|
||||
defs <- branch
|
||||
let curr = currentNS defs
|
||||
let ctxt = gamma defs
|
||||
filteredDefs <-
|
||||
do names <- allNames ctxt
|
||||
defs <- traverse (flip lookupCtxtExact ctxt) names
|
||||
let defs = flip mapMaybe defs $ \ md =>
|
||||
do d <- md
|
||||
guard (visibleIn curr (fullname d) (visibility d))
|
||||
guard (isJust $ userNameRoot (fullname d))
|
||||
pure d
|
||||
allDefs <- traverse (resolved ctxt) defs
|
||||
filterM (\def => fuzzyMatch neg pos def.type) allDefs
|
||||
put Ctxt defs
|
||||
doc <- traverse (docsOrSignature replFC) $ fullname <$> filteredDefs
|
||||
pure $ Printed $ vsep $ pretty <$> (intersperse "\n" $ join doc)
|
||||
where
|
||||
|
||||
data NameOrConst = AName Name
|
||||
| AInt
|
||||
| AInteger
|
||||
| ABits8
|
||||
| ABits16
|
||||
| ABits32
|
||||
| ABits64
|
||||
| AString
|
||||
| AChar
|
||||
| ADouble
|
||||
| AWorld
|
||||
| AType
|
||||
|
||||
eqConst : (x, y : NameOrConst) -> Bool
|
||||
eqConst AInt AInt = True
|
||||
eqConst AInteger AInteger = True
|
||||
eqConst ABits8 ABits8 = True
|
||||
eqConst ABits16 ABits16 = True
|
||||
eqConst ABits32 ABits32 = True
|
||||
eqConst ABits64 ABits64 = True
|
||||
eqConst AString AString = True
|
||||
eqConst AChar AChar = True
|
||||
eqConst ADouble ADouble = True
|
||||
eqConst AWorld AWorld = True
|
||||
eqConst AType AType = True
|
||||
eqConst _ _ = False
|
||||
|
||||
parseNameOrConst : PTerm -> Maybe NameOrConst
|
||||
parseNameOrConst (PRef _ n) = Just (AName n)
|
||||
parseNameOrConst (PPrimVal _ IntType) = Just AInt
|
||||
parseNameOrConst (PPrimVal _ IntegerType) = Just AInteger
|
||||
parseNameOrConst (PPrimVal _ Bits8Type) = Just ABits8
|
||||
parseNameOrConst (PPrimVal _ Bits16Type) = Just ABits16
|
||||
parseNameOrConst (PPrimVal _ Bits32Type) = Just ABits32
|
||||
parseNameOrConst (PPrimVal _ Bits64Type) = Just ABits64
|
||||
parseNameOrConst (PPrimVal _ StringType) = Just AString
|
||||
parseNameOrConst (PPrimVal _ CharType) = Just AChar
|
||||
parseNameOrConst (PPrimVal _ DoubleType) = Just ADouble
|
||||
parseNameOrConst (PPrimVal _ WorldType) = Just AWorld
|
||||
parseNameOrConst (PType _) = Just AType
|
||||
parseNameOrConst _ = Nothing
|
||||
|
||||
parseExpr' : PTerm -> Maybe (List NameOrConst)
|
||||
parseExpr' (PApp _ f x) =
|
||||
[| parseNameOrConst x :: parseExpr' f |]
|
||||
parseExpr' x = (:: []) <$> parseNameOrConst x
|
||||
|
||||
parseExpr : PTerm -> Maybe (List NameOrConst, List NameOrConst)
|
||||
parseExpr (PPi _ _ _ _ a (PImplicit _)) = do
|
||||
a' <- parseExpr' a
|
||||
pure (a', [])
|
||||
parseExpr (PPi _ _ _ _ a b) = do
|
||||
a' <- parseExpr' a
|
||||
b' <- parseExpr' b
|
||||
pure (a', b')
|
||||
parseExpr b = do
|
||||
b' <- parseExpr' b
|
||||
pure ([], b')
|
||||
|
||||
isApproximationOf : (given : Name)
|
||||
-> (candidate : Name)
|
||||
-> Bool
|
||||
isApproximationOf (NS ns n) (NS ns' n') =
|
||||
n == n' && Namespace.isApproximationOf ns ns'
|
||||
isApproximationOf (UN n) (NS ns' (UN n')) =
|
||||
n == n'
|
||||
isApproximationOf (NS ns n) _ =
|
||||
False
|
||||
isApproximationOf (UN n) (UN n') =
|
||||
n == n'
|
||||
isApproximationOf _ _ =
|
||||
False
|
||||
|
||||
isApproximationOf' : (given : NameOrConst)
|
||||
-> (candidate : NameOrConst)
|
||||
-> Bool
|
||||
isApproximationOf' (AName x) (AName y) =
|
||||
isApproximationOf x y
|
||||
isApproximationOf' a b = eqConst a b
|
||||
|
||||
||| Find all name and type literal occurrences.
|
||||
export
|
||||
doFind : List NameOrConst -> Term vars -> List NameOrConst
|
||||
doFind ns (Local fc x idx y) = ns
|
||||
doFind ns (Ref fc x name) = AName name :: ns
|
||||
doFind ns (Meta fc n i xs)
|
||||
= foldl doFind ns xs
|
||||
doFind ns (Bind fc x (Let _ c val ty) scope)
|
||||
= doFind (doFind (doFind ns val) ty) scope
|
||||
doFind ns (Bind fc x b scope)
|
||||
= doFind (doFind ns (binderType b)) scope
|
||||
doFind ns (App fc fn arg)
|
||||
= doFind (doFind ns fn) arg
|
||||
doFind ns (As fc s as tm) = doFind ns tm
|
||||
doFind ns (TDelayed fc x y) = doFind ns y
|
||||
doFind ns (TDelay fc x t y)
|
||||
= doFind (doFind ns t) y
|
||||
doFind ns (TForce fc r x) = doFind ns x
|
||||
doFind ns (PrimVal fc c) =
|
||||
fromMaybe [] ((:: []) <$> parseNameOrConst (PPrimVal fc c)) ++ ns
|
||||
doFind ns (Erased fc i) = ns
|
||||
doFind ns (TType fc) = AType :: ns
|
||||
|
||||
toFullNames' : NameOrConst -> Core NameOrConst
|
||||
toFullNames' (AName x) = AName <$> toFullNames x
|
||||
toFullNames' x = pure x
|
||||
|
||||
fuzzyMatch : (neg : List NameOrConst)
|
||||
-> (pos : List NameOrConst)
|
||||
-> Term vars
|
||||
-> Core Bool
|
||||
fuzzyMatch neg pos (Bind _ _ b sc) = do
|
||||
let refsB = doFind [] (binderType b)
|
||||
refsB <- traverse toFullNames' refsB
|
||||
let neg' = diffBy isApproximationOf' neg refsB
|
||||
fuzzyMatch neg' pos sc
|
||||
fuzzyMatch (_ :: _) pos tm = pure False
|
||||
fuzzyMatch [] pos tm = do
|
||||
let refsB = doFind [] tm
|
||||
refsB <- traverse toFullNames' refsB
|
||||
pure (isNil $ diffBy isApproximationOf' pos refsB)
|
@ -1,4 +1,4 @@
|
||||
module Idris.REPLOpts
|
||||
module Idris.REPL.Opts
|
||||
|
||||
import Compiler.Common
|
||||
import Idris.Syntax
|
@ -453,6 +453,7 @@ data REPLCmd : Type where
|
||||
Exec : PTerm -> REPLCmd
|
||||
Help : REPLCmd
|
||||
TypeSearch : PTerm -> REPLCmd
|
||||
FuzzyTypeSearch : PTerm -> REPLCmd
|
||||
DebugInfo : Name -> REPLCmd
|
||||
SetOpt : REPLOpt -> REPLCmd
|
||||
GetOpts : REPLCmd
|
||||
@ -472,6 +473,7 @@ data REPLCmd : Type where
|
||||
ShowVersion : REPLCmd
|
||||
Quit : REPLCmd
|
||||
NOP : REPLCmd
|
||||
ImportPackage : String -> REPLCmd
|
||||
|
||||
public export
|
||||
record Import where
|
||||
|
@ -45,3 +45,28 @@ zipMaybe : List a -> List b -> Maybe (List (a, b))
|
||||
zipMaybe [] [] = pure []
|
||||
zipMaybe (a::as) (b::bs) = ((a, b) ::) <$> zipMaybe as bs
|
||||
zipMaybe _ _ = Nothing
|
||||
|
||||
export
|
||||
findBy' : (a -> Bool) -> List a -> (List a, Maybe a, List a)
|
||||
findBy' f [] = ([], Nothing, [])
|
||||
findBy' f (x :: xs) =
|
||||
case f x of
|
||||
True => ([], Just x, xs)
|
||||
False =>
|
||||
let (pre, mb, post) = findBy' f xs in
|
||||
(x :: pre, mb, post)
|
||||
|
||||
||| Compute the difference of two lists by the given predicate.
|
||||
||| Lists are treated as bags.
|
||||
export
|
||||
diffBy : (a -> a -> Bool)
|
||||
-> List a
|
||||
-> List a
|
||||
-> List a
|
||||
diffBy f [] ys = []
|
||||
diffBy f (x :: xs) ys =
|
||||
let whole@(pre, mb, post) = findBy' (f x) ys
|
||||
ys' = pre ++ post in
|
||||
case mb of
|
||||
Just _ => diffBy f xs ys'
|
||||
Nothing => x :: diffBy f xs ys'
|
||||
|
@ -12,7 +12,7 @@ import Libraries.Utils.Path
|
||||
------------------------------------------------------------------------------
|
||||
-- Filenames
|
||||
|
||||
||| A `Filename root` is anchored in `root`.
|
||||
||| A `Filename root` is anchored in `root`.
|
||||
||| We use a `data` type so that Idris can easily infer `root` when passing
|
||||
||| a `FileName` around. We do not use a `record` because we do not want to
|
||||
||| allow users to manufacture their own `FileName`.
|
||||
@ -26,6 +26,11 @@ export
|
||||
fileName : FileName root -> String
|
||||
fileName (MkFileName str) = str
|
||||
|
||||
namespace FileName
|
||||
export
|
||||
toRelative : FileName root -> FileName (parse "")
|
||||
toRelative (MkFileName x) = MkFileName x
|
||||
|
||||
||| Convert a filename anchored in `root` to a filepath by appending the name
|
||||
||| to the root path.
|
||||
export
|
||||
@ -63,6 +68,13 @@ export
|
||||
emptyTree : Tree root
|
||||
emptyTree = MkTree [] []
|
||||
|
||||
namespace Tree
|
||||
||| No run time information is changed,
|
||||
||| so we assert the identity.
|
||||
export
|
||||
toRelative : Tree root -> Tree (parse "")
|
||||
toRelative x = believe_me x
|
||||
|
||||
||| Filter out files and directories that do not satisfy a given predicate.
|
||||
export
|
||||
filter : (filePred, dirPred : {root : _} -> FileName root -> Bool) ->
|
||||
@ -127,6 +139,14 @@ go dir acc = case !(dirEntry dir) of
|
||||
else { files $= (entry ::) } acc
|
||||
assert_total (go dir acc)
|
||||
|
||||
||| Depth first traversal of all of the files in a tree
|
||||
export
|
||||
covering
|
||||
depthFirst : ({root : Path} -> FileName root -> Lazy (IO a) -> IO a) ->
|
||||
{root : Path} -> Tree root -> IO a -> IO a
|
||||
depthFirst check t k =
|
||||
let next = foldr (\ (dir ** iot), def => depthFirst check !iot def) k t.subTrees in
|
||||
foldr (\ fn, def => check fn def) next t.files
|
||||
|
||||
||| Display a tree by printing it procedurally. Note that because directory
|
||||
||| trees contain suspended computations corresponding to their subtrees this
|
||||
|
@ -520,7 +520,14 @@ fileStem path = pure $ fst $ splitFileName !(fileName path)
|
||||
||| - Otherwise, the portion of the file name after the last ".".
|
||||
export
|
||||
extension : String -> Maybe String
|
||||
extension path = pure $ snd $ splitFileName !(fileName path)
|
||||
extension path = fileName path >>=
|
||||
filter (/= "") . Just . snd . splitFileName
|
||||
where
|
||||
-- TODO Use Data.Maybe.filter instead when next minor
|
||||
-- release comes out.
|
||||
filter : forall a. (a -> Bool) -> Maybe a -> Maybe a
|
||||
filter f Nothing = Nothing
|
||||
filter f (Just x) = toMaybe (f x) x
|
||||
|
||||
||| Updates the file name in the path.
|
||||
|||
|
||||
|
Loading…
Reference in New Issue
Block a user