Idris2/libs/contrib/Language/JSON/Parser.idr

80 lines
2.0 KiB
Idris
Raw Normal View History

2020-05-18 15:59:07 +03:00
module Language.JSON.Parser
import Language.JSON.Data
import Text.Parser
import Data.List
import public Language.JSON.Tokens
%default total
2020-05-18 15:59:07 +03:00
private
punct : Punctuation -> Grammar state JSONToken True ()
2020-05-18 15:59:07 +03:00
punct p = match $ JTPunct p
private
rawString : Grammar state JSONToken True String
2020-05-18 15:59:07 +03:00
rawString = do mstr <- match JTString
the (Grammar _ _ False _) $
2020-05-18 15:59:07 +03:00
case mstr of
Just str => pure str
Nothing => fail "invalid string"
mutual
private
json : Grammar state JSONToken True JSON
2020-05-18 15:59:07 +03:00
json = object
<|> array
<|> string
<|> boolean
<|> number
<|> null
private
object : Grammar state JSONToken True JSON
2020-05-18 15:59:07 +03:00
object = do punct $ Curly Open
commit
props <- properties
punct $ Curly Close
pure $ JObject props
where
properties : Grammar state JSONToken False (List (String, JSON))
2020-05-18 15:59:07 +03:00
properties = sepBy (punct Comma) $
do key <- rawString
punct Colon
value <- json
pure (key, value)
private
array : Grammar state JSONToken True JSON
2020-05-18 15:59:07 +03:00
array = do punct (Square Open)
commit
vals <- values
punct (Square Close)
pure (JArray vals)
where
values : Grammar state JSONToken False (List JSON)
2020-05-18 15:59:07 +03:00
values = sepBy (punct Comma) json
private
string : Grammar state JSONToken True JSON
2020-05-18 15:59:07 +03:00
string = map JString rawString
private
boolean : Grammar state JSONToken True JSON
2020-05-18 15:59:07 +03:00
boolean = map JBoolean $ match JTBoolean
private
number : Grammar state JSONToken True JSON
2020-05-18 15:59:07 +03:00
number = map JNumber $ match JTNumber
private
null : Grammar state JSONToken True JSON
2020-05-18 15:59:07 +03:00
null = map (const JNull) $ match JTNull
export
parseJSON : List (WithBounds JSONToken) -> Maybe JSON
2020-05-18 15:59:07 +03:00
parseJSON toks = case parse json $ filter (not . ignored) toks of
Right (j, []) => Just j
_ => Nothing