mirror of
https://github.com/idris-lang/Idris2.git
synced 2024-12-19 09:12:34 +03:00
64 lines
1.5 KiB
Idris
64 lines
1.5 KiB
Idris
module Language.JSON.String.Tokens
|
|
|
|
import Data.String.Extra
|
|
import Text.Token
|
|
|
|
%default total
|
|
|
|
public export
|
|
data JSONStringTokenKind
|
|
= JSTQuote
|
|
| JSTChar
|
|
| JSTSimpleEscape
|
|
| JSTUnicodeEscape
|
|
|
|
public export
|
|
JSONStringToken : Type
|
|
JSONStringToken = Token JSONStringTokenKind
|
|
|
|
public export
|
|
Eq JSONStringTokenKind where
|
|
(==) JSTQuote JSTQuote = True
|
|
(==) JSTChar JSTChar = True
|
|
(==) JSTSimpleEscape JSTSimpleEscape = True
|
|
(==) JSTUnicodeEscape JSTUnicodeEscape = True
|
|
(==) _ _ = False
|
|
|
|
private
|
|
charValue : String -> Char
|
|
charValue x = case index 0 x of
|
|
Nothing => '\NUL'
|
|
Just c => c
|
|
|
|
private
|
|
simpleEscapeValue : String -> Char
|
|
simpleEscapeValue x
|
|
= case index 1 x of
|
|
Nothing => '\NUL'
|
|
Just c => case c of
|
|
'"' => '"'
|
|
'\\' => '\\'
|
|
'/' => '/'
|
|
'b' => '\b'
|
|
'f' => '\f'
|
|
'n' => '\n'
|
|
'r' => '\r'
|
|
't' => '\t'
|
|
_ => '\NUL'
|
|
|
|
private
|
|
unicodeEscapeValue : String -> Char
|
|
unicodeEscapeValue x = chr $ cast ("0x" ++ drop 2 x)
|
|
|
|
public export
|
|
TokenKind JSONStringTokenKind where
|
|
TokType JSTQuote = ()
|
|
TokType JSTChar = Char
|
|
TokType JSTSimpleEscape = Char
|
|
TokType JSTUnicodeEscape = Char
|
|
|
|
tokValue JSTQuote = const ()
|
|
tokValue JSTChar = charValue
|
|
tokValue JSTSimpleEscape = simpleEscapeValue
|
|
tokValue JSTUnicodeEscape = unicodeEscapeValue
|