mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-25 12:42:02 +03:00
[ fix #1224 ] moduleIdent must be capitalised
This commit is contained in:
parent
00929deed6
commit
2df49ebdd4
@ -90,6 +90,7 @@ mkModTree loc done modFP mod
|
|||||||
(\err =>
|
(\err =>
|
||||||
case err of
|
case err of
|
||||||
CyclicImports _ => throw err
|
CyclicImports _ => throw err
|
||||||
|
ParseFail _ _ => throw err
|
||||||
_ => pure (MkModTree mod Nothing []))
|
_ => pure (MkModTree mod Nothing []))
|
||||||
|
|
||||||
data DoneMod : Type where
|
data DoneMod : Type where
|
||||||
|
@ -1496,7 +1496,7 @@ import_ fname indents
|
|||||||
= do b <- bounds (do keyword "import"
|
= do b <- bounds (do keyword "import"
|
||||||
reexp <- option False (do keyword "public"
|
reexp <- option False (do keyword "public"
|
||||||
pure True)
|
pure True)
|
||||||
ns <- moduleIdent
|
ns <- mustWork moduleIdent
|
||||||
nsAs <- option (miAsNamespace ns)
|
nsAs <- option (miAsNamespace ns)
|
||||||
(do exactIdent "as"
|
(do exactIdent "as"
|
||||||
mustWork namespaceId)
|
mustWork namespaceId)
|
||||||
@ -1511,7 +1511,7 @@ prog fname
|
|||||||
= do b <- bounds (do doc <- option "" documentation
|
= do b <- bounds (do doc <- option "" documentation
|
||||||
nspace <- option (nsAsModuleIdent mainNS)
|
nspace <- option (nsAsModuleIdent mainNS)
|
||||||
(do keyword "module"
|
(do keyword "module"
|
||||||
moduleIdent)
|
mustWork moduleIdent)
|
||||||
imports <- block (import_ fname)
|
imports <- block (import_ fname)
|
||||||
pure (doc, nspace, imports))
|
pure (doc, nspace, imports))
|
||||||
ds <- block (topDecl fname)
|
ds <- block (topDecl fname)
|
||||||
@ -1525,7 +1525,7 @@ progHdr fname
|
|||||||
= do b <- bounds (do doc <- option "" documentation
|
= do b <- bounds (do doc <- option "" documentation
|
||||||
nspace <- option (nsAsModuleIdent mainNS)
|
nspace <- option (nsAsModuleIdent mainNS)
|
||||||
(do keyword "module"
|
(do keyword "module"
|
||||||
moduleIdent)
|
mustWork moduleIdent)
|
||||||
imports <- block (import_ fname)
|
imports <- block (import_ fname)
|
||||||
pure (doc, nspace, imports))
|
pure (doc, nspace, imports))
|
||||||
(doc, nspace, imports) <- pure b.val
|
(doc, nspace, imports) <- pure b.val
|
||||||
|
@ -234,12 +234,7 @@ namespaceId = do
|
|||||||
|
|
||||||
export
|
export
|
||||||
moduleIdent : Rule ModuleIdent
|
moduleIdent : Rule ModuleIdent
|
||||||
moduleIdent
|
moduleIdent = nsAsModuleIdent <$> namespaceId
|
||||||
= terminal "Expected module identifier"
|
|
||||||
(\x => case x of
|
|
||||||
DotSepIdent ns n => Just (mkModuleIdent (Just ns) n)
|
|
||||||
Ident i => Just (mkModuleIdent Nothing i)
|
|
||||||
_ => Nothing)
|
|
||||||
|
|
||||||
export
|
export
|
||||||
unqualifiedName : Rule String
|
unqualifiedName : Rule String
|
||||||
|
@ -23,7 +23,8 @@ runParserTo fname lit reject str p
|
|||||||
|
|
||||||
export
|
export
|
||||||
runParser : {e : _} ->
|
runParser : {e : _} ->
|
||||||
(fname : String) -> Maybe LiterateStyle -> String -> Grammar Token e ty -> Either Error ty
|
(fname : String) -> Maybe LiterateStyle -> String ->
|
||||||
|
Grammar Token e ty -> Either Error ty
|
||||||
runParser fname lit = runParserTo fname lit (pred $ const False)
|
runParser fname lit = runParserTo fname lit (pred $ const False)
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
|
@ -29,8 +29,10 @@ fromLexError fname (_, l, c, _)
|
|||||||
= LexFail (MkFC fname (l, c) (l, c + 1)) "Can't recognise token."
|
= LexFail (MkFC fname (l, c) (l, c + 1)) "Can't recognise token."
|
||||||
|
|
||||||
export
|
export
|
||||||
fromParsingError : (Show token, Pretty token) => String -> ParsingError token -> Error
|
fromParsingError : (Show token, Pretty token) =>
|
||||||
fromParsingError fname (Error msg Nothing) = ParseFail (MkFC fname (0, 0) (0, 0)) (msg +> '.')
|
String -> ParsingError token -> Error
|
||||||
|
fromParsingError fname (Error msg Nothing)
|
||||||
|
= ParseFail (MkFC fname (0, 0) (0, 0)) (msg +> '.')
|
||||||
fromParsingError fname (Error msg (Just t))
|
fromParsingError fname (Error msg (Just t))
|
||||||
= let l = t.startLine
|
= let l = t.startLine
|
||||||
c = t.startCol in
|
c = t.startCol in
|
||||||
|
5
tests/idris2/perror008/Issue1224a.idr
Normal file
5
tests/idris2/perror008/Issue1224a.idr
Normal file
@ -0,0 +1,5 @@
|
|||||||
|
module ggg
|
||||||
|
|
||||||
|
export
|
||||||
|
yyy : Int
|
||||||
|
yyy = 0
|
4
tests/idris2/perror008/Issue1224b.idr
Normal file
4
tests/idris2/perror008/Issue1224b.idr
Normal file
@ -0,0 +1,4 @@
|
|||||||
|
import ggg
|
||||||
|
|
||||||
|
t : Int
|
||||||
|
t = yyy
|
@ -52,3 +52,15 @@ Issue710f.idr:2:15--2:16
|
|||||||
2 | constructor cons
|
2 | constructor cons
|
||||||
^
|
^
|
||||||
|
|
||||||
|
Uncaught error: Error: Expected a capitalised identifier, got: ggg.
|
||||||
|
|
||||||
|
Issue1224a.idr:1:8--1:9
|
||||||
|
1 | module ggg
|
||||||
|
^
|
||||||
|
|
||||||
|
Uncaught error: Error: Expected a capitalised identifier, got: ggg.
|
||||||
|
|
||||||
|
Issue1224b.idr:1:8--1:9
|
||||||
|
1 | import ggg
|
||||||
|
^
|
||||||
|
|
||||||
|
@ -5,4 +5,7 @@ $1 --no-color --console-width 0 --check Issue710d.idr || true
|
|||||||
$1 --no-color --console-width 0 --check Issue710e.idr || true
|
$1 --no-color --console-width 0 --check Issue710e.idr || true
|
||||||
$1 --no-color --console-width 0 --check Issue710f.idr || true
|
$1 --no-color --console-width 0 --check Issue710f.idr || true
|
||||||
|
|
||||||
|
$1 --no-color --console-width 0 --check Issue1224a.idr || true
|
||||||
|
$1 --no-color --console-width 0 --check Issue1224b.idr || true
|
||||||
|
|
||||||
rm -rf build
|
rm -rf build
|
||||||
|
Loading…
Reference in New Issue
Block a user