mirror of
https://github.com/idris-lang/Idris2.git
synced 2025-01-05 06:37:23 +03:00
parent
405c266b5e
commit
8f71aae705
@ -1379,6 +1379,7 @@ lookupDefTyExact = lookupExactBy (\g => (definition g, type g))
|
|||||||
-- private names are only visible in this namespace if their namespace
|
-- private names are only visible in this namespace if their namespace
|
||||||
-- is the current namespace (or an outer one)
|
-- is the current namespace (or an outer one)
|
||||||
-- that is: the namespace of 'n' is a parent of nspace
|
-- that is: the namespace of 'n' is a parent of nspace
|
||||||
|
export
|
||||||
visibleIn : Namespace -> Name -> Visibility -> Bool
|
visibleIn : Namespace -> Name -> Visibility -> Bool
|
||||||
visibleIn nspace (NS ns n) Private = isParentOf ns nspace
|
visibleIn nspace (NS ns n) Private = isParentOf ns nspace
|
||||||
-- Public and Export names are always visible
|
-- Public and Export names are always visible
|
||||||
@ -1709,7 +1710,7 @@ setExternal fc tyn u
|
|||||||
|
|
||||||
export
|
export
|
||||||
addHintFor : {auto c : Ref Ctxt Defs} ->
|
addHintFor : {auto c : Ref Ctxt Defs} ->
|
||||||
FC -> Name -> Name -> Bool -> Bool -> Core ()
|
FC -> Name -> Name -> Bool -> Bool -> Core ()
|
||||||
addHintFor fc tyn_in hintn_in direct loading
|
addHintFor fc tyn_in hintn_in direct loading
|
||||||
= do defs <- get Ctxt
|
= do defs <- get Ctxt
|
||||||
tyn <- toFullNames tyn_in
|
tyn <- toFullNames tyn_in
|
||||||
|
@ -574,6 +574,11 @@ data Term : List Name -> Type where
|
|||||||
Term vars
|
Term vars
|
||||||
TType : FC -> Term vars
|
TType : FC -> Term vars
|
||||||
|
|
||||||
|
export
|
||||||
|
isErased : Term vars -> Bool
|
||||||
|
isErased (Erased _ _) = True
|
||||||
|
isErased _ = False
|
||||||
|
|
||||||
export
|
export
|
||||||
getLoc : Term vars -> FC
|
getLoc : Term vars -> FC
|
||||||
getLoc (Local fc _ _ _) = fc
|
getLoc (Local fc _ _ _) = fc
|
||||||
|
@ -678,15 +678,22 @@ equivTypes : {auto c : Ref Ctxt Defs} ->
|
|||||||
(ty1 : ClosedTerm) ->
|
(ty1 : ClosedTerm) ->
|
||||||
(ty2 : ClosedTerm) ->
|
(ty2 : ClosedTerm) ->
|
||||||
Core Bool
|
Core Bool
|
||||||
equivTypes ty1 ty2 = do defs <- get Ctxt
|
equivTypes ty1 ty2 =
|
||||||
True <- pure (!(getArity defs [] ty1) == !(getArity defs [] ty2))
|
do let False = isErased ty1
|
||||||
| False => pure False
|
| _ => pure False
|
||||||
_ <- newRef UST initUState
|
logTerm "typesearch.equiv" 10 "Candidate: " ty1
|
||||||
catch (do res <- unify inTerm replFC [] ty1 ty2
|
defs <- get Ctxt
|
||||||
case res of
|
True <- pure (!(getArity defs [] ty1) == !(getArity defs [] ty2))
|
||||||
(MkUnifyResult _ _ _ NoLazy) => pure True
|
| False => pure False
|
||||||
_ => pure False)
|
_ <- newRef UST initUState
|
||||||
(\err => pure False)
|
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`
|
||| Process a single `REPLCmd`
|
||||||
|||
|
|||
|
||||||
@ -807,23 +814,26 @@ process (Exec ctm)
|
|||||||
= execExp ctm
|
= execExp ctm
|
||||||
process Help
|
process Help
|
||||||
= pure RequestedHelp
|
= pure RequestedHelp
|
||||||
process (TypeSearch searchTerm@(PPi _ _ _ _ _ _))
|
process (TypeSearch searchTerm)
|
||||||
= do defs <- branch
|
= do defs <- branch
|
||||||
|
let curr = currentNS defs
|
||||||
let ctxt = gamma defs
|
let ctxt = gamma defs
|
||||||
rawTy <- desugar AnyExpr [] searchTerm
|
rawTy <- desugar AnyExpr [] searchTerm
|
||||||
let bound = piBindNames replFC [] rawTy
|
let bound = piBindNames replFC [] rawTy
|
||||||
(ty, _) <- elabTerm 0 InType [] (MkNested []) [] bound Nothing
|
(ty, _) <- elabTerm 0 InType [] (MkNested []) [] bound Nothing
|
||||||
ty' <- toResolvedNames ty
|
ty' <- toResolvedNames ty
|
||||||
filteredDefs <- do names <- allNames ctxt
|
filteredDefs <-
|
||||||
defs <- catMaybes <$> (traverse (flip lookupCtxtExact ctxt) names)
|
do names <- allNames ctxt
|
||||||
allDefs <- traverse (resolved ctxt) defs
|
defs <- traverse (flip lookupCtxtExact ctxt) names
|
||||||
(flip filterM) allDefs (\def => equivTypes def.type ty')
|
let defs = flip mapMaybe defs $ \ md =>
|
||||||
|
do d <- md
|
||||||
|
guard (visibleIn curr (fullname d) (visibility d))
|
||||||
|
pure d
|
||||||
|
allDefs <- traverse (resolved ctxt) defs
|
||||||
|
filterM (\def => equivTypes def.type ty') allDefs
|
||||||
put Ctxt defs
|
put Ctxt defs
|
||||||
doc <- traverse (docsOrSignature replFC) $ (.fullname) <$> filteredDefs
|
doc <- traverse (docsOrSignature replFC) $ (.fullname) <$> filteredDefs
|
||||||
pure $ Printed $ vsep $ pretty <$> (intersperse "\n" $ join doc)
|
pure $ Printed $ vsep $ pretty <$> (intersperse "\n" $ join doc)
|
||||||
process (TypeSearch _)
|
|
||||||
= pure $ REPLError $ reflow "Could not parse input as a type signature."
|
|
||||||
process (Missing n)
|
process (Missing n)
|
||||||
= do defs <- get Ctxt
|
= do defs <- get Ctxt
|
||||||
case !(lookupCtxtName n (gamma defs)) of
|
case !(lookupCtxtName n (gamma defs)) of
|
||||||
|
Loading…
Reference in New Issue
Block a user