mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-21 02:31:50 +03:00
bd093e9cba
As in Idris 1 - if an interface declares a method as total or covering, then all implementations have to satisfy that. Temporarily turn off %default directive again, at least until totality checking works as it should (this is probably better than removing it from various places because I'll forget to put them back)
27 lines
705 B
Idris
27 lines
705 B
Idris
module Language.JSON.String.Parser
|
|
|
|
import Language.JSON.String.Tokens
|
|
import Text.Lexer
|
|
import Text.Parser
|
|
|
|
%default total
|
|
|
|
private
|
|
stringChar : Grammar JSONStringToken True Char
|
|
stringChar = match JSTChar
|
|
<|> match JSTSimpleEscape
|
|
<|> match JSTUnicodeEscape
|
|
|
|
private
|
|
quotedString : Grammar JSONStringToken True String
|
|
quotedString = let q = match JSTQuote in
|
|
do chars <- between q q (many stringChar)
|
|
eof
|
|
pure $ pack chars
|
|
|
|
export
|
|
parseString : List JSONStringToken -> Maybe String
|
|
parseString toks = case parse quotedString toks of
|
|
Right (str, []) => Just str
|
|
_ => Nothing
|