Skip to content

Commit

Permalink
Merge pull request #4623 from unalos/4292
Browse files Browse the repository at this point in the history
Accepting hyphens in libs declaration in ipkg format.
  • Loading branch information
melted committed Dec 26, 2018
2 parents ecf297a + 78cdbb9 commit be43904
Show file tree
Hide file tree
Showing 2 changed files with 19 additions and 10 deletions.
13 changes: 9 additions & 4 deletions src/Idris/Package/Parser.hs
Original file line number Diff line number Diff line change
Expand Up @@ -13,9 +13,10 @@ import Idris.Imports
import Idris.Options (Opt)
import Idris.Package.Common
import Idris.Parser (moduleName)
import Idris.Parser.Helpers (Parser, Parsing, eol, iName, identifier, isEol,
lchar, packageName, parseErrorDoc, reserved,
runparser, someSpace, stringLiteral)
import Idris.Parser.Helpers (Parser, Parsing, eol, iName, identifier,
identifierWithExtraChars, isEol, lchar,
packageName, parseErrorDoc, reserved, runparser,
someSpace, stringLiteral)

import Control.Applicative
import Control.Monad.State.Strict
Expand Down Expand Up @@ -119,6 +120,9 @@ pOptions = do
Opts.Failure e -> fail $ fst $ Opts.renderFailure e ""
_ -> fail "Unexpected error"

libIdentifier :: Parsing m => m String
libIdentifier = identifierWithExtraChars "_'-."

pClause :: PParser ()
pClause = clause "executable" filename (\st v -> st { execout = Just v })
<|> clause "main" (iName []) (\st v -> st { idris_main = Just v })
Expand All @@ -129,7 +133,7 @@ pClause = clause "executable" filename (\st v -> st { execout = Just v })
in st { pkgdeps = ps `union` pkgdeps st
, idris_opts = pkgs ++ idris_opts st })
<|> clause "modules" (commaSep moduleName) (\st v -> st { modules = modules st ++ v })
<|> clause "libs" (commaSep identifier) (\st v -> st { libdeps = libdeps st ++ v })
<|> clause "libs" (commaSep libIdentifier) (\st v -> st { libdeps = libdeps st ++ v })
<|> clause "objs" (commaSep identifier) (\st v -> st { objs = objs st ++ v })
<|> clause "makefile" (iName []) (\st v -> st { makefile = Just (show v) })
<|> clause "tests" (commaSep (iName [])) (\st v -> st { idris_tests = idris_tests st ++ v })
Expand All @@ -142,3 +146,4 @@ pClause = clause "executable" filename (\st v -> st { execout = Just v })
<|> clause "brief" stringLiteral (\st v -> st { pkgbrief = Just v })
<|> clause "author" textUntilEol (\st v -> st { pkgauthor = Just v })
<|> clause "maintainer" textUntilEol (\st v -> st { pkgmaintainer = Just v })

16 changes: 10 additions & 6 deletions src/Idris/Parser/Helpers.hs
Original file line number Diff line number Diff line change
Expand Up @@ -36,6 +36,7 @@ module Idris.Parser.Helpers
, iName
, name
, identifier
, identifierWithExtraChars
, packageName
-- * Access
, accessibility
Expand Down Expand Up @@ -269,10 +270,10 @@ reservedIdentifiers = HS.fromList
, "rewrite", "syntax", "then", "total", "using", "where", "with"
]

identifierOrReserved :: Parsing m => m String
identifierOrReserved = token $ P.try $ do
identifierOrReservedWithExtraChars :: Parsing m => String -> m String
identifierOrReservedWithExtraChars extraChars = token $ P.try $ do
c <- P.satisfy isAlpha <|> P.oneOf "_"
cs <- P.many (P.satisfy isAlphaNum <|> P.oneOf "_'.")
cs <- P.many (P.satisfy isAlphaNum <|> P.oneOf extraChars)
return $ c : cs

char :: Parsing m => Char -> m Char
Expand All @@ -295,13 +296,16 @@ reserved name = token $ P.try $ do
P.notFollowedBy (P.satisfy isAlphaNum <|> P.oneOf "_'.") <?> "end of " ++ name

-- | Parses an identifier as a token
identifier :: Parsing m => m String
identifier = P.try $ do
ident <- identifierOrReserved
identifierWithExtraChars :: Parsing m => String -> m String
identifierWithExtraChars extraChars = P.try $ do
ident <- identifierOrReservedWithExtraChars extraChars
when (ident `HS.member` reservedIdentifiers) $ P.unexpected . P.Label . NonEmpty.fromList $ "reserved " ++ ident
when (ident == "_") $ P.unexpected . P.Label . NonEmpty.fromList $ "wildcard"
return ident

identifier :: Parsing m => m String
identifier = identifierWithExtraChars "_'."

-- | Parses an identifier with possible namespace as a name
iName :: Parsing m => [String] -> m Name
iName bad = maybeWithNS identifier bad <?> "name"
Expand Down

0 comments on commit be43904

Please sign in to comment.