[ fix ] properly handle Namespace blocks for DocStrings (#1342)

This commit is contained in:
G. Allais 2021-04-28 09:31:01 +01:00 committed by GitHub
parent 4224c58651
commit d32daaf36a
No known key found for this signature in database
GPG Key ID: 4AEE18F83AFDEB23
8 changed files with 33 additions and 6 deletions

View File

@ -1998,6 +1998,20 @@ extendNS ns
= do defs <- get Ctxt = do defs <- get Ctxt
put Ctxt (record { currentNS $= (<.> ns) } defs) put Ctxt (record { currentNS $= (<.> ns) } defs)
export
withExtendedNS : {auto c : Ref Ctxt Defs} ->
Namespace -> Core a -> Core a
withExtendedNS ns act
= do defs <- get Ctxt
let cns = currentNS defs
put Ctxt (record { currentNS = cns <.> ns } defs)
ma <- catch (Right <$> act) (pure . Left)
defs <- get Ctxt
put Ctxt (record { currentNS = cns } defs)
case ma of
Left err => throw err
Right a => pure a
-- Get the name as it would be defined in the current namespace -- Get the name as it would be defined in the current namespace
-- i.e. if it doesn't have an explicit namespace already, add it, -- i.e. if it doesn't have an explicit namespace already, add it,
-- otherwise leave it alone -- otherwise leave it alone

View File

@ -953,7 +953,8 @@ mutual
mds' <- traverse (desugarDecl ps) mds mds' <- traverse (desugarDecl ps) mds
pure (concat mds') pure (concat mds')
desugarDecl ps (PNamespace fc ns decls) desugarDecl ps (PNamespace fc ns decls)
= do ds <- traverse (desugarDecl ps) decls = withExtendedNS ns $ do
ds <- traverse (desugarDecl ps) decls
pure [INamespace fc ns (concat ds)] pure [INamespace fc ns (concat ds)]
desugarDecl ps (PTransform fc n lhs rhs) desugarDecl ps (PTransform fc n lhs rhs)
= do (bound, blhs) <- bindNames False !(desugar LHS ps lhs) = do (bound, blhs) <- bindNames False !(desugar LHS ps lhs)

View File

@ -1,6 +1,7 @@
module Idris.DocString module Idris.DocString
import Core.Context import Core.Context
import Core.Context.Log
import Core.Core import Core.Core
import Core.Env import Core.Env
import Core.TT import Core.TT
@ -68,6 +69,8 @@ addDocString : {auto c : Ref Ctxt Defs} ->
Core () Core ()
addDocString n_in doc addDocString n_in doc
= do n <- inCurrentNS n_in = do n <- inCurrentNS n_in
log "doc.record" 50 $
"Adding doc for " ++ show n_in ++ " (aka " ++ show n ++ " in current NS)"
syn <- get Syn syn <- get Syn
put Syn (record { docstrings $= addName n doc, put Syn (record { docstrings $= addName n doc,
saveDocstrings $= insert n () } syn) saveDocstrings $= insert n () } syn)

View File

@ -47,12 +47,8 @@ process eopts nest env (IParameters fc ps decls)
process eopts nest env (IRecord fc ns vis rec) process eopts nest env (IRecord fc ns vis rec)
= processRecord eopts nest env ns vis rec = processRecord eopts nest env ns vis rec
process eopts nest env (INamespace fc ns decls) process eopts nest env (INamespace fc ns decls)
= do defs <- get Ctxt = withExtendedNS ns $
let cns = currentNS defs
extendNS ns
traverse_ (processDecl eopts nest env) decls traverse_ (processDecl eopts nest env) decls
defs <- get Ctxt
put Ctxt (record { currentNS = cns } defs)
process eopts nest env (ITransform fc n lhs rhs) process eopts nest env (ITransform fc n lhs rhs)
= processTransform eopts nest env fc n lhs rhs = processTransform eopts nest env fc n lhs rhs
process eopts nest env (IRunElabDecl fc tm) process eopts nest env (IRunElabDecl fc tm)

View File

@ -18,6 +18,7 @@ Main> Prelude.List : Type -> Type
Nil : List a Nil : List a
Empty list Empty list
(::) : a -> List a -> List a (::) : a -> List a -> List a
A non-empty list, consisting of a head element and the rest of the list.
Main> Prelude.Show : Type -> Type Main> Prelude.Show : Type -> Type
Things that have a canonical `String` representation. Things that have a canonical `String` representation.
Parameters: ty Parameters: ty

View File

@ -12,3 +12,11 @@ export
||| Hello! ||| Hello!
hello : Int -> Int hello : Int -> Int
hello x = x*2 hello x = x*2
namespace NS
namespace NestedNS
||| A type of Foo
Foo : Type
Foo = ()

View File

@ -3,4 +3,7 @@ Doc> hello : Int -> Int
Hello! Hello!
world : Nat -> Nat world : Nat -> Nat
World! World!
Doc> Doc.NS.NestedNS.Foo : Type
A type of Foo
Doc> Bye for now! Doc> Bye for now!

View File

@ -1,2 +1,3 @@
:browse Doc :browse Doc
:doc Foo
:q :q