author | wenzelm |
Wed, 11 Nov 2020 21:04:22 +0100 | |
changeset 72575 | c7ab83a0c564 |
parent 71776 | 5ef7f374e0f8 |
child 73340 | 0ffcad1f6130 |
permissions | -rw-r--r-- |
63992 | 1 |
/* Title: Pure/General/json.scala |
63644 | 2 |
Author: Makarius |
3 |
||
67819 | 4 |
Support for JSON: https://www.json.org/. |
67874 | 5 |
|
6 |
See also http://seriot.ch/parsing_json.php "Parsing JSON is a Minefield". |
|
63644 | 7 |
*/ |
8 |
||
9 |
package isabelle |
|
10 |
||
11 |
||
71601 | 12 |
import scala.util.matching.Regex |
66925
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
13 |
import scala.util.parsing.combinator.Parsers |
66922 | 14 |
import scala.util.parsing.combinator.lexical.Scanners |
15 |
import scala.util.parsing.input.CharArrayReader.EofCh |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
16 |
|
66922 | 17 |
|
63644 | 18 |
object JSON |
19 |
{ |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
20 |
type T = Any |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
21 |
type S = String |
63644 | 22 |
|
67841 | 23 |
object Object |
24 |
{ |
|
67857 | 25 |
type Entry = (String, JSON.T) |
26 |
||
67841 | 27 |
type T = Map[String, JSON.T] |
67850 | 28 |
val empty: Object.T = Map.empty |
29 |
||
67857 | 30 |
def apply(entries: Entry*): Object.T = Map(entries:_*) |
67841 | 31 |
|
71776 | 32 |
def unapply(obj: Any): Option[Object.T] = |
67841 | 33 |
obj match { |
34 |
case m: Map[_, _] if m.keySet.forall(_.isInstanceOf[String]) => |
|
35 |
Some(m.asInstanceOf[Object.T]) |
|
36 |
case _ => None |
|
37 |
} |
|
38 |
} |
|
66928 | 39 |
|
63644 | 40 |
|
66922 | 41 |
/* lexer */ |
42 |
||
43 |
object Kind extends Enumeration |
|
44 |
{ |
|
45 |
val KEYWORD, STRING, NUMBER, ERROR = Value |
|
46 |
} |
|
47 |
||
48 |
sealed case class Token(kind: Kind.Value, text: String) |
|
49 |
{ |
|
50 |
def is_keyword: Boolean = kind == Kind.KEYWORD |
|
66925
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
51 |
def is_keyword(name: String): Boolean = kind == Kind.KEYWORD && text == name |
66922 | 52 |
def is_string: Boolean = kind == Kind.STRING |
53 |
def is_number: Boolean = kind == Kind.NUMBER |
|
54 |
def is_error: Boolean = kind == Kind.ERROR |
|
55 |
} |
|
56 |
||
57 |
object Lexer extends Scanners with Scan.Parsers |
|
58 |
{ |
|
59 |
override type Elem = Char |
|
60 |
type Token = JSON.Token |
|
61 |
||
62 |
def errorToken(msg: String): Token = Token(Kind.ERROR, msg) |
|
63 |
||
67111
42f290d8ccbd
more accurate JSON parsing according to http://seriot.ch/parsing_json.php
wenzelm
parents:
66928
diff
changeset
|
64 |
val white_space: String = " \t\n\r" |
71601 | 65 |
override val whiteSpace: Regex = ("[" + white_space + "]+").r |
67111
42f290d8ccbd
more accurate JSON parsing according to http://seriot.ch/parsing_json.php
wenzelm
parents:
66928
diff
changeset
|
66 |
def whitespace: Parser[Any] = many(character(white_space.contains(_))) |
66922 | 67 |
|
71601 | 68 |
val letter: Parser[String] = one(character(Symbol.is_ascii_letter)) |
69 |
val letters1: Parser[String] = many1(character(Symbol.is_ascii_letter)) |
|
66922 | 70 |
|
71601 | 71 |
def digits: Parser[String] = many(character(Symbol.is_ascii_digit)) |
72 |
def digits1: Parser[String] = many1(character(Symbol.is_ascii_digit)) |
|
66922 | 73 |
|
74 |
||
75 |
/* keyword */ |
|
76 |
||
77 |
def keyword: Parser[Token] = |
|
66924 | 78 |
(letters1 | one(character("{}[],:".contains(_)))) ^^ (s => Token(Kind.KEYWORD, s)) |
66922 | 79 |
|
80 |
||
81 |
/* string */ |
|
82 |
||
83 |
def string: Parser[Token] = |
|
84 |
'\"' ~> rep(string_body) <~ '\"' ^^ (cs => Token(Kind.STRING, cs.mkString)) |
|
85 |
||
86 |
def string_body: Parser[Char] = |
|
67111
42f290d8ccbd
more accurate JSON parsing according to http://seriot.ch/parsing_json.php
wenzelm
parents:
66928
diff
changeset
|
87 |
elem("", c => c > '\u001f' && c != '\"' && c != '\\' && c != EofCh) | '\\' ~> string_escape |
66922 | 88 |
|
89 |
def string_escape: Parser[Char] = |
|
90 |
elem("", "\"\\/".contains(_)) | |
|
91 |
elem("", "bfnrt".contains(_)) ^^ |
|
92 |
{ case 'b' => '\b' case 'f' => '\f' case 'n' => '\n' case 'r' => '\r' case 't' => '\t' } | |
|
93 |
'u' ~> repeated(character("0123456789abcdefABCDEF".contains(_)), 4, 4) ^^ |
|
94 |
(s => Integer.parseInt(s, 16).toChar) |
|
95 |
||
96 |
def string_failure: Parser[Token] = '\"' ~> failure("Unterminated string") |
|
97 |
||
98 |
||
99 |
/* number */ |
|
100 |
||
101 |
def number: Parser[Token] = |
|
67111
42f290d8ccbd
more accurate JSON parsing according to http://seriot.ch/parsing_json.php
wenzelm
parents:
66928
diff
changeset
|
102 |
opt("-") ~ number_body ~ opt(letter) ^^ { |
66922 | 103 |
case a ~ b ~ None => Token(Kind.NUMBER, a.getOrElse("") + b) |
104 |
case a ~ b ~ Some(c) => |
|
105 |
errorToken("Invalid number format: " + quote(a.getOrElse("") + b + c)) |
|
106 |
} |
|
107 |
||
108 |
def number_body: Parser[String] = |
|
109 |
(zero | positive) ~ opt(number_fract) ~ opt(number_exp) ^^ |
|
110 |
{ case a ~ b ~ c => a + b.getOrElse("") + c.getOrElse("") } |
|
111 |
||
112 |
def number_fract: Parser[String] = "." ~ digits1 ^^ { case a ~ b => a + b } |
|
113 |
||
114 |
def number_exp: Parser[String] = |
|
115 |
one(character("eE".contains(_))) ~ maybe(character("-+".contains(_))) ~ digits1 ^^ |
|
116 |
{ case a ~ b ~ c => a + b + c } |
|
117 |
||
71601 | 118 |
def zero: Parser[String] = one(character(c => c == '0')) |
119 |
def nonzero: Parser[String] = one(character(c => c != '0' && Symbol.is_ascii_digit(c))) |
|
66922 | 120 |
def positive: Parser[String] = nonzero ~ digits ^^ { case a ~ b => a + b } |
121 |
||
122 |
||
123 |
/* token */ |
|
124 |
||
125 |
def token: Parser[Token] = |
|
126 |
keyword | (string | (string_failure | (number | failure("Illegal character")))) |
|
127 |
} |
|
128 |
||
129 |
||
66925
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
130 |
/* parser */ |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
131 |
|
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
132 |
trait Parser extends Parsers |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
133 |
{ |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
134 |
type Elem = Token |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
135 |
|
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
136 |
def $$$(name: String): Parser[Token] = elem(name, _.is_keyword(name)) |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
137 |
def string: Parser[String] = elem("string", _.is_string) ^^ (_.text) |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
138 |
def number: Parser[Double] = elem("number", _.is_number) ^^ (tok => tok.text.toDouble) |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
139 |
|
67841 | 140 |
def json_object: Parser[Object.T] = |
66925
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
141 |
$$$("{") ~> |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
142 |
repsep(string ~ ($$$(":") ~> json_value) ^^ { case a ~ b => (a, b) }, $$$(",")) <~ |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
143 |
$$$("}") ^^ (_.toMap) |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
144 |
|
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
145 |
def json_array: Parser[List[T]] = |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
146 |
$$$("[") ~> repsep(json_value, $$$(",")) <~ $$$("]") |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
147 |
|
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
148 |
def json_value: Parser[T] = |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
149 |
json_object | (json_array | (number | (string | |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
150 |
($$$("true") ^^^ true | ($$$("false") ^^^ false | ($$$("null") ^^^ null)))))) |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
151 |
|
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
152 |
def parse(input: CharSequence, strict: Boolean): T = |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
153 |
{ |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
154 |
val scanner = new Lexer.Scanner(Scan.char_reader(input)) |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
155 |
phrase(if (strict) json_object | json_array else json_value)(scanner) match { |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
156 |
case Success(json, _) => json |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
157 |
case NoSuccess(_, next) => error("Malformed JSON input at " + next.pos) |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
158 |
} |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
159 |
} |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
160 |
} |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
161 |
|
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
162 |
object Parser extends Parser |
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
163 |
|
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
164 |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
165 |
/* standard format */ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
166 |
|
66925
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
167 |
def parse(s: S, strict: Boolean = true): T = Parser.parse(s, strict) |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
168 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
169 |
object Format |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
170 |
{ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
171 |
def unapply(s: S): Option[T] = |
66925
8b117dc73278
separate JSON parser, which is slightly more general than (deprecated) scala.util.parsing.json.JSON;
wenzelm
parents:
66924
diff
changeset
|
172 |
try { Some(parse(s, strict = false)) } |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
173 |
catch { case ERROR(_) => None } |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
174 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
175 |
def apply(json: T): S = |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
176 |
{ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
177 |
val result = new StringBuilder |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
178 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
179 |
def string(s: String) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
180 |
{ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
181 |
result += '"' |
64761 | 182 |
result ++= |
183 |
s.iterator.map { |
|
184 |
case '"' => "\\\"" |
|
185 |
case '\\' => "\\\\" |
|
186 |
case '\b' => "\\b" |
|
187 |
case '\f' => "\\f" |
|
188 |
case '\n' => "\\n" |
|
189 |
case '\r' => "\\r" |
|
190 |
case '\t' => "\\t" |
|
191 |
case c => |
|
192 |
if (c <= '\u001f' || c >= '\u007f' && c <= '\u009f') "\\u%04x".format(c.toInt) |
|
193 |
else c |
|
194 |
}.mkString |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
195 |
result += '"' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
196 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
197 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
198 |
def array(list: List[T]) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
199 |
{ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
200 |
result += '[' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
201 |
Library.separate(None, list.map(Some(_))).foreach({ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
202 |
case None => result += ',' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
203 |
case Some(x) => json_format(x) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
204 |
}) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
205 |
result += ']' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
206 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
207 |
|
67841 | 208 |
def object_(obj: Object.T) |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
209 |
{ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
210 |
result += '{' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
211 |
Library.separate(None, obj.toList.map(Some(_))).foreach({ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
212 |
case None => result += ',' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
213 |
case Some((x, y)) => |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
214 |
string(x) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
215 |
result += ':' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
216 |
json_format(y) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
217 |
}) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
218 |
result += '}' |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
219 |
} |
63644 | 220 |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
221 |
def json_format(x: T) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
222 |
{ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
223 |
x match { |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
224 |
case null => result ++= "null" |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
225 |
case _: Int | _: Long | _: Boolean => result ++= x.toString |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
226 |
case n: Double => |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
227 |
val i = n.toLong |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
228 |
result ++= (if (i.toDouble == n) i.toString else n.toString) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
229 |
case s: String => string(s) |
67841 | 230 |
case Object(m) => object_(m) |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
231 |
case list: List[T] => array(list) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
232 |
case _ => error("Bad JSON value: " + x.toString) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
233 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
234 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
235 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
236 |
json_format(json) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
237 |
result.toString |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
238 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
239 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
240 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
241 |
|
67843 | 242 |
/* typed values */ |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
243 |
|
67842 | 244 |
object Value |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
245 |
{ |
67885 | 246 |
object UUID |
247 |
{ |
|
69458 | 248 |
def unapply(json: T): Option[isabelle.UUID.T] = |
67885 | 249 |
json match { |
69458 | 250 |
case x: java.lang.String => isabelle.UUID.unapply(x) |
67885 | 251 |
case _ => None |
252 |
} |
|
253 |
} |
|
254 |
||
71298 | 255 |
object String |
256 |
{ |
|
67842 | 257 |
def unapply(json: T): Option[java.lang.String] = |
258 |
json match { |
|
259 |
case x: java.lang.String => Some(x) |
|
260 |
case _ => None |
|
261 |
} |
|
262 |
} |
|
263 |
||
71305 | 264 |
object String0 |
265 |
{ |
|
266 |
def unapply(json: T): Option[java.lang.String] = |
|
267 |
json match { |
|
268 |
case null => Some("") |
|
269 |
case x: java.lang.String => Some(x) |
|
270 |
case _ => None |
|
271 |
} |
|
272 |
} |
|
273 |
||
71298 | 274 |
object Double |
275 |
{ |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
276 |
def unapply(json: T): Option[scala.Double] = |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
277 |
json match { |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
278 |
case x: scala.Double => Some(x) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
279 |
case x: scala.Long => Some(x.toDouble) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
280 |
case x: scala.Int => Some(x.toDouble) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
281 |
case _ => None |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
282 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
283 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
284 |
|
71298 | 285 |
object Long |
286 |
{ |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
287 |
def unapply(json: T): Option[scala.Long] = |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
288 |
json match { |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
289 |
case x: scala.Double if x.toLong.toDouble == x => Some(x.toLong) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
290 |
case x: scala.Long => Some(x) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
291 |
case x: scala.Int => Some(x.toLong) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
292 |
case _ => None |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
293 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
294 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
295 |
|
71298 | 296 |
object Int |
297 |
{ |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
298 |
def unapply(json: T): Option[scala.Int] = |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
299 |
json match { |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
300 |
case x: scala.Double if x.toInt.toDouble == x => Some(x.toInt) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
301 |
case x: scala.Long if x.toInt.toLong == x => Some(x.toInt) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
302 |
case x: scala.Int => Some(x) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
303 |
case _ => None |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
304 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
305 |
} |
67842 | 306 |
|
71298 | 307 |
object Boolean |
308 |
{ |
|
67842 | 309 |
def unapply(json: T): Option[scala.Boolean] = |
310 |
json match { |
|
311 |
case x: scala.Boolean => Some(x) |
|
312 |
case _ => None |
|
313 |
} |
|
314 |
} |
|
67843 | 315 |
|
316 |
object List |
|
317 |
{ |
|
318 |
def unapply[A](json: T, unapply: T => Option[A]): Option[List[A]] = |
|
319 |
json match { |
|
320 |
case xs: List[T] => |
|
321 |
val ys = xs.map(unapply) |
|
322 |
if (ys.forall(_.isDefined)) Some(ys.map(_.get)) else None |
|
323 |
case _ => None |
|
324 |
} |
|
325 |
} |
|
68943 | 326 |
|
71298 | 327 |
object Seconds |
328 |
{ |
|
68943 | 329 |
def unapply(json: T): Option[Time] = |
71601 | 330 |
Double.unapply(json).map(Time.seconds) |
68943 | 331 |
} |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
332 |
} |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
333 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
334 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
335 |
/* object values */ |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
336 |
|
67841 | 337 |
def optional(entry: (String, Option[T])): Object.T = |
64878 | 338 |
entry match { |
67850 | 339 |
case (name, Some(x)) => Object(name -> x) |
67841 | 340 |
case (_, None) => Object.empty |
64878 | 341 |
} |
342 |
||
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
343 |
def value(obj: T, name: String): Option[T] = |
63644 | 344 |
obj match { |
67841 | 345 |
case Object(m) => m.get(name) |
63644 | 346 |
case _ => None |
347 |
} |
|
348 |
||
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
349 |
def value[A](obj: T, name: String, unapply: T => Option[A]): Option[A] = |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
350 |
for { |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
351 |
json <- value(obj, name) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
352 |
x <- unapply(json) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
353 |
} yield x |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
354 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
355 |
def array(obj: T, name: String): Option[List[T]] = |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
356 |
value(obj, name) match { |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
357 |
case Some(a: List[T]) => Some(a) |
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
358 |
case _ => None |
63644 | 359 |
} |
360 |
||
67864 | 361 |
def value_default[A](obj: T, name: String, unapply: T => Option[A], default: => A): Option[A] = |
67843 | 362 |
value(obj, name) match { |
363 |
case None => Some(default) |
|
364 |
case Some(json) => unapply(json) |
|
365 |
} |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
366 |
|
69458 | 367 |
def uuid(obj: T, name: String): Option[UUID.T] = |
67885 | 368 |
value(obj, name, Value.UUID.unapply) |
369 |
||
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
370 |
def string(obj: T, name: String): Option[String] = |
67842 | 371 |
value(obj, name, Value.String.unapply) |
67864 | 372 |
def string_default(obj: T, name: String, default: => String = ""): Option[String] = |
67843 | 373 |
value_default(obj, name, Value.String.unapply, default) |
63644 | 374 |
|
71305 | 375 |
def string0(obj: T, name: String): Option[String] = |
376 |
value(obj, name, Value.String0.unapply) |
|
377 |
def string0_default(obj: T, name: String, default: => String = ""): Option[String] = |
|
378 |
value_default(obj, name, Value.String0.unapply, default) |
|
379 |
||
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
380 |
def double(obj: T, name: String): Option[Double] = |
67842 | 381 |
value(obj, name, Value.Double.unapply) |
67864 | 382 |
def double_default(obj: T, name: String, default: => Double = 0.0): Option[Double] = |
67843 | 383 |
value_default(obj, name, Value.Double.unapply, default) |
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
384 |
|
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
385 |
def long(obj: T, name: String): Option[Long] = |
67842 | 386 |
value(obj, name, Value.Long.unapply) |
67864 | 387 |
def long_default(obj: T, name: String, default: => Long = 0): Option[Long] = |
67843 | 388 |
value_default(obj, name, Value.Long.unapply, default) |
63644 | 389 |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
390 |
def int(obj: T, name: String): Option[Int] = |
67842 | 391 |
value(obj, name, Value.Int.unapply) |
67864 | 392 |
def int_default(obj: T, name: String, default: => Int = 0): Option[Int] = |
67843 | 393 |
value_default(obj, name, Value.Int.unapply, default) |
63644 | 394 |
|
64545
25045094d7bb
clarified JSON operations (see isabelle_vscode/a7931dc2a1ab);
wenzelm
parents:
63992
diff
changeset
|
395 |
def bool(obj: T, name: String): Option[Boolean] = |
67842 | 396 |
value(obj, name, Value.Boolean.unapply) |
67864 | 397 |
def bool_default(obj: T, name: String, default: => Boolean = false): Option[Boolean] = |
67843 | 398 |
value_default(obj, name, Value.Boolean.unapply, default) |
399 |
||
400 |
def list[A](obj: T, name: String, unapply: T => Option[A]): Option[List[A]] = |
|
401 |
value(obj, name, Value.List.unapply(_, unapply)) |
|
67864 | 402 |
def list_default[A](obj: T, name: String, unapply: T => Option[A], default: => List[A] = Nil) |
67843 | 403 |
: Option[List[A]] = value_default(obj, name, Value.List.unapply(_, unapply), default) |
68742 | 404 |
|
405 |
def strings(obj: T, name: String): Option[List[String]] = |
|
71601 | 406 |
list(obj, name, JSON.Value.String.unapply) |
68742 | 407 |
def strings_default(obj: T, name: String, default: => List[String] = Nil): Option[List[String]] = |
71601 | 408 |
list_default(obj, name, JSON.Value.String.unapply, default) |
68943 | 409 |
|
410 |
def seconds(obj: T, name: String): Option[Time] = |
|
411 |
value(obj, name, Value.Seconds.unapply) |
|
412 |
def seconds_default(obj: T, name: String, default: => Time = Time.zero): Option[Time] = |
|
413 |
value_default(obj, name, Value.Seconds.unapply, default) |
|
63644 | 414 |
} |