Idris2/libs/contrib/Language/JSON/String/Lexer.idr
Alex Humphreys f3855d7100
Update contrib Text.Parser to match Library.Text.Parser (#1808)
Co-authored-by: Guillaume ALLAIS <guillaume.allais@ens-lyon.org>
2021-08-06 10:03:13 +01:00

43 lines
853 B
Idris

module Language.JSON.String.Lexer
import Data.Nat
import Language.JSON.String.Tokens
import Text.Lexer
%default total
export
quo : Lexer
quo = is '"'
export
esc : Lexer -> Lexer
esc = escape (is '\\')
private
unicodeEscape : Lexer
unicodeEscape = esc $ is 'u' <+> count (exactly 4) hexDigit
private
simpleEscape : Lexer
simpleEscape = esc $ oneOf "\"\\/bfnrt"
private
legalChar : Lexer
legalChar = non (quo <|> is '\\' <|> control)
private
jsonStringTokenMap : TokenMap JSONStringToken
jsonStringTokenMap = toTokenMap $
[ (quo, JSTQuote)
, (unicodeEscape, JSTUnicodeEscape)
, (simpleEscape, JSTSimpleEscape)
, (legalChar, JSTChar)
]
export
lexString : String -> Maybe (List (WithBounds JSONStringToken))
lexString x = case lex jsonStringTokenMap x of
(toks, _, _, "") => Just $ toks
_ => Nothing